| author | bulwahn | 
| Fri, 11 Nov 2011 12:10:49 +0100 | |
| changeset 45461 | 130c90bb80b4 | 
| parent 44058 | ae85c5d64913 | 
| child 46947 | b8c7eb0c2f89 | 
| 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  | 
|
| 
27208
 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
12  | 
val induct_tac: Proof.context -> string -> int -> tactic  | 
| 
 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
13  | 
val exhaust_tac: Proof.context -> string -> int -> tactic  | 
| 12204 | 14  | 
val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
24867 
diff
changeset
 | 
15  | 
val rep_datatype: Facts.ref * Attrib.src list -> Facts.ref * Attrib.src list ->  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
24867 
diff
changeset
 | 
16  | 
(Facts.ref * Attrib.src list) list -> (Facts.ref * Attrib.src list) list -> theory -> theory  | 
| 18708 | 17  | 
val setup: theory -> theory  | 
| 6065 | 18  | 
end;  | 
19  | 
||
20  | 
||
| 6070 | 21  | 
(** Datatype information, e.g. associated theorems **)  | 
22  | 
||
23  | 
type datatype_info =  | 
|
| 12175 | 24  | 
  {inductive: bool,             (*true if inductive, not coinductive*)
 | 
| 6070 | 25  | 
constructors : term list, (*the constructors, as Consts*)  | 
26  | 
rec_rewrites : thm list, (*recursor equations*)  | 
|
27  | 
case_rewrites : thm list, (*case equations*)  | 
|
28  | 
induct : thm,  | 
|
29  | 
mutual_induct : thm,  | 
|
30  | 
exhaustion : thm};  | 
|
31  | 
||
| 33522 | 32  | 
structure DatatypesData = Theory_Data  | 
| 22846 | 33  | 
(  | 
| 6070 | 34  | 
type T = datatype_info Symtab.table;  | 
35  | 
val empty = Symtab.empty;  | 
|
| 16458 | 36  | 
val extend = I;  | 
| 33522 | 37  | 
fun merge data : T = Symtab.merge (K true) data;  | 
| 22846 | 38  | 
);  | 
| 6070 | 39  | 
|
40  | 
||
41  | 
||
42  | 
(** Constructor information: needed to map constructors to datatypes **)  | 
|
43  | 
||
44  | 
type constructor_info =  | 
|
45  | 
  {big_rec_name : string,     (*name of the mutually recursive set*)
 | 
|
46  | 
constructors : term list, (*the constructors, as Consts*)  | 
|
| 6141 | 47  | 
free_iffs : thm list, (*freeness simprules*)  | 
| 6070 | 48  | 
rec_rewrites : thm list}; (*recursor equations*)  | 
49  | 
||
50  | 
||
| 33522 | 51  | 
structure ConstructorsData = Theory_Data  | 
| 22846 | 52  | 
(  | 
| 6070 | 53  | 
type T = constructor_info Symtab.table  | 
54  | 
val empty = Symtab.empty  | 
|
| 16458 | 55  | 
val extend = I  | 
| 33522 | 56  | 
fun merge data = Symtab.merge (K true) data;  | 
| 22846 | 57  | 
);  | 
| 6070 | 58  | 
|
| 6065 | 59  | 
structure DatatypeTactics : DATATYPE_TACTICS =  | 
60  | 
struct  | 
|
61  | 
||
| 16458 | 62  | 
fun datatype_info thy name =  | 
| 17412 | 63  | 
(case Symtab.lookup (DatatypesData.get thy) name of  | 
| 15531 | 64  | 
SOME info => info  | 
65  | 
  | NONE => error ("Unknown datatype " ^ quote name));
 | 
|
| 6065 | 66  | 
|
67  | 
||
68  | 
(*Given a variable, find the inductive set associated it in the assumptions*)  | 
|
| 14153 | 69  | 
exception Find_tname of string  | 
70  | 
||
| 37780 | 71  | 
fun find_tname ctxt var As =  | 
| 24826 | 72  | 
  let fun mk_pair (Const(@{const_name mem},_) $ Free (v,_) $ A) =
 | 
| 6065 | 73  | 
(v, #1 (dest_Const (head_of A)))  | 
| 12175 | 74  | 
| mk_pair _ = raise Match  | 
| 37780 | 75  | 
val pairs = map_filter (try (mk_pair o FOLogic.dest_Trueprop)) As  | 
76  | 
val x =  | 
|
77  | 
(case try (dest_Free o Syntax.read_term ctxt) var of  | 
|
78  | 
SOME (x, _) => x  | 
|
79  | 
        | _ => raise Find_tname ("Bad variable " ^ quote var))
 | 
|
80  | 
in case AList.lookup (op =) pairs x of  | 
|
| 15531 | 81  | 
       NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
 | 
82  | 
| SOME t => t  | 
|
| 6065 | 83  | 
end;  | 
84  | 
||
| 12175 | 85  | 
(** generic exhaustion and induction tactic for datatypes  | 
86  | 
Differences from HOL:  | 
|
| 6065 | 87  | 
(1) no checking if the induction var occurs in premises, since it always  | 
88  | 
appears in one of them, and it's hard to check for other occurrences  | 
|
89  | 
(2) exhaustion works for VARIABLES in the premises, not general terms  | 
|
90  | 
**)  | 
|
91  | 
||
| 
27208
 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
92  | 
fun exhaust_induct_tac exh ctxt var i state =  | 
| 6065 | 93  | 
let  | 
| 42361 | 94  | 
val thy = Proof_Context.theory_of ctxt  | 
| 37780 | 95  | 
    val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i state
 | 
96  | 
val tn = find_tname ctxt' var (map term_of asms)  | 
|
| 12175 | 97  | 
val rule =  | 
| 16458 | 98  | 
if exh then #exhaustion (datatype_info thy tn)  | 
99  | 
else #induct (datatype_info thy tn)  | 
|
| 24826 | 100  | 
    val (Const(@{const_name mem},_) $ Var(ixn,_) $ _) =
 | 
| 6112 | 101  | 
(case prems_of rule of  | 
| 12175 | 102  | 
[] => error "induction is not available for this datatype"  | 
103  | 
| major::_ => FOLogic.dest_Trueprop major)  | 
|
| 6065 | 104  | 
in  | 
| 27239 | 105  | 
eres_inst_tac ctxt [(ixn, var)] rule i state  | 
| 14153 | 106  | 
end  | 
107  | 
handle Find_tname msg =>  | 
|
108  | 
if exh then (*try boolean case analysis instead*)  | 
|
| 
27208
 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
109  | 
case_tac ctxt var i state  | 
| 14153 | 110  | 
else error msg;  | 
| 6065 | 111  | 
|
112  | 
val exhaust_tac = exhaust_induct_tac true;  | 
|
113  | 
val induct_tac = exhaust_induct_tac false;  | 
|
114  | 
||
| 6070 | 115  | 
|
116  | 
(**** declare non-datatype as datatype ****)  | 
|
117  | 
||
118  | 
fun rep_datatype_i elim induct case_eqns recursor_eqns thy =  | 
|
119  | 
let  | 
|
120  | 
(*analyze the LHS of a case equation to get a constructor*)  | 
|
| 41310 | 121  | 
    fun const_of (Const(@{const_name IFOL.eq}, _) $ (_ $ c) $ _) = c
 | 
| 6070 | 122  | 
      | const_of eqn = error ("Ill-formed case equation: " ^
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26618 
diff
changeset
 | 
123  | 
Syntax.string_of_term_global thy eqn);  | 
| 6070 | 124  | 
|
125  | 
val constructors =  | 
|
| 44058 | 126  | 
map (head_of o const_of o FOLogic.dest_Trueprop o Thm.prop_of) case_eqns;  | 
| 6070 | 127  | 
|
| 24826 | 128  | 
    val Const (@{const_name mem}, _) $ _ $ data =
 | 
| 12175 | 129  | 
FOLogic.dest_Trueprop (hd (prems_of elim));  | 
130  | 
||
| 6112 | 131  | 
val Const(big_rec_name, _) = head_of data;  | 
132  | 
||
| 6070 | 133  | 
val simps = case_eqns @ recursor_eqns;  | 
134  | 
||
135  | 
val dt_info =  | 
|
| 12175 | 136  | 
          {inductive = true,
 | 
137  | 
constructors = constructors,  | 
|
138  | 
rec_rewrites = recursor_eqns,  | 
|
139  | 
case_rewrites = case_eqns,  | 
|
140  | 
induct = induct,  | 
|
| 35409 | 141  | 
           mutual_induct = @{thm TrueI},  (*No need for mutual induction*)
 | 
| 12175 | 142  | 
exhaustion = elim};  | 
| 6070 | 143  | 
|
144  | 
val con_info =  | 
|
| 12175 | 145  | 
          {big_rec_name = big_rec_name,
 | 
146  | 
constructors = constructors,  | 
|
147  | 
(*let primrec handle definition by cases*)  | 
|
148  | 
free_iffs = [], (*thus we expect the necessary freeness rewrites  | 
|
149  | 
to be in the simpset already, as is the case for  | 
|
150  | 
Nat and disjoint sum*)  | 
|
151  | 
rec_rewrites = (case recursor_eqns of  | 
|
152  | 
[] => case_eqns | _ => recursor_eqns)};  | 
|
| 6070 | 153  | 
|
154  | 
(*associate with each constructor the datatype name and rewrites*)  | 
|
155  | 
val con_pairs = map (fn c => (#1 (dest_Const c), con_info)) constructors  | 
|
156  | 
||
157  | 
in  | 
|
| 17223 | 158  | 
thy  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
159  | 
|> 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: 
38522 
diff
changeset
 | 
160  | 
|> Global_Theory.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd  | 
| 17412 | 161  | 
|> DatatypesData.put (Symtab.update (big_rec_name, dt_info) (DatatypesData.get thy))  | 
162  | 
|> ConstructorsData.put (fold_rev Symtab.update con_pairs (ConstructorsData.get thy))  | 
|
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
22846 
diff
changeset
 | 
163  | 
|> Sign.parent_path  | 
| 12204 | 164  | 
end;  | 
| 6065 | 165  | 
|
| 12204 | 166  | 
fun rep_datatype raw_elim raw_induct raw_case_eqns raw_recursor_eqns thy =  | 
| 24725 | 167  | 
let  | 
| 42361 | 168  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
24867 
diff
changeset
 | 
169  | 
val elim = Facts.the_single "elimination" (Attrib.eval_thms ctxt [raw_elim]);  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
24867 
diff
changeset
 | 
170  | 
val induct = Facts.the_single "induction" (Attrib.eval_thms ctxt [raw_induct]);  | 
| 24725 | 171  | 
val case_eqns = Attrib.eval_thms ctxt raw_case_eqns;  | 
172  | 
val recursor_eqns = Attrib.eval_thms ctxt raw_recursor_eqns;  | 
|
173  | 
in rep_datatype_i elim induct case_eqns recursor_eqns thy end;  | 
|
| 12175 | 174  | 
|
| 12204 | 175  | 
|
176  | 
(* theory setup *)  | 
|
177  | 
||
178  | 
val setup =  | 
|
| 30541 | 179  | 
  Method.setup @{binding induct_tac}
 | 
180  | 
(Args.goal_spec -- Scan.lift Args.name >>  | 
|
181  | 
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt s)))  | 
|
182  | 
"induct_tac emulation (dynamic instantiation!)" #>  | 
|
183  | 
  Method.setup @{binding case_tac}
 | 
|
184  | 
(Args.goal_spec -- Scan.lift Args.name >>  | 
|
185  | 
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (exhaust_tac ctxt s)))  | 
|
186  | 
"datatype case_tac emulation (dynamic instantiation!)";  | 
|
| 12204 | 187  | 
|
188  | 
||
189  | 
(* outer syntax *)  | 
|
190  | 
||
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
191  | 
val _ = List.app Keyword.keyword ["elimination", "induction", "case_eqns", "recursor_eqns"];  | 
| 24867 | 192  | 
|
| 12204 | 193  | 
val rep_datatype_decl =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
194  | 
(Parse.$$$ "elimination" |-- Parse.!!! Parse_Spec.xthm) --  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
195  | 
(Parse.$$$ "induction" |-- Parse.!!! Parse_Spec.xthm) --  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
196  | 
(Parse.$$$ "case_eqns" |-- Parse.!!! Parse_Spec.xthms1) --  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
197  | 
Scan.optional (Parse.$$$ "recursor_eqns" |-- Parse.!!! Parse_Spec.xthms1) []  | 
| 12204 | 198  | 
>> (fn (((x, y), z), w) => rep_datatype x y z w);  | 
199  | 
||
| 24867 | 200  | 
val _ =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
201  | 
Outer_Syntax.command "rep_datatype" "represent existing set inductively" Keyword.thy_decl  | 
| 12204 | 202  | 
(rep_datatype_decl >> Toplevel.theory);  | 
203  | 
||
204  | 
end;  | 
|
205  | 
||
206  | 
val exhaust_tac = DatatypeTactics.exhaust_tac;  | 
|
207  | 
val induct_tac = DatatypeTactics.induct_tac;  |