| author | blanchet | 
| Wed, 15 Sep 2010 19:55:26 +0200 | |
| changeset 39430 | afb4d5c672bd | 
| parent 39133 | 70d3915c92f0 | 
| child 40242 | bb433b0668b8 | 
| permissions | -rw-r--r-- | 
| 18830 | 1  | 
(* Title: Pure/Isar/local_defs.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 18840 | 4  | 
Local definitions.  | 
| 18830 | 5  | 
*)  | 
6  | 
||
7  | 
signature LOCAL_DEFS =  | 
|
8  | 
sig  | 
|
| 20306 | 9  | 
val cert_def: Proof.context -> term -> (string * typ) * term  | 
| 18830 | 10  | 
val abs_def: term -> (string * typ) * term  | 
| 20306 | 11  | 
val mk_def: Proof.context -> (string * term) list -> term list  | 
| 21684 | 12  | 
val expand: cterm list -> thm -> thm  | 
| 
20224
 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 
wenzelm 
parents: 
20049 
diff
changeset
 | 
13  | 
val def_export: Assumption.export  | 
| 30211 | 14  | 
val add_defs: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->  | 
15  | 
(term * (string * thm)) list * Proof.context  | 
|
| 29581 | 16  | 
val add_def: (binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context  | 
17  | 
val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->  | 
|
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
18  | 
(term * term) * Proof.context  | 
| 21591 | 19  | 
val export: Proof.context -> Proof.context -> thm -> thm list * thm  | 
| 35717 | 20  | 
val export_cterm: Proof.context -> Proof.context -> cterm -> thm list * cterm  | 
| 21844 | 21  | 
val trans_terms: Proof.context -> thm list -> thm  | 
22  | 
val trans_props: Proof.context -> thm list -> thm  | 
|
| 35739 | 23  | 
val contract: Proof.context -> thm list -> cterm -> thm -> thm  | 
| 21506 | 24  | 
val print_rules: Proof.context -> unit  | 
| 18840 | 25  | 
val defn_add: attribute  | 
26  | 
val defn_del: attribute  | 
|
| 23541 | 27  | 
val meta_rewrite_conv: Proof.context -> conv  | 
| 21506 | 28  | 
val meta_rewrite_rule: Proof.context -> thm -> thm  | 
| 20306 | 29  | 
val unfold: Proof.context -> thm list -> thm -> thm  | 
30  | 
val unfold_goals: Proof.context -> thm list -> thm -> thm  | 
|
31  | 
val unfold_tac: Proof.context -> thm list -> tactic  | 
|
32  | 
val fold: Proof.context -> thm list -> thm -> thm  | 
|
33  | 
val fold_tac: Proof.context -> thm list -> tactic  | 
|
34  | 
val derived_def: Proof.context -> bool -> term ->  | 
|
| 20909 | 35  | 
((string * typ) * term) * (Proof.context -> thm -> thm)  | 
| 18830 | 36  | 
end;  | 
37  | 
||
| 35624 | 38  | 
structure Local_Defs: LOCAL_DEFS =  | 
| 18830 | 39  | 
struct  | 
40  | 
||
| 18840 | 41  | 
(** primitive definitions **)  | 
42  | 
||
| 18830 | 43  | 
(* prepare defs *)  | 
44  | 
||
45  | 
fun cert_def ctxt eq =  | 
|
46  | 
let  | 
|
| 25025 | 47  | 
    fun err msg = cat_error msg ("The error(s) above occurred in definition: " ^
 | 
48  | 
quote (Syntax.string_of_term ctxt eq));  | 
|
| 18950 | 49  | 
val ((lhs, _), eq') = eq  | 
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
35739 
diff
changeset
 | 
50  | 
|> Sign.no_vars ctxt  | 
| 33385 | 51  | 
|> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true)  | 
| 18950 | 52  | 
handle TERM (msg, _) => err msg | ERROR msg => err msg;  | 
53  | 
in (Term.dest_Free (Term.head_of lhs), eq') end;  | 
|
| 18830 | 54  | 
|
| 33385 | 55  | 
val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free;  | 
| 18830 | 56  | 
|
| 18875 | 57  | 
fun mk_def ctxt args =  | 
58  | 
let  | 
|
59  | 
val (xs, rhss) = split_list args;  | 
|
60  | 
val (bind, _) = ProofContext.bind_fixes xs ctxt;  | 
|
61  | 
val lhss = map (fn (x, rhs) => bind (Free (x, Term.fastype_of rhs))) args;  | 
|
62  | 
in map Logic.mk_equals (lhss ~~ rhss) end;  | 
|
63  | 
||
| 18830 | 64  | 
|
65  | 
(* export defs *)  | 
|
66  | 
||
| 20021 | 67  | 
val head_of_def =  | 
| 20887 | 68  | 
#1 o Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body;  | 
| 20021 | 69  | 
|
| 18830 | 70  | 
|
| 18875 | 71  | 
(*  | 
| 20887 | 72  | 
[x, x == a]  | 
| 18896 | 73  | 
:  | 
74  | 
B x  | 
|
75  | 
-----------  | 
|
| 20887 | 76  | 
B a  | 
| 18875 | 77  | 
*)  | 
| 21684 | 78  | 
fun expand defs =  | 
79  | 
Drule.implies_intr_list defs  | 
|
80  | 
#> Drule.generalize ([], map (head_of_def o Thm.term_of) defs)  | 
|
81  | 
#> funpow (length defs) (fn th => Drule.reflexive_thm RS th);  | 
|
82  | 
||
| 21801 | 83  | 
val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of);  | 
| 21684 | 84  | 
|
85  | 
fun def_export _ defs = (expand defs, expand_term defs);  | 
|
| 18830 | 86  | 
|
87  | 
||
88  | 
(* add defs *)  | 
|
89  | 
||
| 20887 | 90  | 
fun add_defs defs ctxt =  | 
| 18830 | 91  | 
let  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
92  | 
val ((bvars, mxs), specs) = defs |> split_list |>> split_list;  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
93  | 
val ((bfacts, atts), rhss) = specs |> split_list |>> split_list;  | 
| 
30585
 
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
94  | 
val xs = map Name.of_binding bvars;  | 
| 30434 | 95  | 
val names = map2 Thm.def_binding_optional bvars bfacts;  | 
| 20887 | 96  | 
val eqs = mk_def ctxt (xs ~~ rhss);  | 
97  | 
val lhss = map (fst o Logic.dest_equals) eqs;  | 
|
| 18830 | 98  | 
in  | 
99  | 
ctxt  | 
|
| 
30763
 
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
100  | 
|> ProofContext.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) bvars mxs) |> #2  | 
| 20980 | 101  | 
|> fold Variable.declare_term eqs  | 
| 20887 | 102  | 
|> ProofContext.add_assms_i def_export  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
103  | 
(map2 (fn a => fn eq => (a, [(eq, [])])) (names ~~ atts) eqs)  | 
| 20887 | 104  | 
|>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss  | 
| 18830 | 105  | 
end;  | 
106  | 
||
| 25104 | 107  | 
fun add_def (var, rhs) ctxt =  | 
| 30211 | 108  | 
let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt  | 
| 25104 | 109  | 
in ((lhs, th), ctxt') end;  | 
| 25025 | 110  | 
|
| 18840 | 111  | 
|
| 25119 | 112  | 
(* fixed_abbrev *)  | 
113  | 
||
114  | 
fun fixed_abbrev ((x, mx), rhs) ctxt =  | 
|
115  | 
let  | 
|
116  | 
val T = Term.fastype_of rhs;  | 
|
117  | 
val ([x'], ctxt') = ctxt  | 
|
118  | 
|> Variable.declare_term rhs  | 
|
| 
30763
 
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
119  | 
|> ProofContext.add_fixes [(x, SOME T, mx)];  | 
| 25119 | 120  | 
val lhs = Free (x', T);  | 
121  | 
val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs));  | 
|
122  | 
fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]);  | 
|
123  | 
val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';  | 
|
124  | 
in ((lhs, rhs), ctxt'') end;  | 
|
125  | 
||
126  | 
||
| 24985 | 127  | 
(* specific export -- result based on educated guessing *)  | 
| 21568 | 128  | 
|
| 24985 | 129  | 
(*  | 
130  | 
[xs, xs == as]  | 
|
131  | 
:  | 
|
132  | 
B xs  | 
|
133  | 
--------------  | 
|
134  | 
B as  | 
|
135  | 
*)  | 
|
136  | 
fun export inner outer = (*beware of closure sizes*)  | 
|
| 21568 | 137  | 
let  | 
| 22671 | 138  | 
val exp = Assumption.export false inner outer;  | 
| 
30473
 
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
 
wenzelm 
parents: 
30434 
diff
changeset
 | 
139  | 
val prems = Assumption.all_prems_of inner;  | 
| 22671 | 140  | 
in fn th =>  | 
141  | 
let  | 
|
142  | 
val th' = exp th;  | 
|
| 22691 | 143  | 
val still_fixed = map #1 (Thm.fold_terms Term.add_frees th' []);  | 
| 22671 | 144  | 
val defs = prems |> filter_out (fn prem =>  | 
145  | 
(case try (head_of_def o Thm.prop_of) prem of  | 
|
146  | 
SOME x => member (op =) still_fixed x  | 
|
147  | 
| NONE => true));  | 
|
148  | 
in (map Drule.abs_def defs, th') end  | 
|
149  | 
end;  | 
|
| 21568 | 150  | 
|
| 24985 | 151  | 
(*  | 
152  | 
[xs, xs == as]  | 
|
153  | 
:  | 
|
154  | 
TERM b xs  | 
|
155  | 
-------------- and --------------  | 
|
156  | 
TERM b as b xs == b as  | 
|
157  | 
*)  | 
|
158  | 
fun export_cterm inner outer ct =  | 
|
| 35717 | 159  | 
export inner outer (Drule.mk_term ct) ||> Drule.dest_term;  | 
| 24985 | 160  | 
|
| 21568 | 161  | 
|
| 21844 | 162  | 
(* basic transitivity reasoning -- modulo beta-eta *)  | 
163  | 
||
164  | 
local  | 
|
165  | 
||
166  | 
val is_trivial = Pattern.aeconv o Logic.dest_equals o Thm.prop_of;  | 
|
167  | 
||
168  | 
fun trans_list _ _ [] = raise Empty  | 
|
169  | 
| trans_list trans ctxt (th :: raw_eqs) =  | 
|
170  | 
(case filter_out is_trivial raw_eqs of  | 
|
171  | 
[] => th  | 
|
172  | 
| eqs =>  | 
|
| 
31794
 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
173  | 
let val ((_, th' :: eqs'), ctxt') = Variable.import true (th :: eqs) ctxt  | 
| 21844 | 174  | 
in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end);  | 
175  | 
||
176  | 
in  | 
|
177  | 
||
178  | 
val trans_terms = trans_list (fn eq2 => fn eq1 => eq2 COMP (eq1 COMP Drule.transitive_thm));  | 
|
179  | 
val trans_props = trans_list (fn eq => fn th => th COMP (eq COMP Drule.equal_elim_rule1));  | 
|
180  | 
||
181  | 
end;  | 
|
182  | 
||
| 35739 | 183  | 
fun contract ctxt defs ct th =  | 
184  | 
trans_props ctxt [th, Thm.symmetric (MetaSimplifier.rewrite true defs ct)];  | 
|
| 21844 | 185  | 
|
| 18840 | 186  | 
|
187  | 
(** defived definitions **)  | 
|
188  | 
||
189  | 
(* transformation rules *)  | 
|
190  | 
||
| 33519 | 191  | 
structure Rules = Generic_Data  | 
| 18840 | 192  | 
(  | 
193  | 
type T = thm list;  | 
|
| 33519 | 194  | 
val empty = [];  | 
| 18840 | 195  | 
val extend = I;  | 
| 33519 | 196  | 
val merge = Thm.merge_thms;  | 
| 18840 | 197  | 
);  | 
198  | 
||
| 22846 | 199  | 
fun print_rules ctxt =  | 
200  | 
Pretty.writeln (Pretty.big_list "definitional transformations:"  | 
|
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
201  | 
(map (Display.pretty_thm ctxt) (Rules.get (Context.Proof ctxt))));  | 
| 18840 | 202  | 
|
| 
24039
 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23541 
diff
changeset
 | 
203  | 
val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm);  | 
| 
 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23541 
diff
changeset
 | 
204  | 
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);  | 
| 18840 | 205  | 
|
206  | 
||
| 18875 | 207  | 
(* meta rewrite rules *)  | 
| 18840 | 208  | 
|
| 23541 | 209  | 
fun meta_rewrite_conv ctxt =  | 
| 18840 | 210  | 
MetaSimplifier.rewrite_cterm (false, false, false) (K (K NONE))  | 
| 26424 | 211  | 
(MetaSimplifier.context ctxt MetaSimplifier.empty_ss  | 
212  | 
addeqcongs [Drule.equals_cong] (*protect meta-level equality*)  | 
|
213  | 
addsimps (Rules.get (Context.Proof ctxt)));  | 
|
| 18840 | 214  | 
|
| 23541 | 215  | 
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;  | 
| 18840 | 216  | 
|
| 18875 | 217  | 
|
218  | 
(* rewriting with object-level rules *)  | 
|
219  | 
||
| 21506 | 220  | 
fun meta f ctxt = f o map (meta_rewrite_rule ctxt);  | 
| 18840 | 221  | 
|
| 21708 | 222  | 
val unfold = meta MetaSimplifier.rewrite_rule;  | 
223  | 
val unfold_goals = meta MetaSimplifier.rewrite_goals_rule;  | 
|
224  | 
val unfold_tac = meta MetaSimplifier.rewrite_goals_tac;  | 
|
225  | 
val fold = meta MetaSimplifier.fold_rule;  | 
|
226  | 
val fold_tac = meta MetaSimplifier.fold_goals_tac;  | 
|
| 18840 | 227  | 
|
228  | 
||
229  | 
(* derived defs -- potentially within the object-logic *)  | 
|
230  | 
||
| 18950 | 231  | 
fun derived_def ctxt conditional prop =  | 
| 18840 | 232  | 
let  | 
233  | 
val ((c, T), rhs) = prop  | 
|
| 20049 | 234  | 
|> Thm.cterm_of (ProofContext.theory_of ctxt)  | 
| 23541 | 235  | 
|> meta_rewrite_conv ctxt  | 
| 18840 | 236  | 
|> (snd o Logic.dest_equals o Thm.prop_of)  | 
| 21568 | 237  | 
|> conditional ? Logic.strip_imp_concl  | 
| 18950 | 238  | 
|> (abs_def o #2 o cert_def ctxt);  | 
| 20909 | 239  | 
fun prove ctxt' def =  | 
| 18840 | 240  | 
let  | 
241  | 
val frees = Term.fold_aterms (fn Free (x, _) =>  | 
|
| 20909 | 242  | 
if Variable.is_fixed ctxt' x then I else insert (op =) x | _ => I) prop [];  | 
| 18840 | 243  | 
in  | 
| 20909 | 244  | 
Goal.prove ctxt' frees [] prop (K (ALLGOALS  | 
| 23541 | 245  | 
(CONVERSION (meta_rewrite_conv ctxt') THEN'  | 
246  | 
MetaSimplifier.rewrite_goal_tac [def] THEN'  | 
|
| 21708 | 247  | 
resolve_tac [Drule.reflexive_thm])))  | 
| 18840 | 248  | 
handle ERROR msg => cat_error msg "Failed to prove definitional specification."  | 
249  | 
end;  | 
|
250  | 
in (((c, T), rhs), prove) end;  | 
|
251  | 
||
| 18830 | 252  | 
end;  |