| author | wenzelm | 
| Tue, 31 Dec 2024 15:09:36 +0100 | |
| changeset 81697 | 1629c2ff4880 | 
| parent 78804 | d4e9d6b7f48d | 
| child 81960 | 326ecfc079a6 | 
| permissions | -rw-r--r-- | 
| 74593 | 1  | 
(* Title: Pure/ML/ml_instantiate.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
ML antiquotation to instantiate logical entities.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature ML_INSTANTIATE =  | 
|
8  | 
sig  | 
|
9  | 
val make_ctyp: Proof.context -> typ -> ctyp  | 
|
10  | 
val make_cterm: Proof.context -> term -> cterm  | 
|
11  | 
type insts = ((indexname * sort) * typ) list * ((indexname * typ) * term) list  | 
|
12  | 
type cinsts = ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list  | 
|
13  | 
val instantiate_typ: insts -> typ -> typ  | 
|
14  | 
val instantiate_ctyp: Position.T -> cinsts -> ctyp -> ctyp  | 
|
15  | 
val instantiate_term: insts -> term -> term  | 
|
16  | 
val instantiate_cterm: Position.T -> cinsts -> cterm -> cterm  | 
|
| 74606 | 17  | 
val instantiate_thm: Position.T -> cinsts -> thm -> thm  | 
18  | 
val instantiate_thms: Position.T -> cinsts -> thm list -> thm list  | 
|
19  | 
val get_thms: Proof.context -> int -> thm list  | 
|
20  | 
val get_thm: Proof.context -> int -> thm  | 
|
| 74593 | 21  | 
end;  | 
22  | 
||
23  | 
structure ML_Instantiate: ML_INSTANTIATE =  | 
|
24  | 
struct  | 
|
25  | 
||
26  | 
(* exported operations *)  | 
|
27  | 
||
28  | 
fun make_ctyp ctxt T = Thm.ctyp_of ctxt T |> Thm.trim_context_ctyp;  | 
|
29  | 
fun make_cterm ctxt t = Thm.cterm_of ctxt t |> Thm.trim_context_cterm;  | 
|
30  | 
||
31  | 
type insts = ((indexname * sort) * typ) list * ((indexname * typ) * term) list  | 
|
32  | 
type cinsts = ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list  | 
|
33  | 
||
34  | 
fun instantiate_typ (insts: insts) =  | 
|
35  | 
Term_Subst.instantiateT (TVars.make (#1 insts));  | 
|
36  | 
||
37  | 
fun instantiate_ctyp pos (cinsts: cinsts) cT =  | 
|
38  | 
Thm.instantiate_ctyp (TVars.make (#1 cinsts)) cT  | 
|
39  | 
handle CTERM (msg, args) => Exn.reraise (CTERM (msg ^ Position.here pos, args));  | 
|
40  | 
||
41  | 
fun instantiate_term (insts: insts) =  | 
|
42  | 
let  | 
|
43  | 
val instT = TVars.make (#1 insts);  | 
|
44  | 
val instantiateT = Term_Subst.instantiateT instT;  | 
|
45  | 
val inst = Vars.make ((map o apfst o apsnd) instantiateT (#2 insts));  | 
|
46  | 
in Term_Subst.instantiate_beta (instT, inst) end;  | 
|
47  | 
||
| 74606 | 48  | 
fun make_cinsts (cinsts: cinsts) =  | 
| 74593 | 49  | 
let  | 
50  | 
val cinstT = TVars.make (#1 cinsts);  | 
|
51  | 
val instantiateT = Term_Subst.instantiateT (TVars.map (K Thm.typ_of) cinstT);  | 
|
52  | 
val cinst = Vars.make ((map o apfst o apsnd) instantiateT (#2 cinsts));  | 
|
| 74606 | 53  | 
in (cinstT, cinst) end;  | 
54  | 
||
55  | 
fun instantiate_cterm pos cinsts ct =  | 
|
56  | 
Thm.instantiate_beta_cterm (make_cinsts cinsts) ct  | 
|
| 74593 | 57  | 
handle CTERM (msg, args) => Exn.reraise (CTERM (msg ^ Position.here pos, args));  | 
58  | 
||
| 74606 | 59  | 
fun instantiate_thm pos cinsts th =  | 
60  | 
Thm.instantiate_beta (make_cinsts cinsts) th  | 
|
61  | 
handle THM (msg, i, args) => Exn.reraise (THM (msg ^ Position.here pos, i, args));  | 
|
62  | 
||
63  | 
fun instantiate_thms pos cinsts = map (instantiate_thm pos cinsts);  | 
|
64  | 
||
65  | 
||
66  | 
(* context data *)  | 
|
67  | 
||
68  | 
structure Data = Proof_Data  | 
|
69  | 
(  | 
|
70  | 
type T = int * thm list Inttab.table;  | 
|
71  | 
fun init _ = (0, Inttab.empty);  | 
|
72  | 
);  | 
|
73  | 
||
74  | 
fun put_thms ths ctxt =  | 
|
75  | 
let  | 
|
76  | 
val (i, thms) = Data.get ctxt;  | 
|
| 
74620
 
d622d1dce05c
avoid persistence of static context: instantiation arguments should provide proper dynamic context (see also e2e2bc1f9570);
 
wenzelm 
parents: 
74606 
diff
changeset
 | 
77  | 
val ctxt' = ctxt |> Data.put (i + 1, Inttab.update (i, map Thm.trim_context ths) thms);  | 
| 74606 | 78  | 
in (i, ctxt') end;  | 
79  | 
||
80  | 
fun get_thms ctxt i = the (Inttab.lookup (#2 (Data.get ctxt)) i);  | 
|
81  | 
fun get_thm ctxt i = the_single (get_thms ctxt i);  | 
|
82  | 
||
| 74593 | 83  | 
|
84  | 
(* ML antiquotation *)  | 
|
85  | 
||
86  | 
local  | 
|
87  | 
||
| 74605 | 88  | 
val make_keywords =  | 
89  | 
Thy_Header.get_keywords'  | 
|
90  | 
#> Keyword.no_major_keywords  | 
|
| 74606 | 91  | 
#> Keyword.add_major_keywords ["typ", "term", "prop", "ctyp", "cterm", "cprop", "lemma"];  | 
| 74593 | 92  | 
|
93  | 
val parse_inst_name = Parse.position (Parse.type_ident >> pair true || Parse.name >> pair false);  | 
|
94  | 
||
95  | 
val parse_inst =  | 
|
96  | 
(parse_inst_name -- (Parse.$$$ "=" |-- Parse.!!! Parse.embedded_ml) ||  | 
|
97  | 
Scan.ahead parse_inst_name -- Parse.embedded_ml)  | 
|
98  | 
>> (fn (((b, a), pos), ml) => (b, ((a, pos), ml)));  | 
|
99  | 
||
100  | 
val parse_insts =  | 
|
101  | 
Parse.and_list1 parse_inst >> (List.partition #1 #> apply2 (map #2));  | 
|
102  | 
||
103  | 
val ml = ML_Lex.tokenize_no_range;  | 
|
104  | 
val ml_range = ML_Lex.tokenize_range;  | 
|
105  | 
fun ml_parens x = ml "(" @ x @ ml ")";
 | 
|
106  | 
fun ml_bracks x = ml "[" @ x @ ml "]";  | 
|
107  | 
fun ml_commas xs = flat (separate (ml ", ") xs);  | 
|
108  | 
val ml_list = ml_bracks o ml_commas;  | 
|
109  | 
fun ml_pair (x, y) = ml_parens (ml_commas [x, y]);  | 
|
110  | 
val ml_list_pair = ml_list o ListPair.map ml_pair;  | 
|
| 74606 | 111  | 
val ml_here = ML_Syntax.atomic o ML_Syntax.print_position;  | 
| 74593 | 112  | 
|
113  | 
fun get_tfree envT (a, pos) =  | 
|
114  | 
(case AList.lookup (op =) envT a of  | 
|
115  | 
SOME S => (a, S)  | 
|
116  | 
  | NONE => error ("No occurrence of type variable " ^ quote a ^ Position.here pos));
 | 
|
117  | 
||
118  | 
fun check_free ctxt env (x, pos) =  | 
|
119  | 
(case AList.lookup (op =) env x of  | 
|
120  | 
SOME T =>  | 
|
121  | 
(Context_Position.reports ctxt (map (pair pos) (Syntax_Phases.markup_free ctxt x)); (x, T))  | 
|
122  | 
  | NONE => error ("No occurrence of variable " ^ quote x ^ Position.here pos));
 | 
|
123  | 
||
| 74606 | 124  | 
fun missing_instT pos envT instT =  | 
| 74593 | 125  | 
(case filter_out (fn (a, _) => exists (fn (b, _) => a = b) instT) envT of  | 
126  | 
[] => ()  | 
|
| 74606 | 127  | 
| bad =>  | 
128  | 
      error ("No instantiation for free type variable(s) " ^ commas_quote (map #1 bad) ^
 | 
|
129  | 
Position.here pos));  | 
|
| 74593 | 130  | 
|
| 74606 | 131  | 
fun missing_inst pos env inst =  | 
| 74593 | 132  | 
(case filter_out (fn (a, _) => exists (fn (b, _) => a = b) inst) env of  | 
133  | 
[] => ()  | 
|
| 74606 | 134  | 
| bad =>  | 
135  | 
      error ("No instantiation for free variable(s) " ^ commas_quote (map #1 bad) ^
 | 
|
136  | 
Position.here pos));  | 
|
| 74593 | 137  | 
|
138  | 
fun make_instT (a, pos) T =  | 
|
139  | 
(case try (Term.dest_TVar o Logic.dest_type) T of  | 
|
140  | 
    NONE => error ("Not a free type variable " ^ quote a ^ Position.here pos)
 | 
|
141  | 
| SOME v => ml (ML_Syntax.print_pair ML_Syntax.print_indexname ML_Syntax.print_sort v));  | 
|
142  | 
||
143  | 
fun make_inst (a, pos) t =  | 
|
144  | 
(case try Term.dest_Var t of  | 
|
145  | 
    NONE => error ("Not a free variable " ^ quote a ^ Position.here pos)
 | 
|
146  | 
| SOME v => ml (ML_Syntax.print_pair ML_Syntax.print_indexname ML_Syntax.print_typ v));  | 
|
147  | 
||
| 74606 | 148  | 
fun make_env ts =  | 
149  | 
let  | 
|
150  | 
val envT = fold Term.add_tfrees ts [];  | 
|
151  | 
val env = fold Term.add_frees ts [];  | 
|
152  | 
in (envT, env) end;  | 
|
| 74593 | 153  | 
|
| 74606 | 154  | 
fun prepare_insts pos {schematic} ctxt1 ctxt0 (instT, inst) ts =
 | 
| 74593 | 155  | 
let  | 
| 74606 | 156  | 
val (envT, env) = make_env ts;  | 
| 74593 | 157  | 
val freesT = map (Logic.mk_type o TFree o get_tfree envT) instT;  | 
158  | 
val frees = map (Free o check_free ctxt1 env) inst;  | 
|
| 74606 | 159  | 
val (ts', (varsT, vars)) =  | 
160  | 
Variable.export_terms ctxt1 ctxt0 (ts @ freesT @ frees)  | 
|
161  | 
|> chop (length ts) ||> chop (length freesT);  | 
|
162  | 
val ml_insts = (map2 make_instT instT varsT, map2 make_inst inst vars);  | 
|
163  | 
in  | 
|
164  | 
if schematic then ()  | 
|
165  | 
else  | 
|
166  | 
let val (envT', env') = make_env ts' in  | 
|
167  | 
missing_instT pos (subtract (eq_fst op =) envT' envT) instT;  | 
|
168  | 
missing_inst pos (subtract (eq_fst op =) env' env) inst  | 
|
169  | 
end;  | 
|
170  | 
(ml_insts, ts')  | 
|
171  | 
end;  | 
|
| 74593 | 172  | 
|
| 74606 | 173  | 
fun prepare_ml range kind ml1 ml2 ml_val (ml_instT, ml_inst) ctxt =  | 
| 74593 | 174  | 
let  | 
175  | 
val (ml_name, ctxt') = ML_Context.variant kind ctxt;  | 
|
176  | 
    val ml_env = ml ("val " ^ ml_name ^ " = ") @ ml ml1 @ ml_parens (ml ml_val) @ ml ";\n";
 | 
|
177  | 
fun ml_body (ml_argsT, ml_args) =  | 
|
178  | 
ml_parens (ml ml2 @  | 
|
179  | 
ml_pair (ml_list_pair (ml_instT, ml_argsT), ml_list_pair (ml_inst, ml_args)) @  | 
|
180  | 
ml_range range (ML_Context.struct_name ctxt ^ "." ^ ml_name));  | 
|
181  | 
in ((ml_env, ml_body), ctxt') end;  | 
|
182  | 
||
| 74606 | 183  | 
fun prepare_type range ((((kind, pos), ml1, ml2), schematic), s) insts ctxt =  | 
| 74593 | 184  | 
let  | 
185  | 
val T = Syntax.read_typ ctxt s;  | 
|
186  | 
val t = Logic.mk_type T;  | 
|
187  | 
val ctxt1 = Proof_Context.augment t ctxt;  | 
|
| 74606 | 188  | 
val (ml_insts, T') =  | 
189  | 
prepare_insts pos schematic ctxt1 ctxt insts [t] ||> (the_single #> Logic.dest_type);  | 
|
190  | 
in prepare_ml range kind ml1 ml2 (ML_Syntax.print_typ T') ml_insts ctxt end;  | 
|
| 74593 | 191  | 
|
| 74606 | 192  | 
fun prepare_term read range ((((kind, pos), ml1, ml2), schematic), (s, fixes)) insts ctxt =  | 
| 74593 | 193  | 
let  | 
194  | 
val ctxt' = #2 (Proof_Context.add_fixes_cmd fixes ctxt);  | 
|
195  | 
val t = read ctxt' s;  | 
|
196  | 
val ctxt1 = Proof_Context.augment t ctxt';  | 
|
| 74606 | 197  | 
val (ml_insts, t') = prepare_insts pos schematic ctxt1 ctxt insts [t] ||> the_single;  | 
198  | 
in prepare_ml range kind ml1 ml2 (ML_Syntax.print_term t') ml_insts ctxt end;  | 
|
| 74593 | 199  | 
|
| 74606 | 200  | 
fun prepare_lemma range ((pos, schematic), make_lemma) insts ctxt =  | 
201  | 
let  | 
|
202  | 
val (ths, (props, ctxt1)) = make_lemma ctxt  | 
|
203  | 
val (i, thms_ctxt) = put_thms ths ctxt;  | 
|
204  | 
val ml_insts = #1 (prepare_insts pos schematic ctxt1 ctxt insts props);  | 
|
205  | 
val (ml1, ml2) =  | 
|
206  | 
if length ths = 1  | 
|
207  | 
      then ("ML_Instantiate.get_thm ML_context", "ML_Instantiate.instantiate_thm " ^ ml_here pos)
 | 
|
208  | 
      else ("ML_Instantiate.get_thms ML_context", "ML_Instantiate.instantiate_thms " ^ ml_here pos);
 | 
|
209  | 
in prepare_ml range "lemma" ml1 ml2 (ML_Syntax.print_int i) ml_insts thms_ctxt end;  | 
|
210  | 
||
211  | 
fun typ_ml (kind, pos: Position.T) = ((kind, pos), "", "ML_Instantiate.instantiate_typ ");  | 
|
212  | 
fun term_ml (kind, pos: Position.T) = ((kind, pos), "", "ML_Instantiate.instantiate_term ");  | 
|
| 74593 | 213  | 
fun ctyp_ml (kind, pos) =  | 
| 74606 | 214  | 
((kind, pos),  | 
215  | 
"ML_Instantiate.make_ctyp ML_context", "ML_Instantiate.instantiate_ctyp " ^ ml_here pos);  | 
|
| 74593 | 216  | 
fun cterm_ml (kind, pos) =  | 
| 74606 | 217  | 
((kind, pos),  | 
218  | 
"ML_Instantiate.make_cterm ML_context", "ML_Instantiate.instantiate_cterm " ^ ml_here pos);  | 
|
| 74593 | 219  | 
|
220  | 
val command_name = Parse.position o Parse.command_name;  | 
|
221  | 
||
| 74606 | 222  | 
val parse_schematic = Args.mode "schematic" >> (fn b => {schematic = b});
 | 
223  | 
||
| 74593 | 224  | 
fun parse_body range =  | 
| 74606 | 225  | 
(command_name "typ" >> typ_ml || command_name "ctyp" >> ctyp_ml) -- parse_schematic --  | 
| 74593 | 226  | 
Parse.!!! Parse.typ >> prepare_type range ||  | 
| 74606 | 227  | 
(command_name "term" >> term_ml || command_name "cterm" >> cterm_ml) -- parse_schematic --  | 
228  | 
Parse.!!! (Parse.term -- Parse.for_fixes) >> prepare_term Syntax.read_term range ||  | 
|
229  | 
(command_name "prop" >> term_ml || command_name "cprop" >> cterm_ml) -- parse_schematic --  | 
|
230  | 
Parse.!!! (Parse.term -- Parse.for_fixes) >> prepare_term Syntax.read_prop range ||  | 
|
231  | 
(command_name "lemma" >> #2) -- parse_schematic -- ML_Thms.embedded_lemma >> prepare_lemma range;  | 
|
| 74593 | 232  | 
|
233  | 
val _ = Theory.setup  | 
|
| 78804 | 234  | 
(ML_Context.add_antiquotation_embedded \<^binding>\<open>instantiate\<close>  | 
235  | 
(fn range => fn input => fn ctxt =>  | 
|
| 74593 | 236  | 
let  | 
| 78804 | 237  | 
val (insts, prepare_val) = input  | 
238  | 
|> Parse.read_embedded ctxt (make_keywords ctxt)  | 
|
| 74593 | 239  | 
((parse_insts --| Parse.$$$ "in") -- parse_body range);  | 
240  | 
||
241  | 
val (((ml_env, ml_body), decls), ctxt1) = ctxt  | 
|
242  | 
|> prepare_val (apply2 (map #1) insts)  | 
|
243  | 
||>> ML_Context.expand_antiquotes_list (op @ (apply2 (map #2) insts));  | 
|
244  | 
||
245  | 
fun decl' ctxt' =  | 
|
246  | 
let val (ml_args_env, ml_args) = split_list (decls ctxt')  | 
|
247  | 
in (ml_env @ flat ml_args_env, ml_body (chop (length (#1 insts)) ml_args)) end;  | 
|
248  | 
in (decl', ctxt1) end));  | 
|
249  | 
||
250  | 
in end;  | 
|
251  | 
||
252  | 
end;  |