| author | wenzelm | 
| Thu, 22 Dec 2005 00:29:01 +0100 | |
| changeset 18471 | ca9a864018d6 | 
| parent 18465 | 16dcd36499b8 | 
| child 18496 | ef36f9be255e | 
| permissions | -rw-r--r-- | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 1 | (* Title: Provers/induct_method.ML | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 4 | |
| 11735 | 5 | Proof by cases and induction on sets and types. | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 6 | *) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 7 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 8 | signature INDUCT_METHOD_DATA = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 9 | sig | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 10 | val cases_default: thm | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 11 | val atomize_old: thm list | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 12 | val atomize: thm list | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 13 | val rulify: thm list | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 14 | val rulify_fallback: thm list | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 15 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 16 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 17 | signature INDUCT_METHOD = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 18 | sig | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 19 | val fix_tac: Proof.context -> int -> (string * typ) list -> int -> tactic | 
| 18287 | 20 | val add_defs: (string option * term) option list -> Proof.context -> | 
| 21 | (term option list * thm list) * Proof.context | |
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 22 | val atomize_term: theory -> term -> term | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 23 | val atomize_tac: int -> tactic | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 24 | val rulified_term: thm -> theory * term | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 25 | val rulify_tac: int -> tactic | 
| 18259 | 26 | val guess_instance: thm -> int -> thm -> thm Seq.seq | 
| 16391 | 27 | val cases_tac: Proof.context -> bool -> term option list list -> thm option -> | 
| 18224 | 28 | thm list -> int -> cases_tactic | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 29 | val induct_tac: Proof.context -> bool -> (string option * term) option list list -> | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 30 | (string * typ) list list -> term option list -> thm list option -> thm list -> int -> | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 31 | cases_tactic | 
| 18235 | 32 | val coinduct_tac: Proof.context -> bool -> term option list -> term option list -> | 
| 33 | thm option -> thm list -> int -> cases_tactic | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 34 | val setup: (theory -> theory) list | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 35 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 36 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 37 | functor InductMethodFun(Data: INDUCT_METHOD_DATA): INDUCT_METHOD = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 38 | struct | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 39 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 40 | val meta_spec = thm "Pure.meta_spec"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 41 | val all_conjunction = thm "Pure.all_conjunction"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 42 | val imp_conjunction = thm "Pure.imp_conjunction"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 43 | val conjunction_imp = thm "Pure.conjunction_imp"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 44 | val conjunction_assoc = thm "Pure.conjunction_assoc"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 45 | val conjunction_congs = [all_conjunction, imp_conjunction]; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 46 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 47 | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 48 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 49 | (** misc utils **) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 50 | |
| 18287 | 51 | (* alignment *) | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 52 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 53 | fun align_left msg xs ys = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 54 | let val m = length xs and n = length ys | 
| 11735 | 55 | in if m < n then raise ERROR_MESSAGE msg else (Library.take (n, xs) ~~ ys) end; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 56 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 57 | fun align_right msg xs ys = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 58 | let val m = length xs and n = length ys | 
| 11735 | 59 | in if m < n then raise ERROR_MESSAGE msg else (Library.drop (m - n, xs) ~~ ys) end; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 60 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 61 | |
| 11735 | 62 | (* prep_inst *) | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 63 | |
| 18205 | 64 | fun prep_inst thy align tune (tm, ts) = | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 65 | let | 
| 18205 | 66 | val cert = Thm.cterm_of thy; | 
| 15531 | 67 | fun prep_var (x, SOME t) = | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 68 | let | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 69 | val cx = cert x; | 
| 18147 | 70 |             val {T = xT, thy, ...} = Thm.rep_cterm cx;
 | 
| 12799 
5472afdd3bd3
MetaSimplifier.rewrite_term replaces slow Tactic.rewrite_cterm;
 wenzelm parents: 
12305diff
changeset | 71 | val ct = cert (tune t); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 72 | in | 
| 18147 | 73 | if Sign.typ_instance thy (#T (Thm.rep_cterm ct), xT) then SOME (cx, ct) | 
| 11735 | 74 | else raise ERROR_MESSAGE (Pretty.string_of (Pretty.block | 
| 75 | [Pretty.str "Ill-typed instantiation:", Pretty.fbrk, | |
| 76 | Display.pretty_cterm ct, Pretty.str " ::", Pretty.brk 1, | |
| 77 | Display.pretty_ctyp (#T (Thm.crep_cterm ct))])) | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 78 | end | 
| 15531 | 79 | | prep_var (_, NONE) = NONE; | 
| 11735 | 80 | val xs = InductAttrib.vars_of tm; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 81 | in | 
| 11735 | 82 | align "Rule has fewer variables than instantiations given" xs ts | 
| 15570 | 83 | |> List.mapPartial prep_var | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 84 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 85 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 86 | |
| 18205 | 87 | (* trace_rules *) | 
| 88 | ||
| 89 | fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule")
 | |
| 90 | | trace_rules ctxt _ rules = Method.trace ctxt rules; | |
| 91 | ||
| 92 | ||
| 93 | (* make_cases *) | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 94 | |
| 18224 | 95 | fun make_cases is_open rule = | 
| 96 | RuleCases.make is_open NONE (Thm.theory_of_thm rule, Thm.prop_of rule); | |
| 97 | ||
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 98 | fun warn_open true = warning "Encountered open rule cases -- deprecated" | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 99 | | warn_open false = (); | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 100 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 101 | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 102 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 103 | (** cases method **) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 104 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 105 | (* | 
| 11735 | 106 | rule selection scheme: | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 107 | cases - default case split | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 108 | `x:A` cases ... - set cases | 
| 11735 | 109 | cases t - type cases | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 110 | ... cases ... r - explicit rule | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 111 | *) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 112 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 113 | local | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 114 | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 115 | fun find_casesT ctxt ((SOME t :: _) :: _) = InductAttrib.find_casesT ctxt (Term.fastype_of t) | 
| 11735 | 116 | | find_casesT _ _ = []; | 
| 117 | ||
| 18224 | 118 | fun find_casesS ctxt (fact :: _) = InductAttrib.find_casesS ctxt (Thm.concl_of fact) | 
| 11735 | 119 | | find_casesS _ _ = []; | 
| 120 | ||
| 16391 | 121 | in | 
| 122 | ||
| 123 | fun cases_tac ctxt is_open insts opt_rule facts = | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 124 | let | 
| 18224 | 125 | val _ = warn_open is_open; | 
| 18147 | 126 | val thy = ProofContext.theory_of ctxt; | 
| 127 | val cert = Thm.cterm_of thy; | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 128 | |
| 11735 | 129 | fun inst_rule r = | 
| 18224 | 130 | if null insts then `RuleCases.get r | 
| 11735 | 131 | else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts | 
| 18205 | 132 | |> (List.concat o map (prep_inst thy align_left I)) | 
| 18224 | 133 | |> Drule.cterm_instantiate) r |> pair (RuleCases.get r); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 134 | |
| 11735 | 135 | val ruleq = | 
| 12852 | 136 | (case opt_rule of | 
| 18205 | 137 | SOME r => Seq.single (inst_rule r) | 
| 138 | | NONE => | |
| 139 | (find_casesS ctxt facts @ find_casesT ctxt insts @ [Data.cases_default]) | |
| 140 | |> tap (trace_rules ctxt InductAttrib.casesN) | |
| 18224 | 141 | |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 142 | in | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 143 | fn i => fn st => | 
| 18224 | 144 | ruleq | 
| 18235 | 145 | |> Seq.maps (RuleCases.consume [] facts) | 
| 18224 | 146 | |> Seq.maps (fn ((cases, (_, more_facts)), rule) => | 
| 147 | CASES (make_cases is_open rule cases) | |
| 148 | (Method.insert_tac more_facts i THEN Tactic.rtac rule i) st) | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 149 | end; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 150 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 151 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 152 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 153 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 154 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 155 | (** induct method **) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 156 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 157 | (* atomize *) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 158 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 159 | fun common_atomize_term is_old thy = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 160 | MetaSimplifier.rewrite_term thy (if is_old then Data.atomize_old else Data.atomize) [] | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 161 | #> ObjectLogic.drop_judgment thy; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 162 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 163 | val atomize_term = common_atomize_term false; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 164 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 165 | fun common_atomize_cterm is_old = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 166 | Tactic.rewrite true (if is_old then Data.atomize_old else Data.atomize); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 167 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 168 | fun common_atomize_tac is_old inner = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 169 | if is_old then | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 170 | Tactic.rewrite_goal_tac [conjunction_assoc] | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 171 | THEN' Tactic.rewrite_goal_tac Data.atomize_old | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 172 | else | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 173 | (if inner then Tactic.rewrite_goal_tac (map Thm.symmetric conjunction_congs) else K all_tac) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 174 | THEN' Tactic.rewrite_goal_tac Data.atomize; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 175 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 176 | val atomize_tac = common_atomize_tac false false; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 177 | val inner_atomize_tac = common_atomize_tac false true; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 178 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 179 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 180 | (* rulify *) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 181 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 182 | fun rulify_term thy = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 183 | MetaSimplifier.rewrite_term thy (Data.rulify @ conjunction_congs) [] #> | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 184 | MetaSimplifier.rewrite_term thy Data.rulify_fallback []; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 185 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 186 | fun rulified_term thm = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 187 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 188 | val thy = Thm.theory_of_thm thm; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 189 | val rulify = rulify_term thy; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 190 | val (As, B) = Logic.strip_horn (Thm.prop_of thm); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 191 | in (thy, Logic.list_implies (map rulify As, rulify B)) end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 192 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 193 | val curry_prems_tac = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 194 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 195 | val rule = conjunction_imp; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 196 | val thy = Thm.theory_of_thm rule; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 197 | val proc = Simplifier.simproc_i thy "curry_prems" | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 198 | [#1 (Logic.dest_equals (Thm.prop_of rule))] | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 199 | (fn _ => fn ss => fn _ => (*WORKAROUND: prevent descending into meta conjunctions*) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 200 | if exists (equal propT o #2 o #1) (#2 (#bounds (#1 (Simplifier.rep_ss ss)))) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 201 | then NONE else SOME rule); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 202 | val ss = MetaSimplifier.theory_context thy Simplifier.empty_ss addsimprocs [proc]; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 203 | in Simplifier.full_simp_tac ss end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 204 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 205 | val rulify_tac = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 206 | Tactic.rewrite_goal_tac (Data.rulify @ conjunction_congs) THEN' | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 207 | Tactic.rewrite_goal_tac Data.rulify_fallback THEN' | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 208 | Tactic.conjunction_tac THEN_ALL_NEW | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 209 | (curry_prems_tac THEN' Tactic.norm_hhf_tac); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 210 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 211 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 212 | (* prepare rule *) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 213 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 214 | fun rule_instance thy inst rule = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 215 | Drule.cterm_instantiate (prep_inst thy align_left I (Thm.prop_of rule, inst)) rule; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 216 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 217 | fun mutual_rule ths = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 218 | (case RuleCases.mutual_rule ths of | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 219 | NONE => error "Failed to join given rules into one mutual rule" | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 220 | | SOME res => res); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 221 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 222 | fun internalize is_old k th = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 223 | th |> Thm.permute_prems 0 k | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 224 | |> Drule.fconv_rule (Drule.concl_conv (Thm.nprems_of th - k) (common_atomize_cterm is_old)); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 225 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 226 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 227 | (* guess rule instantiation -- cannot handle pending goal parameters *) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 228 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 229 | local | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 230 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 231 | fun dest_env thy (env as Envir.Envir {iTs, ...}) =
 | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 232 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 233 | val cert = Thm.cterm_of thy; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 234 | val certT = Thm.ctyp_of thy; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 235 | val pairs = Envir.alist_of env; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 236 | val ts = map (cert o Envir.norm_term env o #2 o #2) pairs; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 237 | val xs = map2 (curry (cert o Var)) (map #1 pairs) (map (#T o Thm.rep_cterm) ts); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 238 | in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) (Vartab.dest iTs), xs ~~ ts) end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 239 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 240 | in | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 241 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 242 | fun guess_instance rule i st = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 243 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 244 |     val {thy, maxidx, ...} = Thm.rep_thm st;
 | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 245 | val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 246 | val params = rev (rename_wrt_term goal (Logic.strip_params goal)); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 247 | in | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 248 | if not (null params) then | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 249 |       (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
 | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 250 | commas_quote (map (Sign.string_of_term thy o Syntax.mark_boundT) params)); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 251 | Seq.single rule) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 252 | else | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 253 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 254 | val rule' = Thm.incr_indexes (maxidx + 1) rule; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 255 | val concl = Logic.strip_assums_concl goal; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 256 | in | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 257 | Unify.smash_unifiers (thy, Envir.empty (#maxidx (Thm.rep_thm rule')), | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 258 | [(Thm.concl_of rule', concl)]) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 259 | |> Seq.map (fn env => Drule.instantiate (dest_env thy env) rule') | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 260 | end | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 261 | end handle Subscript => Seq.empty; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 262 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 263 | end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 264 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 265 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 266 | (* special renaming of rule parameters *) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 267 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 268 | fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 269 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 270 | val x = ProofContext.revert_skolem ctxt z; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 271 | fun index i [] = [] | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 272 | | index i (y :: ys) = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 273 | if x = y then x ^ string_of_int i :: index (i + 1) ys | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 274 | else y :: index i ys; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 275 | fun rename_params [] = [] | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 276 | | rename_params ((y, Type (U, _)) :: ys) = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 277 | (if U = T then x else y) :: rename_params ys | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 278 | | rename_params ((y, _) :: ys) = y :: rename_params ys; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 279 | fun rename_asm A = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 280 | let | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 281 | val xs = rename_params (Logic.strip_params A); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 282 | val xs' = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 283 | (case List.filter (equal x) xs of | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 284 | [] => xs | [_] => xs | _ => index 1 xs); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 285 | in Logic.list_rename_params (xs', A) end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 286 | fun rename_prop p = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 287 | let val (As, C) = Logic.strip_horn p | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 288 | in Logic.list_implies (map rename_asm As, C) end; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 289 | val cp' = cterm_fun rename_prop (Thm.cprop_of thm); | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 290 | val thm' = Thm.equal_elim (Thm.reflexive cp') thm; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 291 | in [RuleCases.save thm thm'] end | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 292 | | special_rename_params _ _ ths = ths; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 293 | |
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 294 | |
| 18235 | 295 | (* fix_tac *) | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 296 | |
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 297 | local | 
| 18240 | 298 | |
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 299 | fun goal_prefix k ((c as Const ("all", _)) $ Abs (a, T, B)) = c $ Abs (a, T, goal_prefix k B)
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 300 | | goal_prefix 0 _ = Term.dummy_pattern propT | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 301 |   | goal_prefix k ((c as Const ("==>", _)) $ A $ B) = c $ A $ goal_prefix (k - 1) B
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 302 | | goal_prefix _ _ = Term.dummy_pattern propT; | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 303 | |
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 304 | fun goal_params k (Const ("all", _) $ Abs (_, _, B)) = goal_params k B + 1
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 305 | | goal_params 0 _ = 0 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 306 |   | goal_params k (Const ("==>", _) $ _ $ B) = goal_params (k - 1) B
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 307 | | goal_params _ _ = 0; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 308 | |
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 309 | fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) => | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 310 | let | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 311 | val thy = ProofContext.theory_of ctxt; | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 312 | val cert = Thm.cterm_of thy; | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 313 | val certT = Thm.ctyp_of thy; | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 314 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 315 | val v = Free (x, T); | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 316 | fun spec_rule prfx (xs, body) = | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 317 | meta_spec | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 318 | |> Thm.rename_params_rule ([ProofContext.revert_skolem ctxt x], 1) | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 319 | |> Thm.lift_rule (cert prfx) | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 320 | |> `(Thm.prop_of #> Logic.strip_assums_concl) | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 321 | |-> (fn pred $ arg => | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 322 | Drule.cterm_instantiate | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 323 | [(cert (Term.head_of pred), cert (Unify.rlist_abs (xs, body))), | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 324 | (cert (Term.head_of arg), cert (Unify.rlist_abs (xs, v)))]); | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 325 | |
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 326 |     fun goal_concl k xs (Const ("all", _) $ Abs (a, T, B)) = goal_concl k ((a, T) :: xs) B
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 327 | | goal_concl 0 xs B = | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 328 | if not (Term.exists_subterm (fn t => t aconv v) B) then NONE | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 329 | else SOME (xs, Term.absfree (x, T, Term.incr_boundvars 1 B)) | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 330 |       | goal_concl k xs (Const ("==>", _) $ _ $ B) = goal_concl (k - 1) xs B
 | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 331 | | goal_concl _ _ _ = NONE; | 
| 18205 | 332 | in | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 333 | (case goal_concl n [] goal of | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 334 | SOME concl => | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 335 | (compose_tac (false, spec_rule (goal_prefix n goal) concl, 1) THEN' rtac asm_rl) i | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 336 |     | NONE => (warning ("Induction: no variable " ^ quote (ProofContext.string_of_term ctxt v) ^
 | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 337 | " to be fixed -- ignored"); all_tac)) | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 338 | end); | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 339 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 340 | fun miniscope_tac p i = PRIMITIVE (Drule.fconv_rule | 
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 341 | (Drule.goals_conv (Library.equal i) | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 342 | (Drule.forall_conv p (Tactic.rewrite true [Thm.symmetric Drule.norm_hhf_eq])))); | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 343 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 344 | in | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 345 | |
| 18250 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 346 | fun fix_tac _ _ [] = K all_tac | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 347 | | fix_tac ctxt n xs = SUBGOAL (fn (goal, i) => | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 348 | (EVERY' (map (meta_spec_tac ctxt n) xs) THEN' | 
| 
dfe5d09514eb
fix_tac: proper treatment of major premises in goal;
 wenzelm parents: 
18240diff
changeset | 349 | (miniscope_tac (goal_params n goal))) i); | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 350 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 351 | end; | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 352 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 353 | |
| 18235 | 354 | (* add_defs *) | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 355 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 356 | fun add_defs def_insts = | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 357 | let | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 358 | fun add (SOME (SOME x, t)) ctxt = | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 359 | let val ((lhs, def), ctxt') = ProofContext.add_def (x, t) ctxt | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 360 | in ((SOME (Free lhs), [def]), ctxt') end | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 361 | | add (SOME (NONE, t)) ctxt = ((SOME t, []), ctxt) | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 362 | | add NONE ctxt = ((NONE, []), ctxt); | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 363 | in fold_map add def_insts #> apfst (split_list #> apsnd List.concat) end; | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 364 | |
| 11790 
42393a11642d
simplified resolveq_cases_tac for cases, separate version for induct;
 wenzelm parents: 
11781diff
changeset | 365 | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 366 | (* induct_tac *) | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 367 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 368 | (* | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 369 | rule selection scheme: | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 370 | `x:A` induct ... - set induction | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 371 | induct x - type induction | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 372 | ... induct ... r - explicit rule | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 373 | *) | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 374 | |
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 375 | local | 
| 15235 
614a804d7116
Induction now preserves the name of the induction variable.
 nipkow parents: 
14981diff
changeset | 376 | |
| 11735 | 377 | fun find_inductT ctxt insts = | 
| 18147 | 378 | fold_rev multiply (insts |> List.mapPartial (fn [] => NONE | ts => List.last ts) | 
| 18205 | 379 | |> map (InductAttrib.find_inductT ctxt o Term.fastype_of)) [[]] | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 380 | |> filter_out (forall Drule.is_internal); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 381 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 382 | fun find_inductS ctxt (fact :: _) = map single (InductAttrib.find_inductS ctxt (Thm.concl_of fact)) | 
| 11735 | 383 | | find_inductS _ _ = []; | 
| 384 | ||
| 16391 | 385 | in | 
| 11790 
42393a11642d
simplified resolveq_cases_tac for cases, separate version for induct;
 wenzelm parents: 
11781diff
changeset | 386 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 387 | fun common_induct_tac is_old ctxt is_open def_insts fixing taking opt_rule facts = | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 388 | let | 
| 18224 | 389 | val _ = warn_open is_open; | 
| 18147 | 390 | val thy = ProofContext.theory_of ctxt; | 
| 391 | val cert = Thm.cterm_of thy; | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 392 | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 393 | val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list; | 
| 18259 | 394 | val atomized_defs = map (map ObjectLogic.atomize_thm) defs; | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 395 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 396 | fun inst_rule (concls, r) = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 397 | (if null insts then `RuleCases.get r | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 398 | else (align_right "Rule has fewer conclusions than arguments given" | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 399 | (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 400 | |> (List.concat o map (prep_inst thy align_right (common_atomize_term is_old thy))) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 401 | |> Drule.cterm_instantiate) r |> pair (RuleCases.get r)) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 402 | |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th)); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 403 | |
| 11735 | 404 | val ruleq = | 
| 405 | (case opt_rule of | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 406 | SOME rs => Seq.single (inst_rule (mutual_rule rs)) | 
| 18205 | 407 | | NONE => | 
| 408 | (find_inductS ctxt facts @ | |
| 409 | map (special_rename_params defs_ctxt insts) (find_inductT ctxt insts)) | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 410 | |> List.mapPartial RuleCases.mutual_rule | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 411 | |> tap (trace_rules ctxt InductAttrib.inductN o map #2) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 412 | |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 413 | |
| 18224 | 414 | fun rule_cases rule = RuleCases.make is_open (SOME (Thm.prop_of rule)) (rulified_term rule); | 
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 415 | in | 
| 18205 | 416 | (fn i => fn st => | 
| 18224 | 417 | ruleq | 
| 18235 | 418 | |> Seq.maps (RuleCases.consume (List.concat defs) facts) | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 419 | |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 420 | (CONJUNCTS (length concls) (ALLGOALS (fn j => | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 421 | (CONJUNCTS ~1 (ALLGOALS | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 422 | (Method.insert_tac (more_facts @ nth_list atomized_defs (j - 1)) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 423 | THEN' fix_tac defs_ctxt | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 424 | (nth concls (j - 1) + more_consumes) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 425 | (nth_list fixing (j - 1)))) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 426 | THEN' common_atomize_tac is_old true) j)) | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 427 | THEN' common_atomize_tac is_old false) i st |> Seq.maps (fn st' => | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 428 | guess_instance (internalize is_old more_consumes rule) i st' | 
| 18235 | 429 | |> Seq.map (rule_instance thy taking) | 
| 430 | |> Seq.maps (fn rule' => | |
| 18224 | 431 | CASES (rule_cases rule' cases) | 
| 432 | (Tactic.rtac rule' i THEN | |
| 433 | PRIMSEQ (ProofContext.exports defs_ctxt ctxt)) st')))) | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 434 | THEN_ALL_NEW_CASES rulify_tac | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 435 | end; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 436 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 437 | val induct_tac = common_induct_tac false; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 438 | |
| 18205 | 439 | end; | 
| 440 | ||
| 441 | ||
| 442 | ||
| 443 | (** coinduct method **) | |
| 444 | ||
| 445 | (* | |
| 446 | rule selection scheme: | |
| 18224 | 447 | goal "x:A" coinduct ... - set coinduction | 
| 448 | coinduct x - type coinduction | |
| 449 | coinduct ... r - explicit rule | |
| 18205 | 450 | *) | 
| 451 | ||
| 452 | local | |
| 453 | ||
| 454 | fun find_coinductT ctxt (SOME t :: _) = InductAttrib.find_coinductT ctxt (Term.fastype_of t) | |
| 455 | | find_coinductT _ _ = []; | |
| 456 | ||
| 18224 | 457 | fun find_coinductS ctxt goal = InductAttrib.find_coinductS ctxt (Logic.strip_assums_concl goal); | 
| 18205 | 458 | |
| 459 | in | |
| 460 | ||
| 18235 | 461 | fun coinduct_tac ctxt is_open inst taking opt_rule facts = | 
| 18205 | 462 | let | 
| 18224 | 463 | val _ = warn_open is_open; | 
| 18205 | 464 | val thy = ProofContext.theory_of ctxt; | 
| 465 | val cert = Thm.cterm_of thy; | |
| 466 | ||
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 467 | fun inst_rule r = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 468 | if null inst then `RuleCases.get r | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 469 | else Drule.cterm_instantiate (prep_inst thy align_left I (Thm.concl_of r, inst)) r | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 470 | |> pair (RuleCases.get r); | 
| 18205 | 471 | |
| 18224 | 472 | fun ruleq goal = | 
| 18205 | 473 | (case opt_rule of | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 474 | SOME r => Seq.single (inst_rule r) | 
| 18205 | 475 | | NONE => | 
| 18224 | 476 | (find_coinductS ctxt goal @ find_coinductT ctxt inst) | 
| 18205 | 477 | |> tap (trace_rules ctxt InductAttrib.coinductN) | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 478 | |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); | 
| 18205 | 479 | in | 
| 18224 | 480 | SUBGOAL_CASES (fn (goal, i) => fn st => | 
| 481 | ruleq goal | |
| 18235 | 482 | |> Seq.maps (RuleCases.consume [] facts) | 
| 18224 | 483 | |> Seq.maps (fn ((cases, (_, more_facts)), rule) => | 
| 18259 | 484 | guess_instance rule i st | 
| 18235 | 485 | |> Seq.map (rule_instance thy taking) | 
| 486 | |> Seq.maps (fn rule' => | |
| 18224 | 487 | CASES (make_cases is_open rule' cases) | 
| 488 | (Method.insert_tac more_facts i THEN Tactic.rtac rule' i) st))) | |
| 18205 | 489 | end; | 
| 490 | ||
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 491 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 492 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 493 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 494 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 495 | (** concrete syntax **) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 496 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 497 | val openN = "open"; | 
| 18205 | 498 | val fixingN = "fixing"; | 
| 18235 | 499 | val takingN = "taking"; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 500 | val ruleN = "rule"; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 501 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 502 | local | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 503 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 504 | fun single_rule [rule] = rule | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 505 | | single_rule _ = error "Single rule expected"; | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 506 | |
| 15703 | 507 | fun named_rule k arg get = | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 508 | Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :-- | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 509 | (fn names => Scan.peek (fn ctxt => Scan.succeed (names |> map (fn name => | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 510 | (case get ctxt name of SOME x => x | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 511 |       | NONE => error ("No rule for " ^ k ^ " " ^ quote name)))))) >> #2;
 | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 512 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 513 | fun rule get_type get_set = | 
| 15703 | 514 | named_rule InductAttrib.typeN Args.local_tyname get_type || | 
| 515 | named_rule InductAttrib.setN Args.local_const get_set || | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 516 | Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.local_thmss; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 517 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 518 | val cases_rule = rule InductAttrib.lookup_casesT InductAttrib.lookup_casesS >> single_rule; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 519 | val induct_rule = rule InductAttrib.lookup_inductT InductAttrib.lookup_inductS; | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 520 | val coinduct_rule = | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 521 | rule InductAttrib.lookup_coinductT InductAttrib.lookup_coinductS >> single_rule; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 522 | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 523 | val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.local_term >> SOME; | 
| 18147 | 524 | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 525 | val def_inst = | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 526 | ((Scan.lift (Args.name --| (Args.$$$ "\\<equiv>" || Args.$$$ "==")) >> SOME) | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 527 | -- Args.local_term) >> SOME || | 
| 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 528 | inst >> Option.map (pair NONE); | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 529 | |
| 18147 | 530 | val free = Scan.state -- Args.local_term >> (fn (_, Free v) => v | (ctxt, t) => | 
| 531 |   error ("Bad free variable: " ^ ProofContext.string_of_term ctxt t));
 | |
| 532 | ||
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 533 | fun unless_more_args scan = Scan.unless (Scan.lift | 
| 18235 | 534 | ((Args.$$$ fixingN || Args.$$$ takingN || Args.$$$ InductAttrib.typeN || | 
| 535 | Args.$$$ InductAttrib.setN || Args.$$$ ruleN) -- Args.colon)) scan; | |
| 18205 | 536 | |
| 537 | val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |-- | |
| 538 | Args.and_list1 (Scan.repeat (unless_more_args free))) []; | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 539 | |
| 18235 | 540 | val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |-- | 
| 541 | Scan.repeat1 (unless_more_args inst)) []; | |
| 542 | ||
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 543 | in | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 544 | |
| 18235 | 545 | fun cases_meth src = | 
| 546 | Method.syntax (Args.mode openN -- | |
| 547 | (Args.and_list (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule)) src | |
| 548 | #> (fn (ctxt, (is_open, (insts, opt_rule))) => | |
| 549 | Method.METHOD_CASES (fn facts => | |
| 550 | Seq.DETERM (HEADGOAL (cases_tac ctxt is_open insts opt_rule facts)))); | |
| 18178 
9e4dfe031525
induct: support local definitions to be passed through the induction;
 wenzelm parents: 
18147diff
changeset | 551 | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 552 | fun induct_meth is_old src = | 
| 18235 | 553 | Method.syntax (Args.mode openN -- | 
| 554 | (Args.and_list (Scan.repeat (unless_more_args def_inst)) -- | |
| 555 | (fixing -- taking -- Scan.option induct_rule))) src | |
| 556 | #> (fn (ctxt, (is_open, (insts, ((fixing, taking), opt_rule)))) => | |
| 557 | Method.RAW_METHOD_CASES (fn facts => | |
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 558 | Seq.DETERM (HEADGOAL | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 559 | (common_induct_tac is_old ctxt is_open insts fixing taking opt_rule facts)))); | 
| 18205 | 560 | |
| 18235 | 561 | fun coinduct_meth src = | 
| 562 | Method.syntax (Args.mode openN -- | |
| 563 | (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule)) src | |
| 564 | #> (fn (ctxt, (is_open, ((insts, taking), opt_rule))) => | |
| 565 | Method.RAW_METHOD_CASES (fn facts => | |
| 566 | Seq.DETERM (HEADGOAL (coinduct_tac ctxt is_open insts taking opt_rule facts)))); | |
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 567 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 568 | end; | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 569 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 570 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 571 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 572 | (** theory setup **) | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 573 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 574 | val setup = | 
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 575 | [Method.add_methods | 
| 18235 | 576 | [(InductAttrib.casesN, cases_meth, "case analysis on types or sets"), | 
| 18465 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 577 | (InductAttrib.inductN, induct_meth false, "induction on types or sets"), | 
| 
16dcd36499b8
simplified setup: removed dest_concls, local_impI, conjI;
 wenzelm parents: 
18330diff
changeset | 578 |      ("old_" ^ InductAttrib.inductN, induct_meth true, "induction on types or sets (legacy)"),
 | 
| 18235 | 579 | (InductAttrib.coinductN, coinduct_meth, "coinduction on types or sets")]]; | 
| 11670 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 580 | |
| 
59f79df42d1f
proof by cases and induction on types and sets (used to be specific for HOL);
 wenzelm parents: diff
changeset | 581 | end; |