Nominal/Ex/Multi_Recs2.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 31 Mar 2011 15:25:35 +0200
changeset 2751 3b8232f56941
parent 2630 8268b277d240
child 2787 1a6593bc494d
permissions -rw-r--r--
final version of slides
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
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
    20
| LetRec l::lrbs e::exp bind (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 =
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    22
  K x::name p::pat f::exp bind (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
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
    86
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
    87
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
    88
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
    89
lemma
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
    90
  fixes c::"'a::fs"
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
    91
  assumes "\<And>name c. P1 c (Var name)"
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
    92
  and "\<And>c. P1 c Unit"
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
    93
  and "\<And>exp1 exp2 c. \<lbrakk>\<And>c. P1 c exp1; \<And>c. P1 c exp2\<rbrakk> \<Longrightarrow> P1 c (Multi_Recs2.Pair exp1 exp2)"
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
    94
  and "\<And>lrbs exp c. \<lbrakk>b_lrbs lrbs \<sharp>* c; \<And>c. P5 c lrbs; \<And>c. P1 c exp\<rbrakk> \<Longrightarrow> P1 c (LetRec lrbs exp)"
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
    95
  and "\<And>name pat exp c. \<lbrakk>b_pat pat \<sharp>* c; \<And>c. P6 c pat; \<And>c. P1 c exp\<rbrakk> \<Longrightarrow> P2 c (K name pat exp)"
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
    96
  and "\<And>fnclause c. (\<And>c. P2 c fnclause) \<Longrightarrow> P3 c (S fnclause)"
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
    97
  and "\<And>fnclause fnclauses c. \<lbrakk>\<And>c. P2 c fnclause; \<And>c. P3 c fnclauses\<rbrakk> \<Longrightarrow> 
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
    98
    P3 c (ORs fnclause fnclauses)"
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
    99
  and "\<And>fnclauses c. (\<And>c. P3 c fnclauses) \<Longrightarrow> P4 c (Clause fnclauses)"
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
   100
  and "\<And>lrb c. (\<And>c. P4 c lrb) \<Longrightarrow> P5 c (Single lrb)"
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
   101
  and "\<And>lrb lrbs c. \<lbrakk>\<And>c. P4 c lrb; \<And>c. P5 c lrbs\<rbrakk> \<Longrightarrow> P5 c (More lrb lrbs)"
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
   102
  and "\<And>name c. P6 c (PVar name)"
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
   103
  and "\<And>c. P6 c PUnit"
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
   104
  and "\<And>pat1 pat2 c. \<lbrakk>\<And>c. P6 c pat1; \<And>c. P6 c pat2\<rbrakk> \<Longrightarrow> P6 c (PPair pat1 pat2)"
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
   105
  shows "P1 c exp" and "P2 c fnclause" and "P3 c fnclauses" and "P4 c lrb" and "P5 c lrbs" and "P6 c pat"
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
   106
  apply(raw_tactic {* Induction_Schema.induction_schema_tac @{context} @{thms assms} *})
2630
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   107
  apply(rule_tac y="exp" and c="c" in fun_recs.strong_exhaust(1))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   108
  apply(simp_all)[4]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   109
  apply(rule_tac ya="fnclause" and c="c" in fun_recs.strong_exhaust(2))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   110
  apply(simp_all)[1]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   111
  apply(rule_tac yb="fnclauses" in fun_recs.strong_exhaust(3))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   112
  apply(simp_all)[2]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   113
  apply(rule_tac yc="lrb" in fun_recs.strong_exhaust(4))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   114
  apply(simp_all)[1]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   115
  apply(rule_tac yd="lrbs" in fun_recs.strong_exhaust(5))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   116
  apply(simp_all)[2]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   117
  apply(rule_tac ye="pat" in fun_recs.strong_exhaust(6))
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   118
  apply(simp_all)[3]
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   119
  apply(tactic {* prove_termination_ind @{context} 1 *})
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   120
  apply(simp_all add: fun_recs.size)
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2629
diff changeset
   121
  done
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
   122
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
   123
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
   124
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126