| author | wenzelm | 
| Wed, 06 Dec 2023 21:28:12 +0100 | |
| changeset 79158 | 3c7ab17380a8 | 
| parent 77879 | dd222e2af01a | 
| child 79232 | 99bc2dd45111 | 
| permissions | -rw-r--r-- | 
| 53707 | 1  | 
(* Title: Pure/Tools/rule_insts.ML  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 53708 | 4  | 
Rule instantiations -- operations within implicit rule / subgoal context.  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 59763 | 7  | 
signature RULE_INSTS =  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 59763 | 9  | 
val where_rule: Proof.context ->  | 
10  | 
((indexname * Position.T) * string) list ->  | 
|
11  | 
(binding * string option * mixfix) list -> thm -> thm  | 
|
12  | 
val of_rule: Proof.context -> string option list * string option list ->  | 
|
13  | 
(binding * string option * mixfix) list -> thm -> thm  | 
|
14  | 
val read_instantiate: Proof.context ->  | 
|
15  | 
((indexname * Position.T) * string) list -> string list -> thm -> thm  | 
|
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
16  | 
val read_term: string -> Proof.context -> term * Proof.context  | 
| 59829 | 17  | 
val goal_context: term -> Proof.context -> (string * typ) list * Proof.context  | 
| 59755 | 18  | 
val res_inst_tac: Proof.context ->  | 
| 59855 | 19  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
20  | 
thm -> int -> tactic  | 
|
| 59755 | 21  | 
val eres_inst_tac: Proof.context ->  | 
| 59855 | 22  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
23  | 
thm -> int -> tactic  | 
|
| 59755 | 24  | 
val cut_inst_tac: Proof.context ->  | 
| 59855 | 25  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
26  | 
thm -> int -> tactic  | 
|
| 59755 | 27  | 
val forw_inst_tac: Proof.context ->  | 
| 59855 | 28  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
29  | 
thm -> int -> tactic  | 
|
| 59755 | 30  | 
val dres_inst_tac: Proof.context ->  | 
| 59855 | 31  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
32  | 
thm -> int -> tactic  | 
|
| 59780 | 33  | 
val thin_tac: Proof.context -> string -> (binding * string option * mixfix) list ->  | 
34  | 
int -> tactic  | 
|
35  | 
val subgoal_tac: Proof.context -> string -> (binding * string option * mixfix) list ->  | 
|
36  | 
int -> tactic  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58027 
diff
changeset
 | 
37  | 
val make_elim_preserve: Proof.context -> thm -> thm  | 
| 59755 | 38  | 
val method:  | 
| 59780 | 39  | 
(Proof.context -> ((indexname * Position.T) * string) list ->  | 
40  | 
(binding * string option * mixfix) list -> thm -> int -> tactic) ->  | 
|
| 59855 | 41  | 
(Proof.context -> thm list -> int -> tactic) ->  | 
42  | 
(Proof.context -> Proof.method) context_parser  | 
|
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
end;  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
|
| 42806 | 45  | 
structure Rule_Insts: RULE_INSTS =  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
struct  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
48  | 
(** read instantiations **)  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
50  | 
local  | 
| 59755 | 51  | 
|
52  | 
fun error_var msg (xi, pos) =  | 
|
53  | 
error (msg ^ quote (Term.string_of_vname xi) ^ Position.here pos);  | 
|
| 59754 | 54  | 
|
| 74229 | 55  | 
fun the_sort tvars (ai, pos) : sort =  | 
| 74269 | 56  | 
(case TVars.get_first (fn ((bi, S), _) => if ai = bi then SOME S else NONE) tvars of  | 
| 45611 | 57  | 
SOME S => S  | 
| 74229 | 58  | 
| NONE => error_var "No such type variable in theorem: " (ai, pos));  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
|
| 59755 | 60  | 
fun the_type vars (xi, pos) : typ =  | 
| 74229 | 61  | 
(case Vartab.lookup vars xi of  | 
| 45611 | 62  | 
SOME T => T  | 
| 59755 | 63  | 
| NONE => error_var "No such variable in theorem: " (xi, pos));  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
|
| 59774 | 65  | 
fun read_type ctxt tvars ((xi, pos), s) =  | 
| 59759 | 66  | 
let  | 
67  | 
val S = the_sort tvars (xi, pos);  | 
|
68  | 
val T = Syntax.read_typ ctxt s;  | 
|
69  | 
in  | 
|
70  | 
if Sign.of_sort (Proof_Context.theory_of ctxt) (T, S) then ((xi, S), T)  | 
|
| 59768 | 71  | 
else error_var "Bad sort for instantiation of type variable: " (xi, pos)  | 
| 59759 | 72  | 
end;  | 
73  | 
||
| 74266 | 74  | 
fun make_instT f (tvars: TVars.set) =  | 
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
75  | 
let  | 
| 74229 | 76  | 
fun add v =  | 
77  | 
let  | 
|
78  | 
val T = TVar v;  | 
|
79  | 
val T' = f T;  | 
|
80  | 
in if T = T' then I else cons (v, T') end;  | 
|
| 74266 | 81  | 
in TVars.fold (add o #1) tvars [] end;  | 
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
82  | 
|
| 74229 | 83  | 
fun make_inst f vars =  | 
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
84  | 
let  | 
| 74229 | 85  | 
fun add v =  | 
86  | 
let  | 
|
87  | 
val t = Var v;  | 
|
88  | 
val t' = f t;  | 
|
89  | 
in if t aconv t' then I else cons (v, t') end;  | 
|
90  | 
in fold add vars [] end;  | 
|
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
91  | 
|
| 59796 | 92  | 
fun read_terms ss Ts ctxt =  | 
| 25329 | 93  | 
let  | 
94  | 
fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt;  | 
|
| 59796 | 95  | 
val (ts, ctxt') = fold_map Variable.fix_dummy_patterns (map2 parse Ts ss) ctxt;  | 
| 25329 | 96  | 
val ts' =  | 
| 39288 | 97  | 
map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts  | 
| 59796 | 98  | 
|> Syntax.check_terms ctxt'  | 
99  | 
|> Variable.polymorphic ctxt';  | 
|
| 25329 | 100  | 
val Ts' = map Term.fastype_of ts';  | 
| 74232 | 101  | 
val tyenv = Vartab.build (fold (Sign.typ_match (Proof_Context.theory_of ctxt)) (Ts ~~ Ts'));  | 
| 59759 | 102  | 
val tyenv' = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), T)) tyenv [];  | 
| 59796 | 103  | 
in ((ts', tyenv'), ctxt') end;  | 
| 25329 | 104  | 
|
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
105  | 
in  | 
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
106  | 
|
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
107  | 
fun read_term s ctxt =  | 
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
108  | 
let  | 
| 
59827
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
109  | 
val (t, ctxt') = Variable.fix_dummy_patterns (Syntax.parse_term ctxt s) ctxt;  | 
| 
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
110  | 
val t' = Syntax.check_term ctxt' t;  | 
| 
 
04e569577c18
proper Rule_Insts.read_term, e.g. to enable case_tac using "_";
 
wenzelm 
parents: 
59825 
diff
changeset
 | 
111  | 
in (t', ctxt') end;  | 
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
112  | 
|
| 59853 | 113  | 
fun read_insts thm raw_insts raw_fixes ctxt =  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
let  | 
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
115  | 
val (type_insts, term_insts) =  | 
| 59853 | 116  | 
List.partition (fn (((x, _), _), _) => String.isPrefix "'" x) raw_insts;  | 
| 59768 | 117  | 
|
| 74266 | 118  | 
    val tvars = TVars.build (Thm.fold_terms {hyps = false} TVars.add_tvars thm);
 | 
119  | 
    val vars = Vars.build (Thm.fold_terms {hyps = false} Vars.add_vars thm);
 | 
|
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
|
| 59853 | 121  | 
(*eigen-context*)  | 
| 60469 | 122  | 
val (_, ctxt1) = ctxt  | 
| 74266 | 123  | 
|> TVars.fold (Variable.declare_internal o Logic.mk_type o TVar o #1) tvars  | 
124  | 
|> Vars.fold (Variable.declare_internal o Var o #1) vars  | 
|
| 60469 | 125  | 
|> Proof_Context.add_fixes_cmd raw_fixes;  | 
| 59853 | 126  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
127  | 
(*explicit type instantiations*)  | 
| 
74220
 
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
 
wenzelm 
parents: 
67147 
diff
changeset
 | 
128  | 
val instT1 =  | 
| 74266 | 129  | 
Term_Subst.instantiateT (TVars.make (map (read_type ctxt1 tvars) type_insts));  | 
| 74229 | 130  | 
val vars1 =  | 
| 74269 | 131  | 
Vartab.build (vars |> Vars.fold (fn ((v, T), _) =>  | 
| 74232 | 132  | 
Vartab.insert (K true) (v, instT1 T)));  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
134  | 
(*term instantiations*)  | 
| 
45613
 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 
wenzelm 
parents: 
45611 
diff
changeset
 | 
135  | 
val (xs, ss) = split_list term_insts;  | 
| 
 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 
wenzelm 
parents: 
45611 
diff
changeset
 | 
136  | 
val Ts = map (the_type vars1) xs;  | 
| 59825 | 137  | 
val ((ts, inferred), ctxt2) = read_terms ss Ts ctxt1;  | 
| 
45613
 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 
wenzelm 
parents: 
45611 
diff
changeset
 | 
138  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
139  | 
(*implicit type instantiations*)  | 
| 74266 | 140  | 
val instT2 = Term_Subst.instantiateT (TVars.make inferred);  | 
| 74229 | 141  | 
val vars2 = Vartab.fold (fn (v, T) => cons (v, instT2 T)) vars1 [];  | 
| 59774 | 142  | 
val inst2 =  | 
| 74266 | 143  | 
Term_Subst.instantiate (TVars.empty,  | 
144  | 
Vars.build (fold2 (fn (xi, _) => fn t => Vars.add ((xi, Term.fastype_of t), t)) xs ts))  | 
|
| 59774 | 145  | 
#> Envir.beta_norm;  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
|
| 74229 | 147  | 
val inst_tvars = make_instT (instT2 o instT1) tvars;  | 
148  | 
val inst_vars = make_inst inst2 vars2;  | 
|
| 59825 | 149  | 
in ((inst_tvars, inst_vars), ctxt2) end;  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
150  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
151  | 
end;  | 
| 
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
152  | 
|
| 
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
153  | 
|
| 
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
154  | 
|
| 
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
155  | 
(** forward rules **)  | 
| 
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
156  | 
|
| 59853 | 157  | 
fun where_rule ctxt raw_insts raw_fixes thm =  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
158  | 
let  | 
| 59853 | 159  | 
val ((inst_tvars, inst_vars), ctxt') = read_insts thm raw_insts raw_fixes ctxt;  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
160  | 
in  | 
| 59769 | 161  | 
thm  | 
162  | 
|> Drule.instantiate_normalize  | 
|
| 74282 | 163  | 
(TVars.make (map (apsnd (Thm.ctyp_of ctxt')) inst_tvars),  | 
164  | 
Vars.make (map (apsnd (Thm.cterm_of ctxt')) inst_vars))  | 
|
| 59853 | 165  | 
|> singleton (Variable.export ctxt' ctxt)  | 
| 
45613
 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 
wenzelm 
parents: 
45611 
diff
changeset
 | 
166  | 
|> Rule_Cases.save thm  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
167  | 
end;  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
|
| 
55143
 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 
wenzelm 
parents: 
55111 
diff
changeset
 | 
169  | 
fun of_rule ctxt (args, concl_args) fixes thm =  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
170  | 
let  | 
| 
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
171  | 
fun zip_vars _ [] = []  | 
| 
45613
 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 
wenzelm 
parents: 
45611 
diff
changeset
 | 
172  | 
| zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest  | 
| 59755 | 173  | 
| zip_vars ((x, _) :: xs) (SOME t :: rest) = ((x, Position.none), t) :: zip_vars xs rest  | 
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
174  | 
| zip_vars [] _ = error "More instantiations than variables in theorem";  | 
| 
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
175  | 
val insts =  | 
| 74281 | 176  | 
zip_vars (Vars.build (Vars.add_vars (Thm.full_prop_of thm)) |> Vars.list_set) args @  | 
177  | 
zip_vars (Vars.build (Vars.add_vars (Thm.concl_of thm)) |> Vars.list_set) concl_args;  | 
|
| 
55143
 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 
wenzelm 
parents: 
55111 
diff
changeset
 | 
178  | 
in where_rule ctxt insts fixes thm end;  | 
| 27236 | 179  | 
|
| 
55143
 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 
wenzelm 
parents: 
55111 
diff
changeset
 | 
180  | 
fun read_instantiate ctxt insts xs =  | 
| 
 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 
wenzelm 
parents: 
55111 
diff
changeset
 | 
181  | 
where_rule ctxt insts (map (fn x => (Binding.name x, NONE, NoSyn)) xs);  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
183  | 
|
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
184  | 
|
| 
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
185  | 
(** attributes **)  | 
| 
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
186  | 
|
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
(* where: named instantiation *)  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
|
| 59855 | 189  | 
val named_insts =  | 
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
62012 
diff
changeset
 | 
190  | 
Parse.and_list1  | 
| 74563 | 191  | 
(Parse.position Args.var -- (Args.$$$ "=" |-- Parse.!!! Parse.embedded_inner_syntax))  | 
| 59853 | 192  | 
-- Parse.for_fixes;  | 
193  | 
||
| 53171 | 194  | 
val _ = Theory.setup  | 
| 67147 | 195  | 
(Attrib.setup \<^binding>\<open>where\<close>  | 
| 59855 | 196  | 
(Scan.lift named_insts >> (fn args =>  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
197  | 
Thm.rule_attribute [] (fn context => uncurry (where_rule (Context.proof_of context)) args)))  | 
| 53171 | 198  | 
"named instantiation of theorem");  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
|
| 
20343
 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 
wenzelm 
parents: 
20336 
diff
changeset
 | 
201  | 
(* of: positional instantiation (terms only) *)  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
local  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
|
| 74563 | 205  | 
val inst = Args.maybe Parse.embedded_inner_syntax;  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
val concl = Args.$$$ "concl" -- Args.colon;  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
val insts =  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
Scan.repeat (Scan.unless concl inst) --  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
Scan.optional (concl |-- Scan.repeat inst) [];  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
in  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
|
| 53171 | 214  | 
val _ = Theory.setup  | 
| 67147 | 215  | 
(Attrib.setup \<^binding>\<open>of\<close>  | 
| 59855 | 216  | 
(Scan.lift (insts -- Parse.for_fixes) >> (fn args =>  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
217  | 
Thm.rule_attribute [] (fn context => uncurry (of_rule (Context.proof_of context)) args)))  | 
| 53171 | 218  | 
"positional instantiation of theorem");  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
end;  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
224  | 
(** tactics **)  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 59825 | 226  | 
(* goal context *)  | 
227  | 
||
| 59829 | 228  | 
fun goal_context goal ctxt =  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
229  | 
let  | 
| 
59828
 
0e9baaf0e0bb
prefer Variable.focus, despite subtle differences of Logic.strip_params vs. Term.strip_all_vars;
 
wenzelm 
parents: 
59827 
diff
changeset
 | 
230  | 
val ((_, params), ctxt') = ctxt  | 
| 59829 | 231  | 
|> Variable.declare_constraints goal  | 
| 59790 | 232  | 
|> Variable.improper_fixes  | 
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
233  | 
|> Variable.focus_params NONE goal  | 
| 60331 | 234  | 
||> Variable.restore_proper_fixes ctxt;  | 
| 
59828
 
0e9baaf0e0bb
prefer Variable.focus, despite subtle differences of Logic.strip_params vs. Term.strip_all_vars;
 
wenzelm 
parents: 
59827 
diff
changeset
 | 
235  | 
in (params, ctxt') end;  | 
| 59825 | 236  | 
|
237  | 
||
238  | 
(* resolution after lifting and instantiation; may refer to parameters of the subgoal *)  | 
|
239  | 
||
| 59853 | 240  | 
fun bires_inst_tac bires_flag ctxt raw_insts raw_fixes thm i st = CSUBGOAL (fn (cgoal, _) =>  | 
| 59825 | 241  | 
let  | 
| 59855 | 242  | 
(*goal context*)  | 
| 59834 | 243  | 
val (params, goal_ctxt) = goal_context (Thm.term_of cgoal) ctxt;  | 
| 59825 | 244  | 
val paramTs = map #2 params;  | 
| 59754 | 245  | 
|
| 59855 | 246  | 
(*instantiation context*)  | 
| 59853 | 247  | 
val ((inst_tvars, inst_vars), inst_ctxt) = read_insts thm raw_insts raw_fixes goal_ctxt;  | 
| 59846 | 248  | 
val fixed = map #1 (fold (Variable.add_newly_fixed inst_ctxt goal_ctxt o #2) inst_vars []);  | 
| 59805 | 249  | 
|
| 59780 | 250  | 
|
251  | 
(* lift and instantiate rule *)  | 
|
| 59754 | 252  | 
|
| 59774 | 253  | 
val inc = Thm.maxidx_of st + 1;  | 
| 59825 | 254  | 
val lift_type = Logic.incr_tvar inc;  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60469 
diff
changeset
 | 
255  | 
fun lift_var ((a, j), T) = ((a, j + inc), paramTs ---> lift_type T);  | 
| 59825 | 256  | 
fun lift_term t = fold_rev Term.absfree params (Logic.incr_indexes (fixed, paramTs, inc) t);  | 
| 59754 | 257  | 
|
| 74282 | 258  | 
val inst_tvars' =  | 
259  | 
TVars.build (inst_tvars |> fold (fn (((a, i), S), T) =>  | 
|
260  | 
TVars.add (((a, i + inc), S), Thm.ctyp_of inst_ctxt (lift_type T))));  | 
|
261  | 
val inst_vars' =  | 
|
262  | 
Vars.build (inst_vars |> fold (fn (v, t) =>  | 
|
263  | 
Vars.add (lift_var v, Thm.cterm_of inst_ctxt (lift_term t))));  | 
|
| 
59771
 
c6e60787ffe2
read instantiations uniformly for rules and tactics;
 
wenzelm 
parents: 
59770 
diff
changeset
 | 
264  | 
|
| 59825 | 265  | 
val thm' = Thm.lift_rule cgoal thm  | 
266  | 
|> Drule.instantiate_normalize (inst_tvars', inst_vars')  | 
|
| 59836 | 267  | 
|> singleton (Variable.export inst_ctxt ctxt);  | 
268  | 
in compose_tac ctxt (bires_flag, thm', Thm.nprems_of thm) i end) i st;  | 
|
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
269  | 
|
| 27120 | 270  | 
val res_inst_tac = bires_inst_tac false;  | 
271  | 
val eres_inst_tac = bires_inst_tac true;  | 
|
272  | 
||
273  | 
||
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
274  | 
(* forward resolution *)  | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
275  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58027 
diff
changeset
 | 
276  | 
fun make_elim_preserve ctxt rl =  | 
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
277  | 
let  | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
278  | 
val maxidx = Thm.maxidx_of rl;  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60469 
diff
changeset
 | 
279  | 
fun var x = ((x, 0), propT);  | 
| 59623 | 280  | 
fun cvar xi = Thm.cterm_of ctxt (Var (xi, propT));  | 
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
281  | 
val revcut_rl' =  | 
| 74282 | 282  | 
Drule.revcut_rl |> Drule.instantiate_normalize  | 
| 77879 | 283  | 
        (TVars.empty, Vars.make2 (var "V", cvar ("V", maxidx + 1)) (var "W", cvar ("W", maxidx + 1)));
 | 
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
284  | 
in  | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
46476 
diff
changeset
 | 
285  | 
(case Seq.list_of  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58027 
diff
changeset
 | 
286  | 
      (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false}
 | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
46476 
diff
changeset
 | 
287  | 
(false, rl, Thm.nprems_of rl) 1 revcut_rl')  | 
| 
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
46476 
diff
changeset
 | 
288  | 
of  | 
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
289  | 
[th] => th  | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
290  | 
    | _ => raise THM ("make_elim_preserve", 1, [rl]))
 | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
291  | 
end;  | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
292  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
293  | 
(*instantiate and cut -- for atomic fact*)  | 
| 59780 | 294  | 
fun cut_inst_tac ctxt insts fixes rule =  | 
295  | 
res_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule);  | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
296  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
297  | 
(*forward tactic applies a rule to an assumption without deleting it*)  | 
| 59780 | 298  | 
fun forw_inst_tac ctxt insts fixes rule =  | 
299  | 
cut_inst_tac ctxt insts fixes rule THEN' assume_tac ctxt;  | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
300  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
301  | 
(*dresolve tactic applies a rule to replace an assumption*)  | 
| 59780 | 302  | 
fun dres_inst_tac ctxt insts fixes rule =  | 
303  | 
eres_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule);  | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
304  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
305  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
306  | 
(* derived tactics *)  | 
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
307  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
308  | 
(*deletion of an assumption*)  | 
| 59780 | 309  | 
fun thin_tac ctxt s fixes =  | 
310  | 
  eres_inst_tac ctxt [((("V", 0), Position.none), s)] fixes Drule.thin_rl;
 | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
311  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
312  | 
(*Introduce the given proposition as lemma and subgoal*)  | 
| 59780 | 313  | 
fun subgoal_tac ctxt A fixes =  | 
314  | 
  DETERM o res_inst_tac ctxt [((("psi", 0), Position.none), A)] fixes cut_rl;
 | 
|
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
315  | 
|
| 
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
316  | 
|
| 53708 | 317  | 
(* method wrapper *)  | 
| 
27245
 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 
wenzelm 
parents: 
27236 
diff
changeset
 | 
318  | 
|
| 30545 | 319  | 
fun method inst_tac tac =  | 
| 59855 | 320  | 
Args.goal_spec -- Scan.optional (Scan.lift (named_insts --| Args.$$$ "in")) ([], []) --  | 
321  | 
Attrib.thms >> (fn ((quant, (insts, fixes)), thms) => fn ctxt => METHOD (fn facts =>  | 
|
| 59780 | 322  | 
if null insts andalso null fixes  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
60695 
diff
changeset
 | 
323  | 
then quant (Method.insert_tac ctxt facts THEN' tac ctxt thms)  | 
| 30515 | 324  | 
else  | 
325  | 
(case thms of  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
60695 
diff
changeset
 | 
326  | 
[thm] => quant (Method.insert_tac ctxt facts THEN' inst_tac ctxt insts fixes thm)  | 
| 30515 | 327  | 
| _ => error "Cannot have instantiations with multiple rules")));  | 
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
328  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
329  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
(* setup *)  | 
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
331  | 
|
| 53708 | 332  | 
(*warning: rule_tac etc. refer to dynamic subgoal context!*)  | 
333  | 
||
| 53171 | 334  | 
val _ = Theory.setup  | 
| 67147 | 335  | 
(Method.setup \<^binding>\<open>rule_tac\<close> (method res_inst_tac resolve_tac)  | 
| 53708 | 336  | 
"apply rule (dynamic instantiation)" #>  | 
| 67147 | 337  | 
Method.setup \<^binding>\<open>erule_tac\<close> (method eres_inst_tac eresolve_tac)  | 
| 30515 | 338  | 
"apply rule in elimination manner (dynamic instantiation)" #>  | 
| 67147 | 339  | 
Method.setup \<^binding>\<open>drule_tac\<close> (method dres_inst_tac dresolve_tac)  | 
| 30515 | 340  | 
"apply rule in destruct manner (dynamic instantiation)" #>  | 
| 67147 | 341  | 
Method.setup \<^binding>\<open>frule_tac\<close> (method forw_inst_tac forward_tac)  | 
| 30515 | 342  | 
"apply rule in forward manner (dynamic instantiation)" #>  | 
| 67147 | 343  | 
Method.setup \<^binding>\<open>cut_tac\<close> (method cut_inst_tac (K cut_rules_tac))  | 
| 53708 | 344  | 
"cut rule (dynamic instantiation)" #>  | 
| 67147 | 345  | 
Method.setup \<^binding>\<open>subgoal_tac\<close>  | 
| 74563 | 346  | 
(Args.goal_spec -- Scan.lift (Scan.repeat1 Parse.embedded_inner_syntax -- Parse.for_fixes) >>  | 
| 59780 | 347  | 
(fn (quant, (props, fixes)) => fn ctxt =>  | 
348  | 
SIMPLE_METHOD'' quant (EVERY' (map (fn prop => subgoal_tac ctxt prop fixes) props))))  | 
|
| 30515 | 349  | 
"insert subgoal (dynamic instantiation)" #>  | 
| 67147 | 350  | 
Method.setup \<^binding>\<open>thin_tac\<close>  | 
| 74563 | 351  | 
(Args.goal_spec -- Scan.lift (Parse.embedded_inner_syntax -- Parse.for_fixes) >>  | 
| 59780 | 352  | 
(fn (quant, (prop, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (thin_tac ctxt prop fixes)))  | 
353  | 
"remove premise (dynamic instantiation)");  | 
|
| 
20336
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
354  | 
|
| 
 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 
wenzelm 
parents:  
diff
changeset
 | 
355  | 
end;  |