Nominal/Ex/Multi_Recs2.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 17 Dec 2011 17:08:47 +0000
branchNominal2-Isabelle2011-1
changeset 3071 11f6a561eb4b
parent 2950 0911cb7bf696
permissions -rw-r--r--
cleaned examples for stable branch
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     1
theory Multi_Recs2
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2440
diff changeset
     2
imports "../Nominal2"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     5
(* 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     6
  multiple recursive binders - multiple letrecs with multiple 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     7
  clauses for each functions
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     8
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
     9
  example 8 from Peter Sewell's bestiary (originally due
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    10
  to James Cheney) 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    11
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    12
*)
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
atom_decl name
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    16
nominal_datatype fun_recs: exp =
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    17
  Var name
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    18
| Unit 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    19
| Pair exp exp
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2787
diff changeset
    20
| LetRec l::lrbs e::exp binds (set) "b_lrbs l" in l e
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
and fnclause =
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2787
diff changeset
    22
  K x::name p::pat f::exp binds (set) "b_pat p" in f
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
and fnclauses =
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
  S fnclause
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
| ORs fnclause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    26
and lrb =
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
  Clause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    28
and lrbs =
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    29
  Single lrb
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    30
| More lrb lrbs
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    31
and pat =
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
  PVar name
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
| PUnit
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    34
| PPair pat pat
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
binder
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    36
  b_lrbs :: "lrbs \<Rightarrow> atom set" and
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    37
  b_pat :: "pat \<Rightarrow> atom set" and
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
  b_fnclauses :: "fnclauses \<Rightarrow> atom set" and
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  b_fnclause :: "fnclause \<Rightarrow> atom set" and
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    40
  b_lrb :: "lrb \<Rightarrow> atom set"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
where
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    42
  "b_lrbs (Single l) = b_lrb l"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    43
| "b_lrbs (More l ls) = b_lrb l \<union> b_lrbs ls"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
| "b_pat (PVar x) = {atom x}"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
| "b_pat (PUnit) = {}"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
| "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
| "b_fnclauses (S fc) = (b_fnclause fc)"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
| "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    49
| "b_lrb (Clause fcs) = (b_fnclauses fcs)"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    50
| "b_fnclause (K x pat exp) = {atom x}"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
2593
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    52
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    53
thm fun_recs.permute_bn
2593
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    54
thm fun_recs.perm_bn_alpha
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    55
thm fun_recs.perm_bn_simps
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    56
thm fun_recs.bn_finite
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    57
thm fun_recs.inducts
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    58
thm fun_recs.distinct
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    59
thm fun_recs.induct
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    60
thm fun_recs.inducts
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    61
thm fun_recs.exhaust
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    62
thm fun_recs.fv_defs
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    63
thm fun_recs.bn_defs
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    64
thm fun_recs.perm_simps
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    65
thm fun_recs.eq_iff
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    66
thm fun_recs.fv_bn_eqvt
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    67
thm fun_recs.size_eqvt
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    68
thm fun_recs.supports
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    69
thm fun_recs.fsupp
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    70
thm fun_recs.supp
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    71
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    72
thm fun_recs.distinct
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    73
thm fun_recs.induct
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    74
thm fun_recs.inducts
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    75
thm fun_recs.exhaust
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    76
thm fun_recs.fv_defs
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    77
thm fun_recs.bn_defs
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    78
thm fun_recs.perm_simps
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    79
thm fun_recs.eq_iff
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    80
thm fun_recs.fv_bn_eqvt
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    81
thm fun_recs.size_eqvt
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    82
thm fun_recs.supports
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    83
thm fun_recs.fsupp
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2477
diff changeset
    84
thm fun_recs.supp
2475
486d4647bb37 supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    85
2629
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
    86
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
end
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90