| author | wenzelm | 
| Wed, 21 Dec 2022 13:22:57 +0100 | |
| changeset 76727 | 6d95e8a636e2 | 
| parent 74575 | ccf599864beb | 
| child 78086 | 5edd5b12017d | 
| 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  | 
|
| 63395 | 9  | 
val cert_def: Proof.context -> (string -> Position.T list) -> term -> (string * typ) * term  | 
| 18830 | 10  | 
val abs_def: term -> (string * typ) * term  | 
| 21684 | 11  | 
val expand: cterm list -> thm -> thm  | 
| 
20224
 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 
wenzelm 
parents: 
20049 
diff
changeset
 | 
12  | 
val def_export: Assumption.export  | 
| 63344 | 13  | 
val define: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->  | 
| 30211 | 14  | 
(term * (string * thm)) list * Proof.context  | 
| 29581 | 15  | 
val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
16  | 
(term * term) * Proof.context  | 
| 
45407
 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 
wenzelm 
parents: 
45406 
diff
changeset
 | 
17  | 
val export: Proof.context -> Proof.context -> thm -> (thm list * thm list) * thm  | 
| 
 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 
wenzelm 
parents: 
45406 
diff
changeset
 | 
18  | 
val export_cterm: Proof.context -> Proof.context -> cterm -> (thm list * thm list) * cterm  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
19  | 
val contract: Proof.context -> thm list -> cterm -> thm -> thm  | 
| 21506 | 20  | 
val print_rules: Proof.context -> unit  | 
| 18840 | 21  | 
val defn_add: attribute  | 
22  | 
val defn_del: attribute  | 
|
| 23541 | 23  | 
val meta_rewrite_conv: Proof.context -> conv  | 
| 21506 | 24  | 
val meta_rewrite_rule: Proof.context -> thm -> thm  | 
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
25  | 
val abs_def_rule: Proof.context -> thm -> thm  | 
| 
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
26  | 
val unfold_abs_def: bool Config.T  | 
| 20306 | 27  | 
val unfold: Proof.context -> thm list -> thm -> thm  | 
28  | 
val unfold_goals: Proof.context -> thm list -> thm -> thm  | 
|
29  | 
val unfold_tac: Proof.context -> thm list -> tactic  | 
|
| 63169 | 30  | 
val unfold0: Proof.context -> thm list -> thm -> thm  | 
31  | 
val unfold0_goals: Proof.context -> thm list -> thm -> thm  | 
|
32  | 
val unfold0_tac: Proof.context -> thm list -> tactic  | 
|
| 20306 | 33  | 
val fold: Proof.context -> thm list -> thm -> thm  | 
34  | 
val fold_tac: Proof.context -> thm list -> tactic  | 
|
| 63395 | 35  | 
  val derived_def: Proof.context -> (string -> Position.T list) -> {conditional: bool} ->
 | 
36  | 
term -> ((string * typ) * term) * (Proof.context -> thm -> thm)  | 
|
| 18830 | 37  | 
end;  | 
38  | 
||
| 35624 | 39  | 
structure Local_Defs: LOCAL_DEFS =  | 
| 18830 | 40  | 
struct  | 
41  | 
||
| 18840 | 42  | 
(** primitive definitions **)  | 
43  | 
||
| 18830 | 44  | 
(* prepare defs *)  | 
45  | 
||
| 63395 | 46  | 
fun cert_def ctxt get_pos eq =  | 
| 18830 | 47  | 
let  | 
| 40242 | 48  | 
fun err msg =  | 
49  | 
      cat_error msg ("The error(s) above occurred in definition:\n" ^
 | 
|
50  | 
quote (Syntax.string_of_term ctxt eq));  | 
|
| 63395 | 51  | 
val ((lhs, _), args, 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
 | 
52  | 
|> Sign.no_vars ctxt  | 
| 63042 | 53  | 
|> Primitive_Defs.dest_def ctxt  | 
54  | 
        {check_head = Term.is_Free,
 | 
|
55  | 
check_free_lhs = not o Variable.is_fixed ctxt,  | 
|
56  | 
check_free_rhs = if Variable.is_body ctxt then K true else Variable.is_fixed ctxt,  | 
|
57  | 
check_tfree = K true}  | 
|
| 18950 | 58  | 
handle TERM (msg, _) => err msg | ERROR msg => err msg;  | 
| 63395 | 59  | 
val _ =  | 
60  | 
Context_Position.reports ctxt  | 
|
61  | 
(maps (fn Free (x, _) => Syntax_Phases.reports_of_scope (get_pos x) | _ => []) args);  | 
|
| 18950 | 62  | 
in (Term.dest_Free (Term.head_of lhs), eq') end;  | 
| 18830 | 63  | 
|
| 33385 | 64  | 
val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free;  | 
| 18830 | 65  | 
|
| 18875 | 66  | 
fun mk_def ctxt args =  | 
67  | 
let  | 
|
| 
42501
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
68  | 
val (bs, rhss) = split_list args;  | 
| 
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
69  | 
val Ts = map Term.fastype_of rhss;  | 
| 63344 | 70  | 
val (xs, _) = ctxt  | 
71  | 
|> Context_Position.set_visible false  | 
|
72  | 
|> Proof_Context.add_fixes (map2 (fn b => fn T => (b, SOME T, NoSyn)) bs Ts);  | 
|
| 
42501
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
73  | 
val lhss = ListPair.map Free (xs, Ts);  | 
| 18875 | 74  | 
in map Logic.mk_equals (lhss ~~ rhss) end;  | 
75  | 
||
| 18830 | 76  | 
|
77  | 
(* export defs *)  | 
|
78  | 
||
| 20021 | 79  | 
val head_of_def =  | 
| 
45406
 
b24ecaa46edb
clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
 
wenzelm 
parents: 
42501 
diff
changeset
 | 
80  | 
Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body;  | 
| 20021 | 81  | 
|
| 18830 | 82  | 
|
| 18875 | 83  | 
(*  | 
| 67721 | 84  | 
[x, x \<equiv> a]  | 
| 18896 | 85  | 
:  | 
86  | 
B x  | 
|
87  | 
-----------  | 
|
| 20887 | 88  | 
B a  | 
| 18875 | 89  | 
*)  | 
| 21684 | 90  | 
fun expand defs =  | 
91  | 
Drule.implies_intr_list defs  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
70314 
diff
changeset
 | 
92  | 
#> Drule.generalize  | 
| 74266 | 93  | 
(Names.empty, Names.build (fold (Names.add_set o #1 o head_of_def o Thm.term_of) defs))  | 
| 21684 | 94  | 
#> funpow (length defs) (fn th => Drule.reflexive_thm RS th);  | 
95  | 
||
| 74575 | 96  | 
val expand_term = Envir.expand_term_defs dest_Free o map (abs_def o Thm.term_of);  | 
| 21684 | 97  | 
|
98  | 
fun def_export _ defs = (expand defs, expand_term defs);  | 
|
| 18830 | 99  | 
|
100  | 
||
| 63344 | 101  | 
(* define *)  | 
| 18830 | 102  | 
|
| 63344 | 103  | 
fun define defs ctxt =  | 
| 18830 | 104  | 
let  | 
| 42496 | 105  | 
val ((xs, mxs), specs) = defs |> split_list |>> split_list;  | 
| 
49748
 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 
wenzelm 
parents: 
47296 
diff
changeset
 | 
106  | 
val (bs, rhss) = specs |> split_list;  | 
| 63344 | 107  | 
val eqs = mk_def ctxt (xs ~~ rhss);  | 
| 20887 | 108  | 
val lhss = map (fst o Logic.dest_equals) eqs;  | 
| 18830 | 109  | 
in  | 
110  | 
ctxt  | 
|
| 42496 | 111  | 
|> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2  | 
| 20980 | 112  | 
|> fold Variable.declare_term eqs  | 
| 60377 | 113  | 
|> Proof_Context.add_assms def_export (map2 (fn b => fn eq => (b, [(eq, [])])) bs eqs)  | 
| 20887 | 114  | 
|>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss  | 
| 18830 | 115  | 
end;  | 
116  | 
||
| 18840 | 117  | 
|
| 25119 | 118  | 
(* fixed_abbrev *)  | 
119  | 
||
120  | 
fun fixed_abbrev ((x, mx), rhs) ctxt =  | 
|
121  | 
let  | 
|
122  | 
val T = Term.fastype_of rhs;  | 
|
123  | 
val ([x'], ctxt') = ctxt  | 
|
124  | 
|> Variable.declare_term rhs  | 
|
| 42360 | 125  | 
|> Proof_Context.add_fixes [(x, SOME T, mx)];  | 
| 25119 | 126  | 
val lhs = Free (x', T);  | 
| 63395 | 127  | 
val _ = cert_def ctxt' (K []) (Logic.mk_equals (lhs, rhs));  | 
| 74575 | 128  | 
fun abbrev_export _ _ = (I, Envir.expand_term_defs dest_Free [((x', T), rhs)]);  | 
| 25119 | 129  | 
val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';  | 
130  | 
in ((lhs, rhs), ctxt'') end;  | 
|
131  | 
||
132  | 
||
| 24985 | 133  | 
(* specific export -- result based on educated guessing *)  | 
| 21568 | 134  | 
|
| 24985 | 135  | 
(*  | 
| 67721 | 136  | 
[xs, xs \<equiv> as]  | 
| 24985 | 137  | 
:  | 
138  | 
B xs  | 
|
139  | 
--------------  | 
|
140  | 
B as  | 
|
141  | 
*)  | 
|
| 60823 | 142  | 
fun export inner outer th =  | 
| 21568 | 143  | 
let  | 
| 60823 | 144  | 
val defs_asms =  | 
145  | 
Assumption.local_assms_of inner outer  | 
|
146  | 
|> filter_out (Drule.is_sort_constraint o Thm.term_of)  | 
|
147  | 
|> map (Thm.assume #> (fn asm =>  | 
|
148  | 
(case try (head_of_def o Thm.prop_of) asm of  | 
|
149  | 
NONE => (asm, false)  | 
|
150  | 
| SOME x =>  | 
|
151  | 
let val t = Free x in  | 
|
152  | 
(case try (Assumption.export_term inner outer) t of  | 
|
153  | 
NONE => (asm, false)  | 
|
154  | 
| SOME u =>  | 
|
155  | 
if t aconv u then (asm, false)  | 
|
156  | 
else (Drule.abs_def (Variable.gen_all outer asm), true))  | 
|
157  | 
end)));  | 
|
158  | 
in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export false inner outer th) end;  | 
|
| 21568 | 159  | 
|
| 24985 | 160  | 
(*  | 
| 67721 | 161  | 
[xs, xs \<equiv> as]  | 
| 24985 | 162  | 
:  | 
163  | 
TERM b xs  | 
|
164  | 
-------------- and --------------  | 
|
| 67721 | 165  | 
TERM b as b xs \<equiv> b as  | 
| 24985 | 166  | 
*)  | 
167  | 
fun export_cterm inner outer ct =  | 
|
| 35717 | 168  | 
export inner outer (Drule.mk_term ct) ||> Drule.dest_term;  | 
| 24985 | 169  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
170  | 
fun contract ctxt defs ct th =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
171  | 
th COMP (Raw_Simplifier.rewrite ctxt true defs ct COMP_INCR Drule.equal_elim_rule2);  | 
| 21844 | 172  | 
|
| 18840 | 173  | 
|
| 
47237
 
b61a11dea74c
more precise Local_Defs.expand wrt. *local* prems only;
 
wenzelm 
parents: 
47060 
diff
changeset
 | 
174  | 
|
| 18840 | 175  | 
(** defived definitions **)  | 
176  | 
||
| 51585 | 177  | 
(* transformation via rewrite rules *)  | 
| 18840 | 178  | 
|
| 33519 | 179  | 
structure Rules = Generic_Data  | 
| 18840 | 180  | 
(  | 
181  | 
type T = thm list;  | 
|
| 33519 | 182  | 
val empty = [];  | 
183  | 
val merge = Thm.merge_thms;  | 
|
| 18840 | 184  | 
);  | 
185  | 
||
| 22846 | 186  | 
fun print_rules ctxt =  | 
| 51585 | 187  | 
Pretty.writeln (Pretty.big_list "definitional rewrite rules:"  | 
| 61268 | 188  | 
(map (Thm.pretty_thm_item ctxt) (Rules.get (Context.Proof ctxt))));  | 
| 18840 | 189  | 
|
| 61091 | 190  | 
val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm o Thm.trim_context);  | 
| 67627 | 191  | 
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);  | 
| 18840 | 192  | 
|
193  | 
||
| 18875 | 194  | 
(* meta rewrite rules *)  | 
| 18840 | 195  | 
|
| 23541 | 196  | 
fun meta_rewrite_conv ctxt =  | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
40242 
diff
changeset
 | 
197  | 
Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))  | 
| 
63221
 
7d43fbbaba28
avoid warnings on duplicate rules in the given list;
 
wenzelm 
parents: 
63169 
diff
changeset
 | 
198  | 
(ctxt  | 
| 
 
7d43fbbaba28
avoid warnings on duplicate rules in the given list;
 
wenzelm 
parents: 
63169 
diff
changeset
 | 
199  | 
|> Raw_Simplifier.init_simpset (Rules.get (Context.Proof ctxt))  | 
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45407 
diff
changeset
 | 
200  | 
|> Raw_Simplifier.add_eqcong Drule.equals_cong); (*protect meta-level equality*)  | 
| 18840 | 201  | 
|
| 23541 | 202  | 
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;  | 
| 18840 | 203  | 
|
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
204  | 
fun abs_def_rule ctxt = meta_rewrite_rule ctxt #> Drule.abs_def;  | 
| 
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
205  | 
|
| 18875 | 206  | 
|
| 63169 | 207  | 
(* unfold object-level rules *)  | 
| 18875 | 208  | 
|
| 69575 | 209  | 
val unfold_abs_def = Config.declare_bool ("unfold_abs_def", \<^here>) (K true);
 | 
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
210  | 
|
| 
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
211  | 
local  | 
| 
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
212  | 
|
| 63169 | 213  | 
fun gen_unfold rewrite ctxt rews =  | 
214  | 
let val meta_rews = map (meta_rewrite_rule ctxt) rews in  | 
|
215  | 
if Config.get ctxt unfold_abs_def then  | 
|
216  | 
rewrite ctxt meta_rews #>  | 
|
217  | 
rewrite ctxt (map (perhaps (try Drule.abs_def)) meta_rews)  | 
|
218  | 
else rewrite ctxt meta_rews  | 
|
219  | 
end;  | 
|
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
220  | 
|
| 63169 | 221  | 
val no_unfold_abs_def = Config.put unfold_abs_def false;  | 
| 18840 | 222  | 
|
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
223  | 
in  | 
| 
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
224  | 
|
| 63169 | 225  | 
val unfold = gen_unfold Raw_Simplifier.rewrite_rule;  | 
226  | 
val unfold_goals = gen_unfold Raw_Simplifier.rewrite_goals_rule;  | 
|
227  | 
val unfold_tac = PRIMITIVE oo unfold_goals;  | 
|
228  | 
||
229  | 
val unfold0 = unfold o no_unfold_abs_def;  | 
|
230  | 
val unfold0_goals = unfold_goals o no_unfold_abs_def;  | 
|
231  | 
val unfold0_tac = unfold_tac o no_unfold_abs_def;  | 
|
| 18840 | 232  | 
|
| 63169 | 233  | 
end  | 
234  | 
||
235  | 
||
236  | 
(* fold object-level rules *)  | 
|
237  | 
||
238  | 
fun fold ctxt rews = Raw_Simplifier.fold_rule ctxt (map (meta_rewrite_rule ctxt) rews);  | 
|
239  | 
fun fold_tac ctxt rews = Raw_Simplifier.fold_goals_tac ctxt (map (meta_rewrite_rule ctxt) rews);  | 
|
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
63042 
diff
changeset
 | 
240  | 
|
| 18840 | 241  | 
|
242  | 
(* derived defs -- potentially within the object-logic *)  | 
|
243  | 
||
| 63395 | 244  | 
fun derived_def ctxt get_pos {conditional} prop =
 | 
| 18840 | 245  | 
let  | 
246  | 
val ((c, T), rhs) = prop  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59573 
diff
changeset
 | 
247  | 
|> Thm.cterm_of ctxt  | 
| 23541 | 248  | 
|> meta_rewrite_conv ctxt  | 
| 18840 | 249  | 
|> (snd o Logic.dest_equals o Thm.prop_of)  | 
| 21568 | 250  | 
|> conditional ? Logic.strip_imp_concl  | 
| 63395 | 251  | 
|> (abs_def o #2 o cert_def ctxt get_pos);  | 
| 
70306
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
252  | 
fun prove def_ctxt0 def =  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
253  | 
let  | 
| 70308 | 254  | 
val def_ctxt = Proof_Context.augment prop def_ctxt0;  | 
| 
70306
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
255  | 
val def_thm =  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
256  | 
Goal.prove def_ctxt [] [] prop  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
257  | 
            (fn {context = goal_ctxt, ...} =>
 | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
258  | 
ALLGOALS  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
259  | 
(CONVERSION (meta_rewrite_conv goal_ctxt) THEN'  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
260  | 
rewrite_goal_tac goal_ctxt [def] THEN'  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
261  | 
resolve_tac goal_ctxt [Drule.reflexive_thm]))  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
262  | 
handle ERROR msg => cat_error msg "Failed to prove definitional specification";  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
263  | 
in  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
264  | 
def_thm  | 
| 
70314
 
6d6839a948cf
proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
 
wenzelm 
parents: 
70308 
diff
changeset
 | 
265  | 
|> singleton (Proof_Context.export def_ctxt def_ctxt0)  | 
| 
70306
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
266  | 
|> Drule.zero_var_indexes  | 
| 
 
61621dc439d4
clarified context: prefer abstract Variable.auto_fixes;
 
wenzelm 
parents: 
70305 
diff
changeset
 | 
267  | 
end;  | 
| 18840 | 268  | 
in (((c, T), rhs), prove) end;  | 
269  | 
||
| 18830 | 270  | 
end;  |