| author | smolkas | 
| Wed, 09 Jan 2013 20:46:32 +0100 | |
| changeset 50786 | af8ecf09a58c | 
| parent 47432 | e1576d13e933 | 
| child 51717 | 9e7d1c139569 | 
| permissions | -rw-r--r-- | 
| 33471 | 1 | (* Title: HOL/Tools/Function/induction_schema.ML | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | ||
| 4 | A method to prove induction schemas. | |
| 5 | *) | |
| 6 | ||
| 7 | signature INDUCTION_SCHEMA = | |
| 8 | sig | |
| 9 | val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic) | |
| 10 | -> Proof.context -> thm list -> tactic | |
| 11 | val induction_schema_tac : Proof.context -> thm list -> tactic | |
| 12 | end | |
| 13 | ||
| 14 | ||
| 15 | structure Induction_Schema : INDUCTION_SCHEMA = | |
| 16 | struct | |
| 17 | ||
| 18 | open Function_Lib | |
| 19 | ||
| 20 | type rec_call_info = int * (string * typ) list * term list * term list | |
| 21 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 22 | datatype scheme_case = SchemeCase of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 23 |  {bidx : int,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 24 | qs: (string * typ) list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 25 | oqnames: string list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 26 | gs: term list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 27 | lhs: term list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 28 | rs: rec_call_info list} | 
| 33471 | 29 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 30 | datatype scheme_branch = SchemeBranch of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 31 |  {P : term,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 32 | xs: (string * typ) list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 33 | ws: (string * typ) list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 34 | Cs: term list} | 
| 33471 | 35 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 36 | datatype ind_scheme = IndScheme of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 37 |  {T: typ, (* sum of products *)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 38 | branches: scheme_branch list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 39 | cases: scheme_case list} | 
| 33471 | 40 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41225diff
changeset | 41 | val ind_atomize = Raw_Simplifier.rewrite true @{thms induct_atomize}
 | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41225diff
changeset | 42 | val ind_rulify = Raw_Simplifier.rewrite true @{thms induct_rulify}
 | 
| 33471 | 43 | |
| 44 | fun meta thm = thm RS eq_reflection | |
| 45 | ||
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41225diff
changeset | 46 | val sum_prod_conv = Raw_Simplifier.rewrite true | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 47 |   (map meta (@{thm split_conv} :: @{thms sum.cases}))
 | 
| 33471 | 48 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 49 | fun term_conv thy cv t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 50 | cv (cterm_of thy t) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 51 | |> prop_of |> Logic.dest_equals |> snd | 
| 33471 | 52 | |
| 53 | fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)) | |
| 54 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 55 | fun dest_hhf ctxt t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 56 | let | 
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42361diff
changeset | 57 | val ((params, imp), ctxt') = Variable.focus t ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 58 | in | 
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42361diff
changeset | 59 | (ctxt', map #2 params, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 60 | end | 
| 33471 | 61 | |
| 62 | fun mk_scheme' ctxt cases concl = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 63 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 64 | fun mk_branch concl = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 65 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 66 | val (_, ws, Cs, _ $ Pxs) = dest_hhf ctxt concl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 67 | val (P, xs) = strip_comb Pxs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 68 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 69 |         SchemeBranch { P=P, xs=map dest_Free xs, ws=ws, Cs=Cs }
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 70 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 71 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 72 | val (branches, cases') = (* correction *) | 
| 41418 | 73 | case Logic.dest_conjunctions concl of | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 74 | [conc] => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 75 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 76 | val _ $ Pxs = Logic.strip_assums_concl conc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 77 | val (P, _) = strip_comb Pxs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 78 | val (cases', conds) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 79 | take_prefix (Term.exists_subterm (curry op aconv P)) cases | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 80 | val concl' = fold_rev (curry Logic.mk_implies) conds conc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 81 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 82 | ([mk_branch concl'], cases') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 83 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 84 | | concls => (map mk_branch concls, cases) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 85 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 86 | fun mk_case premise = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 87 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 88 | val (ctxt', qs, prems, _ $ Plhs) = dest_hhf ctxt premise | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 89 | val (P, lhs) = strip_comb Plhs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 90 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 91 | fun bidx Q = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 92 |           find_index (fn SchemeBranch {P=P',...} => Q aconv P') branches
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 93 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 94 | fun mk_rcinfo pr = | 
| 33471 | 95 | let | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 96 | val (_, Gvs, Gas, _ $ Phyp) = dest_hhf ctxt' pr | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 97 | val (P', rcs) = strip_comb Phyp | 
| 33471 | 98 | in | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 99 | (bidx P', Gvs, Gas, rcs) | 
| 33471 | 100 | end | 
| 101 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 102 |         fun is_pred v = exists (fn SchemeBranch {P,...} => v aconv P) branches
 | 
| 33471 | 103 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 104 | val (gs, rcprs) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 105 | take_prefix (not o Term.exists_subterm is_pred) prems | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 106 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 107 |         SchemeCase {bidx=bidx P, qs=qs, oqnames=map fst qs(*FIXME*),
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 108 | gs=gs, lhs=lhs, rs=map mk_rcinfo rcprs} | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 109 | end | 
| 33471 | 110 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 111 |     fun PT_of (SchemeBranch { xs, ...}) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 112 | foldr1 HOLogic.mk_prodT (map snd xs) | 
| 33471 | 113 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 114 | val ST = Balanced_Tree.make (uncurry SumTree.mk_sumT) (map PT_of branches) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 115 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 116 |     IndScheme {T=ST, cases=map mk_case cases', branches=branches }
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 117 | end | 
| 33471 | 118 | |
| 119 | fun mk_completeness ctxt (IndScheme {cases, branches, ...}) bidx =
 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 120 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 121 |     val SchemeBranch { xs, ws, Cs, ... } = nth branches bidx
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 122 |     val relevant_cases = filter (fn SchemeCase {bidx=bidx', ...} => bidx' = bidx) cases
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 123 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 124 |     val allqnames = fold (fn SchemeCase {qs, ...} => fold (insert (op =) o Free) qs) relevant_cases []
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 125 |     val (Pbool :: xs') = map Free (Variable.variant_frees ctxt allqnames (("P", HOLogic.boolT) :: xs))
 | 
| 42361 | 126 | val Cs' = map (Pattern.rewrite_term (Proof_Context.theory_of ctxt) (filter_out (op aconv) (map Free xs ~~ xs')) []) Cs | 
| 33471 | 127 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 128 |     fun mk_case (SchemeCase {qs, oqnames, gs, lhs, ...}) =
 | 
| 33471 | 129 | HOLogic.mk_Trueprop Pbool | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 130 | |> fold_rev (fn x_l => curry Logic.mk_implies (HOLogic.mk_Trueprop(HOLogic.mk_eq x_l))) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 131 | (xs' ~~ lhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 132 | |> fold_rev (curry Logic.mk_implies) gs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 133 | |> fold_rev mk_forall_rename (oqnames ~~ map Free qs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 134 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 135 | HOLogic.mk_Trueprop Pbool | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 136 | |> fold_rev (curry Logic.mk_implies o mk_case) relevant_cases | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 137 | |> fold_rev (curry Logic.mk_implies) Cs' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 138 | |> fold_rev (Logic.all o Free) ws | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 139 | |> fold_rev mk_forall_rename (map fst xs ~~ xs') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 140 |     |> mk_forall_rename ("P", Pbool)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 141 | end | 
| 33471 | 142 | |
| 33855 
cd8acf137c9c
eliminated dead code and some unused bindings, reported by polyml
 krauss parents: 
33697diff
changeset | 143 | fun mk_wf R (IndScheme {T, ...}) =
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 144 |   HOLogic.Trueprop $ (Const (@{const_name wf}, mk_relT T --> HOLogic.boolT) $ R)
 | 
| 33471 | 145 | |
| 146 | fun mk_ineqs R (IndScheme {T, cases, branches}) =
 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 147 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 148 | fun inject i ts = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 149 | SumTree.mk_inj T (length branches) (i + 1) (foldr1 HOLogic.mk_prod ts) | 
| 33471 | 150 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 151 |     val thesis = Free ("thesis", HOLogic.boolT) (* FIXME *)
 | 
| 33471 | 152 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 153 | fun mk_pres bdx args = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 154 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 155 |         val SchemeBranch { xs, ws, Cs, ... } = nth branches bdx
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 156 | fun replace (x, v) t = betapply (lambda (Free x) t, v) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 157 | val Cs' = map (fold replace (xs ~~ args)) Cs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 158 | val cse = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 159 | HOLogic.mk_Trueprop thesis | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 160 | |> fold_rev (curry Logic.mk_implies) Cs' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 161 | |> fold_rev (Logic.all o Free) ws | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 162 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 163 | Logic.mk_implies (cse, HOLogic.mk_Trueprop thesis) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 164 | end | 
| 33471 | 165 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 166 |     fun f (SchemeCase {bidx, qs, oqnames, gs, lhs, rs, ...}) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 167 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 168 | fun g (bidx', Gvs, Gas, rcarg) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 169 | let val export = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 170 | fold_rev (curry Logic.mk_implies) Gas | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 171 | #> fold_rev (curry Logic.mk_implies) gs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 172 | #> fold_rev (Logic.all o Free) Gvs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 173 | #> fold_rev mk_forall_rename (oqnames ~~ map Free qs) | 
| 33471 | 174 | in | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 175 | (HOLogic.mk_mem (HOLogic.mk_prod (inject bidx' rcarg, inject bidx lhs), R) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 176 | |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 177 | |> export, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 178 | mk_pres bidx' rcarg | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 179 | |> export | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 180 | |> Logic.all thesis) | 
| 33471 | 181 | end | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 182 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 183 | map g rs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 184 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 185 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 186 | map f cases | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 187 | end | 
| 33471 | 188 | |
| 189 | ||
| 190 | fun mk_ind_goal thy branches = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 191 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 192 |     fun brnch (SchemeBranch { P, xs, ws, Cs, ... }) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 193 | HOLogic.mk_Trueprop (list_comb (P, map Free xs)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 194 | |> fold_rev (curry Logic.mk_implies) Cs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 195 | |> fold_rev (Logic.all o Free) ws | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 196 | |> term_conv thy ind_atomize | 
| 35625 | 197 | |> Object_Logic.drop_judgment thy | 
| 39756 
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
 krauss parents: 
37677diff
changeset | 198 | |> HOLogic.tupled_lambda (foldr1 HOLogic.mk_prod (map Free xs)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 199 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 200 | SumTree.mk_sumcases HOLogic.boolT (map brnch branches) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 201 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 202 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 203 | fun mk_induct_rule ctxt R x complete_thms wf_thm ineqss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 204 |   (IndScheme {T, cases=scases, branches}) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 205 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 206 | val n = length branches | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 207 | val scases_idx = map_index I scases | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 208 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 209 | fun inject i ts = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 210 | SumTree.mk_inj T n (i + 1) (foldr1 HOLogic.mk_prod ts) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 211 |     val P_of = nth (map (fn (SchemeBranch { P, ... }) => P) branches)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 212 | |
| 42361 | 213 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 214 | val cert = cterm_of thy | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 215 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 216 | val P_comp = mk_ind_goal thy branches | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 217 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 218 | (* Inductive Hypothesis: !!z. (z,x):R ==> P z *) | 
| 46217 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 wenzelm parents: 
42495diff
changeset | 219 |     val ihyp = Logic.all_const T $ Abs ("z", T,
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 220 | Logic.mk_implies | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 221 | (HOLogic.mk_Trueprop ( | 
| 37677 | 222 |           Const (@{const_name Set.member}, HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT) 
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 223 | $ (HOLogic.pair_const T T $ Bound 0 $ x) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 224 | $ R), | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 225 | HOLogic.mk_Trueprop (P_comp $ Bound 0))) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 226 | |> cert | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 227 | |
| 36945 | 228 | val aihyp = Thm.assume ihyp | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 229 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 230 | (* Rule for case splitting along the sum types *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 231 |     val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 232 | val pats = map_index (uncurry inject) xss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 233 | val sum_split_rule = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 234 | Pat_Completeness.prove_completeness thy [x] (P_comp $ x) xss (map single pats) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 235 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 236 |     fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 237 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 238 | val fxs = map Free xs | 
| 36945 | 239 | val branch_hyp = Thm.assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat)))) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 240 | |
| 36945 | 241 | val C_hyps = map (cert #> Thm.assume) Cs | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 242 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 243 | val (relevant_cases, ineqss') = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 244 | (scases_idx ~~ ineqss) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 245 |           |> filter (fn ((_, SchemeCase {bidx=bidx', ...}), _) => bidx' = bidx)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 246 | |> split_list | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 247 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 248 |         fun prove_case (cidx, SchemeCase {qs, gs, lhs, rs, ...}) ineq_press =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 249 | let | 
| 36945 | 250 | val case_hyps = | 
| 251 | map (Thm.assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 252 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 253 | val cqs = map (cert o Free) qs | 
| 36945 | 254 | val ags = map (Thm.assume o cert) gs | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 255 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 256 | val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 257 | val sih = full_simplify replace_x_ss aihyp | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 258 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 259 | fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 260 | let | 
| 36945 | 261 | val cGas = map (Thm.assume o cert) Gas | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 262 | val cGvs = map (cert o Free) Gvs | 
| 36945 | 263 | val import = fold Thm.forall_elim (cqs @ cGvs) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 264 | #> fold Thm.elim_implies (ags @ cGas) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 265 | val ipres = pres | 
| 36945 | 266 | |> Thm.forall_elim (cert (list_comb (P_of idx, rcargs))) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 267 | |> import | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 268 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 269 | sih | 
| 36945 | 270 | |> Thm.forall_elim (cert (inject idx rcargs)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 271 | |> Thm.elim_implies (import ineq) (* Psum rcargs *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 272 | |> Conv.fconv_rule sum_prod_conv | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 273 | |> Conv.fconv_rule ind_rulify | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 274 | |> (fn th => th COMP ipres) (* P rs *) | 
| 36945 | 275 | |> fold_rev (Thm.implies_intr o cprop_of) cGas | 
| 276 | |> fold_rev Thm.forall_intr cGvs | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 277 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 278 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 279 | val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 280 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 281 | val step = HOLogic.mk_Trueprop (list_comb (P, lhs)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 282 | |> fold_rev (curry Logic.mk_implies o prop_of) P_recs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 283 | |> fold_rev (curry Logic.mk_implies) gs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 284 | |> fold_rev (Logic.all o Free) qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 285 | |> cert | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 286 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 287 | val Plhs_to_Pxs_conv = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 288 | foldl1 (uncurry Conv.combination_conv) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 289 | (Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 290 | |
| 36945 | 291 | val res = Thm.assume step | 
| 292 | |> fold Thm.forall_elim cqs | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 293 | |> fold Thm.elim_implies ags | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 294 | |> fold Thm.elim_implies P_recs (* P lhs *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 295 | |> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *) | 
| 36945 | 296 | |> fold_rev (Thm.implies_intr o cprop_of) (ags @ case_hyps) | 
| 297 | |> fold_rev Thm.forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 298 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 299 | (res, (cidx, step)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 300 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 301 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 302 | val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 303 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 304 | val bstep = complete_thm | 
| 36945 | 305 | |> Thm.forall_elim (cert (list_comb (P, fxs))) | 
| 306 | |> fold (Thm.forall_elim o cert) (fxs @ map Free ws) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 307 | |> fold Thm.elim_implies C_hyps | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 308 | |> fold Thm.elim_implies cases (* P xs *) | 
| 36945 | 309 | |> fold_rev (Thm.implies_intr o cprop_of) C_hyps | 
| 310 | |> fold_rev (Thm.forall_intr o cert o Free) ws | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 311 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 312 | val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 313 | |> Goal.init | 
| 41225 | 314 |           |> (Simplifier.rewrite_goals_tac (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.cases}))
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 315 | THEN CONVERSION ind_rulify 1) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 316 | |> Seq.hd | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 317 | |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 318 | |> Goal.finish ctxt | 
| 36945 | 319 | |> Thm.implies_intr (cprop_of branch_hyp) | 
| 320 | |> fold_rev (Thm.forall_intr o cert) fxs | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 321 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 322 | (Pxs, steps) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 323 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 324 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 325 | val (branches, steps) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 326 | map_index prove_branch (branches ~~ (complete_thms ~~ pats)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 327 | |> split_list |> apsnd flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 328 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 329 | val istep = sum_split_rule | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 330 | |> fold (fn b => fn th => Drule.compose_single (b, 1, th)) branches | 
| 36945 | 331 | |> Thm.implies_intr ihyp | 
| 332 | |> Thm.forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 333 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 334 | val induct_rule = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 335 |       @{thm "wf_induct_rule"}
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 336 | |> (curry op COMP) wf_thm | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 337 | |> (curry op COMP) istep | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 338 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 339 | val steps_sorted = map snd (sort (int_ord o pairself fst) steps) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 340 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 341 | (steps_sorted, induct_rule) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 342 | end | 
| 33471 | 343 | |
| 344 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 345 | fun mk_ind_tac comp_tac pres_tac term_tac ctxt facts = | 
| 46467 | 346 | (* FIXME proper use of facts!? *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 347 | (ALLGOALS (Method.insert_tac facts)) THEN HEADGOAL (SUBGOAL (fn (t, i) => | 
| 33471 | 348 | let | 
| 349 | val (ctxt', _, cases, concl) = dest_hhf ctxt t | |
| 350 |     val scheme as IndScheme {T=ST, branches, ...} = mk_scheme' ctxt' cases concl
 | |
| 351 | val ([Rn,xn], ctxt'') = Variable.variant_fixes ["R","x"] ctxt' | |
| 352 | val R = Free (Rn, mk_relT ST) | |
| 353 | val x = Free (xn, ST) | |
| 42361 | 354 | val cert = cterm_of (Proof_Context.theory_of ctxt) | 
| 33471 | 355 | |
| 356 | val ineqss = mk_ineqs R scheme | |
| 36945 | 357 | |> map (map (pairself (Thm.assume o cert))) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 358 | val complete = | 
| 36945 | 359 | map_range (mk_completeness ctxt scheme #> cert #> Thm.assume) (length branches) | 
| 360 | val wf_thm = mk_wf R scheme |> cert |> Thm.assume | |
| 33471 | 361 | |
| 362 | val (descent, pres) = split_list (flat ineqss) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 363 | val newgoals = complete @ pres @ wf_thm :: descent | 
| 33471 | 364 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 365 | val (steps, indthm) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 366 | mk_induct_rule ctxt'' R x complete wf_thm ineqss scheme | 
| 33471 | 367 | |
| 368 |     fun project (i, SchemeBranch {xs, ...}) =
 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 369 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 370 | val inst = (foldr1 HOLogic.mk_prod (map Free xs)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 371 | |> SumTree.mk_inj ST (length branches) (i + 1) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 372 | |> cert | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 373 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 374 | indthm | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 375 | |> Drule.instantiate' [] [SOME inst] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 376 | |> simplify SumTree.sumcase_split_ss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 377 | |> Conv.fconv_rule ind_rulify | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 378 | end | 
| 33471 | 379 | |
| 380 | val res = Conjunction.intr_balanced (map_index project branches) | |
| 36945 | 381 | |> fold_rev Thm.implies_intr (map cprop_of newgoals @ steps) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 382 | |> Drule.generalize ([], [Rn]) | 
| 33471 | 383 | |
| 384 | val nbranches = length branches | |
| 385 | val npres = length pres | |
| 386 | in | |
| 387 | Thm.compose_no_flatten false (res, length newgoals) i | |
| 388 | THEN term_tac (i + nbranches + npres) | |
| 389 | THEN (EVERY (map (TRY o pres_tac) ((i + nbranches + npres - 1) downto (i + nbranches)))) | |
| 390 | THEN (EVERY (map (TRY o comp_tac) ((i + nbranches - 1) downto i))) | |
| 391 | end)) | |
| 392 | ||
| 393 | ||
| 394 | fun induction_schema_tac ctxt = | |
| 395 | mk_ind_tac (K all_tac) (assume_tac APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt; | |
| 396 | ||
| 397 | end |