| author | haftmann | 
| Sun, 16 Oct 2016 09:31:05 +0200 | |
| changeset 64243 | aee949f6642d | 
| parent 63120 | 629a4c5e953e | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 6070 | 1 | (* Title: ZF/Tools/induct_tacs.ML | 
| 6065 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1994 University of Cambridge | |
| 4 | ||
| 12204 | 5 | Induction and exhaustion tactics for Isabelle/ZF. The theory | 
| 6 | information needed to support them (and to support primrec). Also a | |
| 7 | function to install other sets as if they were datatypes. | |
| 6065 | 8 | *) | 
| 9 | ||
| 10 | signature DATATYPE_TACTICS = | |
| 11 | sig | |
| 59788 | 12 | val exhaust_tac: Proof.context -> string -> (binding * string option * mixfix) list -> | 
| 13 | int -> tactic | |
| 14 | val induct_tac: Proof.context -> string -> (binding * string option * mixfix) list -> | |
| 15 | int -> tactic | |
| 12204 | 16 | val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
56231diff
changeset | 17 | val rep_datatype: Facts.ref * Token.src list -> Facts.ref * Token.src list -> | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
56231diff
changeset | 18 | (Facts.ref * Token.src list) list -> (Facts.ref * Token.src list) list -> theory -> theory | 
| 6065 | 19 | end; | 
| 20 | ||
| 21 | ||
| 6070 | 22 | (** Datatype information, e.g. associated theorems **) | 
| 23 | ||
| 24 | type datatype_info = | |
| 12175 | 25 |   {inductive: bool,             (*true if inductive, not coinductive*)
 | 
| 6070 | 26 | constructors : term list, (*the constructors, as Consts*) | 
| 27 | rec_rewrites : thm list, (*recursor equations*) | |
| 28 | case_rewrites : thm list, (*case equations*) | |
| 29 | induct : thm, | |
| 30 | mutual_induct : thm, | |
| 31 | exhaustion : thm}; | |
| 32 | ||
| 33522 | 33 | structure DatatypesData = Theory_Data | 
| 22846 | 34 | ( | 
| 6070 | 35 | type T = datatype_info Symtab.table; | 
| 36 | val empty = Symtab.empty; | |
| 16458 | 37 | val extend = I; | 
| 33522 | 38 | fun merge data : T = Symtab.merge (K true) data; | 
| 22846 | 39 | ); | 
| 6070 | 40 | |
| 41 | ||
| 42 | ||
| 43 | (** Constructor information: needed to map constructors to datatypes **) | |
| 44 | ||
| 45 | type constructor_info = | |
| 46 |   {big_rec_name : string,     (*name of the mutually recursive set*)
 | |
| 47 | constructors : term list, (*the constructors, as Consts*) | |
| 6141 | 48 | free_iffs : thm list, (*freeness simprules*) | 
| 6070 | 49 | rec_rewrites : thm list}; (*recursor equations*) | 
| 50 | ||
| 51 | ||
| 33522 | 52 | structure ConstructorsData = Theory_Data | 
| 22846 | 53 | ( | 
| 6070 | 54 | type T = constructor_info Symtab.table | 
| 55 | val empty = Symtab.empty | |
| 16458 | 56 | val extend = I | 
| 33522 | 57 | fun merge data = Symtab.merge (K true) data; | 
| 22846 | 58 | ); | 
| 6070 | 59 | |
| 6065 | 60 | structure DatatypeTactics : DATATYPE_TACTICS = | 
| 61 | struct | |
| 62 | ||
| 16458 | 63 | fun datatype_info thy name = | 
| 17412 | 64 | (case Symtab.lookup (DatatypesData.get thy) name of | 
| 15531 | 65 | SOME info => info | 
| 66 |   | NONE => error ("Unknown datatype " ^ quote name));
 | |
| 6065 | 67 | |
| 68 | ||
| 69 | (*Given a variable, find the inductive set associated it in the assumptions*) | |
| 14153 | 70 | exception Find_tname of string | 
| 71 | ||
| 37780 | 72 | fun find_tname ctxt var As = | 
| 24826 | 73 |   let fun mk_pair (Const(@{const_name mem},_) $ Free (v,_) $ A) =
 | 
| 6065 | 74 | (v, #1 (dest_Const (head_of A))) | 
| 12175 | 75 | | mk_pair _ = raise Match | 
| 37780 | 76 | val pairs = map_filter (try (mk_pair o FOLogic.dest_Trueprop)) As | 
| 77 | val x = | |
| 78 | (case try (dest_Free o Syntax.read_term ctxt) var of | |
| 79 | SOME (x, _) => x | |
| 80 |         | _ => raise Find_tname ("Bad variable " ^ quote var))
 | |
| 81 | in case AList.lookup (op =) pairs x of | |
| 15531 | 82 |        NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
 | 
| 83 | | SOME t => t | |
| 6065 | 84 | end; | 
| 85 | ||
| 12175 | 86 | (** generic exhaustion and induction tactic for datatypes | 
| 87 | Differences from HOL: | |
| 6065 | 88 | (1) no checking if the induction var occurs in premises, since it always | 
| 89 | appears in one of them, and it's hard to check for other occurrences | |
| 90 | (2) exhaustion works for VARIABLES in the premises, not general terms | |
| 91 | **) | |
| 92 | ||
| 59788 | 93 | fun exhaust_induct_tac exh ctxt var fixes i state = SUBGOAL (fn _ => | 
| 6065 | 94 | let | 
| 42361 | 95 | val thy = Proof_Context.theory_of ctxt | 
| 60695 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 wenzelm parents: 
59936diff
changeset | 96 |     val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i NONE state
 | 
| 59582 | 97 | val tn = find_tname ctxt' var (map Thm.term_of asms) | 
| 12175 | 98 | val rule = | 
| 16458 | 99 | if exh then #exhaustion (datatype_info thy tn) | 
| 100 | else #induct (datatype_info thy tn) | |
| 24826 | 101 |     val (Const(@{const_name mem},_) $ Var(ixn,_) $ _) =
 | 
| 59582 | 102 | (case Thm.prems_of rule of | 
| 12175 | 103 | [] => error "induction is not available for this datatype" | 
| 104 | | major::_ => FOLogic.dest_Trueprop major) | |
| 6065 | 105 | in | 
| 59788 | 106 | Rule_Insts.eres_inst_tac ctxt [((ixn, Position.none), var)] fixes rule i | 
| 14153 | 107 | end | 
| 108 | handle Find_tname msg => | |
| 109 | if exh then (*try boolean case analysis instead*) | |
| 59788 | 110 | case_tac ctxt var fixes i | 
| 56231 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 wenzelm parents: 
55740diff
changeset | 111 | else error msg) i state; | 
| 6065 | 112 | |
| 113 | val exhaust_tac = exhaust_induct_tac true; | |
| 114 | val induct_tac = exhaust_induct_tac false; | |
| 115 | ||
| 6070 | 116 | |
| 117 | (**** declare non-datatype as datatype ****) | |
| 118 | ||
| 119 | fun rep_datatype_i elim induct case_eqns recursor_eqns thy = | |
| 120 | let | |
| 121 | (*analyze the LHS of a case equation to get a constructor*) | |
| 41310 | 122 |     fun const_of (Const(@{const_name IFOL.eq}, _) $ (_ $ c) $ _) = c
 | 
| 6070 | 123 |       | const_of eqn = error ("Ill-formed case equation: " ^
 | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26618diff
changeset | 124 | Syntax.string_of_term_global thy eqn); | 
| 6070 | 125 | |
| 126 | val constructors = | |
| 44058 | 127 | map (head_of o const_of o FOLogic.dest_Trueprop o Thm.prop_of) case_eqns; | 
| 6070 | 128 | |
| 24826 | 129 |     val Const (@{const_name mem}, _) $ _ $ data =
 | 
| 59582 | 130 | FOLogic.dest_Trueprop (hd (Thm.prems_of elim)); | 
| 12175 | 131 | |
| 6112 | 132 | val Const(big_rec_name, _) = head_of data; | 
| 133 | ||
| 6070 | 134 | val simps = case_eqns @ recursor_eqns; | 
| 135 | ||
| 136 | val dt_info = | |
| 12175 | 137 |           {inductive = true,
 | 
| 138 | constructors = constructors, | |
| 139 | rec_rewrites = recursor_eqns, | |
| 140 | case_rewrites = case_eqns, | |
| 141 | induct = induct, | |
| 35409 | 142 |            mutual_induct = @{thm TrueI},  (*No need for mutual induction*)
 | 
| 12175 | 143 | exhaustion = elim}; | 
| 6070 | 144 | |
| 145 | val con_info = | |
| 12175 | 146 |           {big_rec_name = big_rec_name,
 | 
| 147 | constructors = constructors, | |
| 148 | (*let primrec handle definition by cases*) | |
| 149 | free_iffs = [], (*thus we expect the necessary freeness rewrites | |
| 150 | to be in the simpset already, as is the case for | |
| 151 | Nat and disjoint sum*) | |
| 152 | rec_rewrites = (case recursor_eqns of | |
| 153 | [] => case_eqns | _ => recursor_eqns)}; | |
| 6070 | 154 | |
| 155 | (*associate with each constructor the datatype name and rewrites*) | |
| 156 | val con_pairs = map (fn c => (#1 (dest_Const c), con_info)) constructors | |
| 157 | ||
| 158 | in | |
| 17223 | 159 | thy | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 160 | |> Sign.add_path (Long_Name.base_name big_rec_name) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38522diff
changeset | 161 | |> Global_Theory.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd | 
| 17412 | 162 | |> DatatypesData.put (Symtab.update (big_rec_name, dt_info) (DatatypesData.get thy)) | 
| 163 | |> ConstructorsData.put (fold_rev Symtab.update con_pairs (ConstructorsData.get thy)) | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
22846diff
changeset | 164 | |> Sign.parent_path | 
| 12204 | 165 | end; | 
| 6065 | 166 | |
| 12204 | 167 | fun rep_datatype raw_elim raw_induct raw_case_eqns raw_recursor_eqns thy = | 
| 24725 | 168 | let | 
| 42361 | 169 | val ctxt = Proof_Context.init_global thy; | 
| 55740 | 170 |     val elim = Facts.the_single ("elimination", Position.none) (Attrib.eval_thms ctxt [raw_elim]);
 | 
| 171 |     val induct = Facts.the_single ("induction", Position.none) (Attrib.eval_thms ctxt [raw_induct]);
 | |
| 24725 | 172 | val case_eqns = Attrib.eval_thms ctxt raw_case_eqns; | 
| 173 | val recursor_eqns = Attrib.eval_thms ctxt raw_recursor_eqns; | |
| 174 | in rep_datatype_i elim induct case_eqns recursor_eqns thy end; | |
| 12175 | 175 | |
| 12204 | 176 | |
| 177 | (* theory setup *) | |
| 178 | ||
| 58828 | 179 | val _ = | 
| 180 | Theory.setup | |
| 181 |     (Method.setup @{binding induct_tac}
 | |
| 63120 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 182 | (Args.goal_spec -- Scan.lift (Args.embedded -- Parse.for_fixes) >> | 
| 59788 | 183 | (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt s xs))) | 
| 58828 | 184 | "induct_tac emulation (dynamic instantiation!)" #> | 
| 185 |     Method.setup @{binding case_tac}
 | |
| 63120 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 186 | (Args.goal_spec -- Scan.lift (Args.embedded -- Parse.for_fixes) >> | 
| 59788 | 187 | (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (exhaust_tac ctxt s xs))) | 
| 58828 | 188 | "datatype case_tac emulation (dynamic instantiation!)"); | 
| 12204 | 189 | |
| 190 | ||
| 191 | (* outer syntax *) | |
| 192 | ||
| 24867 | 193 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59788diff
changeset | 194 |   Outer_Syntax.command @{command_keyword rep_datatype} "represent existing set inductively"
 | 
| 62969 | 195 |     ((@{keyword "elimination"} |-- Parse.!!! Parse.thm) --
 | 
| 196 |      (@{keyword "induction"} |-- Parse.!!! Parse.thm) --
 | |
| 197 |      (@{keyword "case_eqns"} |-- Parse.!!! Parse.thms1) --
 | |
| 198 |      Scan.optional (@{keyword "recursor_eqns"} |-- Parse.!!! Parse.thms1) []
 | |
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 199 | >> (fn (((x, y), z), w) => Toplevel.theory (rep_datatype x y z w))); | 
| 12204 | 200 | |
| 201 | end; | |
| 202 | ||
| 203 | val exhaust_tac = DatatypeTactics.exhaust_tac; | |
| 204 | val induct_tac = DatatypeTactics.induct_tac; |