author | blanchet |
Wed, 03 Nov 2010 22:33:23 +0100 | |
changeset 40342 | 3154f63e2bda |
parent 39288 | f1ae2493d93f |
child 42360 | da8817d01e7c |
permissions | -rw-r--r-- |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/Isar/rule_insts.ML |
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 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
4 |
Rule instantiations -- operations within a rule/subgoal context. |
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 |
|
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
|
7 |
signature BASIC_RULE_INSTS = |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
8 |
sig |
27236 | 9 |
val read_instantiate: Proof.context -> (indexname * string) list -> thm -> thm |
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
|
10 |
val instantiate_tac: Proof.context -> (indexname * string) list -> tactic |
27120 | 11 |
val res_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic |
12 |
val eres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic |
|
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
|
13 |
val cut_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
14 |
val forw_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
15 |
val dres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
16 |
val thin_tac: Proof.context -> string -> int -> tactic |
27219 | 17 |
val subgoal_tac: Proof.context -> string -> int -> tactic |
18 |
val subgoals_tac: Proof.context -> string list -> int -> tactic |
|
30545 | 19 |
val method: (Proof.context -> (indexname * string) list -> thm -> int -> tactic) -> |
30551 | 20 |
(Proof.context -> thm list -> int -> tactic) -> (Proof.context -> Proof.method) context_parser |
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
|
21 |
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
|
22 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
23 |
signature RULE_INSTS = |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
24 |
sig |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
25 |
include BASIC_RULE_INSTS |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
26 |
val make_elim_preserve: thm -> thm |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
27 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
28 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
29 |
structure RuleInsts: RULE_INSTS = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
30 |
struct |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
31 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
32 |
(** reading instantiations **) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
33 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
34 |
local |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
35 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
36 |
fun is_tvar (x, _) = String.isPrefix "'" x; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
37 |
|
22681
9d42e5365ad1
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents:
21879
diff
changeset
|
38 |
fun error_var msg xi = error (msg ^ Term.string_of_vname xi); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
39 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
40 |
fun the_sort tvars xi = the (AList.lookup (op =) tvars xi) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
41 |
handle Option.Option => error_var "No such type variable in theorem: " xi; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
42 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
43 |
fun the_type vars xi = the (AList.lookup (op =) vars xi) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
44 |
handle Option.Option => error_var "No such variable in theorem: " xi; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
45 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
46 |
fun unify_vartypes thy vars (xi, u) (unifier, maxidx) = |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
47 |
let |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
48 |
val T = the_type vars xi; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
49 |
val U = Term.fastype_of u; |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
50 |
val maxidx' = Term.maxidx_term u (Int.max (#2 xi, maxidx)); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
51 |
in |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
52 |
Sign.typ_unify thy (T, U) (unifier, maxidx') |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
53 |
handle Type.TUNIFY => error_var "Incompatible type for instantiation of " xi |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
54 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
55 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
56 |
fun instantiate inst = |
31977 | 57 |
Term_Subst.instantiate ([], map (fn (xi, t) => ((xi, Term.fastype_of t), t)) inst) #> |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
58 |
Envir.beta_norm; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
59 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
60 |
fun make_instT f v = |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
61 |
let |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
62 |
val T = TVar v; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
63 |
val T' = f T; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
64 |
in if T = T' then NONE else SOME (T, T') end; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
65 |
|
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
66 |
fun make_inst f v = |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
67 |
let |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
68 |
val t = Var v; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
69 |
val t' = f t; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
70 |
in if t aconv t' then NONE else SOME (t, t') end; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
71 |
|
27282 | 72 |
val add_used = |
73 |
(Thm.fold_terms o fold_types o fold_atyps) |
|
74 |
(fn TFree (a, _) => insert (op =) a |
|
75 |
| TVar ((a, _), _) => insert (op =) a |
|
76 |
| _ => I); |
|
77 |
||
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
78 |
in |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
79 |
|
25333 | 80 |
fun read_termTs ctxt schematic ss Ts = |
25329 | 81 |
let |
82 |
fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt; |
|
83 |
val ts = map2 parse Ts ss; |
|
84 |
val ts' = |
|
39288 | 85 |
map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts |
25333 | 86 |
|> Syntax.check_terms ((schematic ? ProofContext.set_mode ProofContext.mode_schematic) ctxt) |
25329 | 87 |
|> Variable.polymorphic ctxt; |
88 |
val Ts' = map Term.fastype_of ts'; |
|
89 |
val tyenv = fold Type.raw_match (Ts ~~ Ts') Vartab.empty; |
|
90 |
in (ts', map (apsnd snd) (Vartab.dest tyenv)) end; |
|
91 |
||
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
92 |
fun read_insts ctxt mixed_insts (tvars, vars) = |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
93 |
let |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
94 |
val thy = ProofContext.theory_of ctxt; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
95 |
val cert = Thm.cterm_of thy; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
96 |
val certT = Thm.ctyp_of thy; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
97 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
98 |
val (type_insts, term_insts) = List.partition (is_tvar o fst) mixed_insts; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
99 |
val internal_insts = term_insts |> map_filter |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
100 |
(fn (xi, Token.Term t) => SOME (xi, t) |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
101 |
| (_, Token.Text _) => NONE |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
102 |
| (xi, _) => error_var "Term argument expected for " xi); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
103 |
val external_insts = term_insts |> map_filter |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
104 |
(fn (xi, Token.Text s) => SOME (xi, s) | _ => NONE); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
105 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
106 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
107 |
(* mixed type instantiations *) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
108 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
109 |
fun readT (xi, arg) = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
110 |
let |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
111 |
val S = the_sort tvars xi; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
112 |
val T = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
113 |
(case arg of |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
114 |
Token.Text s => Syntax.read_typ ctxt s |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
115 |
| Token.Typ T => T |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
116 |
| _ => error_var "Type argument expected for " xi); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
117 |
in |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
118 |
if Sign.of_sort thy (T, S) then ((xi, S), T) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
119 |
else error_var "Incompatible sort for typ instantiation of " xi |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
120 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
121 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
122 |
val type_insts1 = map readT type_insts; |
31977 | 123 |
val instT1 = Term_Subst.instantiateT type_insts1; |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
124 |
val vars1 = map (apsnd instT1) vars; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
125 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
126 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
127 |
(* internal term instantiations *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
128 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
129 |
val instT2 = Envir.norm_type |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
130 |
(#1 (fold (unify_vartypes thy vars1) internal_insts (Vartab.empty, 0))); |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
131 |
val vars2 = map (apsnd instT2) vars1; |
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20509
diff
changeset
|
132 |
val internal_insts2 = map (apsnd (map_types instT2)) internal_insts; |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
133 |
val inst2 = instantiate internal_insts2; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
134 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
135 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
136 |
(* external term instantiations *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
137 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
138 |
val (xs, strs) = split_list external_insts; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
139 |
val Ts = map (the_type vars2) xs; |
25354 | 140 |
val (ts, inferred) = read_termTs ctxt false strs Ts; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
141 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
142 |
val instT3 = Term.typ_subst_TVars inferred; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
143 |
val vars3 = map (apsnd instT3) vars2; |
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20509
diff
changeset
|
144 |
val internal_insts3 = map (apsnd (map_types instT3)) internal_insts2; |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
145 |
val external_insts3 = xs ~~ ts; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
146 |
val inst3 = instantiate external_insts3; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
147 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
148 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
149 |
(* results *) |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
150 |
|
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
151 |
val type_insts3 = map (fn ((a, _), T) => (a, instT3 (instT2 T))) type_insts1; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
152 |
val term_insts3 = internal_insts3 @ external_insts3; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
153 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
154 |
val inst_tvars = map_filter (make_instT (instT3 o instT2 o instT1)) tvars; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
155 |
val inst_vars = map_filter (make_inst (inst3 o inst2)) vars3; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
156 |
in |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
157 |
((type_insts3, term_insts3), |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
158 |
(map (pairself certT) inst_tvars, map (pairself cert) inst_vars)) |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
159 |
end; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
160 |
|
27236 | 161 |
fun read_instantiate_mixed ctxt mixed_insts thm = |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
162 |
let |
20487
6ac7a4fc32a0
read_instantiate: declare names of TVars as well (temporary workaround for no-freeze feature of type inference);
wenzelm
parents:
20343
diff
changeset
|
163 |
val ctxt' = ctxt |> Variable.declare_thm thm |
27282 | 164 |
|> fold (fn a => Variable.declare_names (Logic.mk_type (TFree (a, dummyS)))) (add_used thm []); (* FIXME tmp *) |
22692 | 165 |
val tvars = Thm.fold_terms Term.add_tvars thm []; |
166 |
val vars = Thm.fold_terms Term.add_vars thm []; |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
167 |
val ((type_insts, term_insts), insts) = read_insts ctxt' (map snd mixed_insts) (tvars, vars); |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
168 |
|
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
169 |
val _ = (*assign internalized values*) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
170 |
mixed_insts |> List.app (fn (arg, (xi, _)) => |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
171 |
if is_tvar xi then |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
172 |
Token.assign (SOME (Token.Typ (the (AList.lookup (op =) type_insts xi)))) arg |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
173 |
else |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
174 |
Token.assign (SOME (Token.Term (the (AList.lookup (op =) term_insts xi)))) arg); |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
175 |
in |
33368 | 176 |
Drule.instantiate insts thm |> Rule_Cases.save thm |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
177 |
end; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
178 |
|
27236 | 179 |
fun read_instantiate_mixed' ctxt (args, concl_args) thm = |
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
180 |
let |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
181 |
fun zip_vars _ [] = [] |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
182 |
| zip_vars (_ :: xs) ((_, NONE) :: rest) = zip_vars xs rest |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
183 |
| zip_vars ((x, _) :: xs) ((arg, SOME t) :: rest) = (arg, (x, t)) :: zip_vars xs rest |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
184 |
| zip_vars [] _ = error "More instantiations than variables in theorem"; |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
185 |
val insts = |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
186 |
zip_vars (rev (Term.add_vars (Thm.full_prop_of thm) [])) args @ |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
187 |
zip_vars (rev (Term.add_vars (Thm.concl_of thm) [])) concl_args; |
27236 | 188 |
in read_instantiate_mixed ctxt insts thm end; |
189 |
||
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
|
190 |
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
|
191 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
192 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
193 |
(* instantiation of rule or goal state *) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
194 |
|
27236 | 195 |
fun read_instantiate ctxt args thm = |
196 |
read_instantiate_mixed (ctxt |> ProofContext.set_mode ProofContext.mode_schematic) (* FIXME !? *) |
|
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
197 |
(map (fn (x, y) => (Token.eof, (x, Token.Text y))) args) thm; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
198 |
|
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
|
199 |
fun instantiate_tac ctxt args = PRIMITIVE (read_instantiate ctxt args); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
200 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
201 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
202 |
|
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
203 |
(** attributes **) |
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
204 |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
205 |
(* where: named instantiation *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
206 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
207 |
local |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
208 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
209 |
val value = |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
210 |
Args.internal_typ >> Token.Typ || |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
211 |
Args.internal_term >> Token.Term || |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
212 |
Args.name_source >> Token.Text; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
213 |
|
36950 | 214 |
val inst = Args.var -- (Args.$$$ "=" |-- Scan.ahead Parse.not_eof -- value) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
215 |
>> (fn (xi, (a, v)) => (a, (xi, v))); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
216 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
217 |
in |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
218 |
|
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
219 |
val _ = Context.>> (Context.map_theory |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
220 |
(Attrib.setup (Binding.name "where") |
36950 | 221 |
(Scan.lift (Parse.and_list inst) >> (fn args => |
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
222 |
Thm.rule_attribute (fn context => read_instantiate_mixed (Context.proof_of context) args))) |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
223 |
"named instantiation of theorem")); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
224 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
225 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
226 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
227 |
|
20343
e093a54bf25e
reworked read_instantiate -- separate read_insts;
wenzelm
parents:
20336
diff
changeset
|
228 |
(* of: positional instantiation (terms only) *) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
229 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
230 |
local |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
231 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
232 |
val value = |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
233 |
Args.internal_term >> Token.Term || |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
234 |
Args.name_source >> Token.Text; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
235 |
|
36950 | 236 |
val inst = Scan.ahead Parse.not_eof -- Args.maybe value; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
237 |
val concl = Args.$$$ "concl" -- Args.colon; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
238 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
239 |
val insts = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
240 |
Scan.repeat (Scan.unless concl inst) -- |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
241 |
Scan.optional (concl |-- Scan.repeat inst) []; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
242 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
243 |
in |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
244 |
|
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
245 |
val _ = Context.>> (Context.map_theory |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
246 |
(Attrib.setup (Binding.name "of") |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
247 |
(Scan.lift insts >> (fn args => |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
248 |
Thm.rule_attribute (fn context => read_instantiate_mixed' (Context.proof_of context) args))) |
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30551
diff
changeset
|
249 |
"positional instantiation of theorem")); |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
250 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
251 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
252 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
253 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
254 |
|
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
|
255 |
(** tactics **) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
256 |
|
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
|
257 |
(* resolution after lifting and instantation; may refer to parameters of the subgoal *) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
258 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
259 |
(* FIXME cleanup this mess!!! *) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
260 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
261 |
fun bires_inst_tac bires_flag ctxt insts thm = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
262 |
let |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
263 |
val thy = ProofContext.theory_of ctxt; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
264 |
(* Separate type and term insts *) |
32784 | 265 |
fun has_type_var ((x, _), _) = |
266 |
(case Symbol.explode x of "'" :: _ => true | _ => false); |
|
33317 | 267 |
val Tinsts = filter has_type_var insts; |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
268 |
val tinsts = filter_out has_type_var insts; |
25333 | 269 |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
270 |
(* Tactic *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
271 |
fun tac i st = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
272 |
let |
25333 | 273 |
val (_, _, Bi, _) = Thm.dest_state (st, i); |
274 |
val params = Logic.strip_params Bi; (*params of subgoal i as string typ pairs*) |
|
275 |
val params = rev (Term.rename_wrt_term Bi params) |
|
276 |
(*as they are printed: bound variables with*) |
|
277 |
(*the same name are renamed during printing*) |
|
278 |
||
279 |
val (param_names, ctxt') = ctxt |
|
280 |
|> Variable.declare_thm thm |
|
281 |
|> Thm.fold_terms Variable.declare_constraints st |
|
30763
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents:
30722
diff
changeset
|
282 |
|> ProofContext.add_fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) params); |
25333 | 283 |
|
284 |
(* Process type insts: Tinsts_env *) |
|
285 |
fun absent xi = error |
|
286 |
("No such variable in theorem: " ^ Term.string_of_vname xi); |
|
287 |
val (rtypes, rsorts) = Drule.types_sorts thm; |
|
288 |
fun readT (xi, s) = |
|
289 |
let val S = case rsorts xi of SOME S => S | NONE => absent xi; |
|
290 |
val T = Syntax.read_typ ctxt' s; |
|
291 |
val U = TVar (xi, S); |
|
292 |
in if Sign.typ_instance thy (T, U) then (U, T) |
|
293 |
else error ("Instantiation of " ^ Term.string_of_vname xi ^ " fails") |
|
294 |
end; |
|
295 |
val Tinsts_env = map readT Tinsts; |
|
296 |
(* Preprocess rule: extract vars and their types, apply Tinsts *) |
|
297 |
fun get_typ xi = |
|
298 |
(case rtypes xi of |
|
299 |
SOME T => typ_subst_atomic Tinsts_env T |
|
300 |
| NONE => absent xi); |
|
301 |
val (xis, ss) = Library.split_list tinsts; |
|
302 |
val Ts = map get_typ xis; |
|
303 |
||
304 |
val (ts, envT) = read_termTs ctxt' true ss Ts; |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
305 |
val envT' = map (fn (ixn, T) => |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
306 |
(TVar (ixn, the (rsorts ixn)), T)) envT @ Tinsts_env; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
307 |
val cenv = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
308 |
map |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
309 |
(fn (xi, t) => |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
310 |
pairself (Thm.cterm_of thy) (Var (xi, fastype_of t), t)) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
311 |
(distinct |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
312 |
(fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
313 |
(xis ~~ ts)); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
314 |
(* Lift and instantiate rule *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
315 |
val {maxidx, ...} = rep_thm st; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
316 |
val paramTs = map #2 params |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
317 |
and inc = maxidx+1 |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
318 |
fun liftvar (Var ((a,j), T)) = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
319 |
Var((a, j+inc), paramTs ---> Logic.incr_tvar inc T) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
320 |
| liftvar t = raise TERM("Variable expected", [t]); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
321 |
fun liftterm t = list_abs_free |
25333 | 322 |
(param_names ~~ paramTs, Logic.incr_indexes(paramTs,inc) t) |
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
323 |
fun liftpair (cv,ct) = |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
324 |
(cterm_fun liftvar cv, cterm_fun liftterm ct) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
325 |
val lifttvar = pairself (ctyp_of thy o Logic.incr_tvar inc); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
326 |
val rule = Drule.instantiate |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
327 |
(map lifttvar envT', map liftpair cenv) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
328 |
(Thm.lift_rule (Thm.cprem_of st i) thm) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
329 |
in |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
330 |
if i > nprems_of st then no_tac st |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
331 |
else st |> |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
332 |
compose_tac (bires_flag, rule, nprems_of thm) i |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
333 |
end |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
334 |
handle TERM (msg,_) => (warning msg; no_tac st) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
335 |
| THM (msg,_,_) => (warning msg; no_tac st); |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
336 |
in tac end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
337 |
|
27120 | 338 |
val res_inst_tac = bires_inst_tac false; |
339 |
val eres_inst_tac = bires_inst_tac true; |
|
340 |
||
341 |
||
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
|
342 |
(* 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
|
343 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
344 |
fun make_elim_preserve 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
|
345 |
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
|
346 |
val cert = Thm.cterm_of (Thm.theory_of_thm 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
|
347 |
val maxidx = Thm.maxidx_of 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
|
348 |
fun cvar xi = cert (Var (xi, propT)); |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
349 |
val revcut_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
|
350 |
instantiate ([], [(cvar ("V", 0), cvar ("V", maxidx + 1)), |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
351 |
(cvar ("W", 0), cvar ("W", maxidx + 1))]) Drule.revcut_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
|
352 |
in |
31945 | 353 |
(case Seq.list_of (Thm.bicompose false (false, rl, Thm.nprems_of rl) 1 revcut_rl') 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
|
354 |
[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
|
355 |
| _ => 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
|
356 |
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
|
357 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
358 |
(*instantiate and cut -- for atomic fact*) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
359 |
fun cut_inst_tac ctxt insts rule = res_inst_tac ctxt insts (make_elim_preserve rule); |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
360 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
361 |
(*forward tactic applies a rule to an assumption without deleting it*) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
362 |
fun forw_inst_tac ctxt insts rule = cut_inst_tac ctxt insts rule THEN' assume_tac; |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
363 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
364 |
(*dresolve tactic applies a rule to replace an assumption*) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
365 |
fun dres_inst_tac ctxt insts rule = eres_inst_tac ctxt insts (make_elim_preserve rule); |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
366 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
367 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
368 |
(* 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
|
369 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
370 |
(*deletion of an assumption*) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
371 |
fun thin_tac ctxt s = eres_inst_tac ctxt [(("V", 0), s)] Drule.thin_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
|
372 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
373 |
(*Introduce the given proposition as lemma and subgoal*) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
374 |
fun subgoal_tac ctxt A = DETERM o res_inst_tac ctxt [(("psi", 0), A)] cut_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
|
375 |
fun subgoals_tac ctxt As = EVERY' (map (subgoal_tac ctxt) As); |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
376 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
377 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
378 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
379 |
(** methods **) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
380 |
|
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
381 |
(* rule_tac etc. -- refer to dynamic goal state! *) |
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
wenzelm
parents:
27236
diff
changeset
|
382 |
|
30545 | 383 |
fun method inst_tac tac = |
30515 | 384 |
Args.goal_spec -- |
385 |
Scan.optional (Scan.lift |
|
36950 | 386 |
(Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.name_source)) --| |
387 |
Args.$$$ "in")) [] -- |
|
30515 | 388 |
Attrib.thms >> |
389 |
(fn ((quant, insts), thms) => fn ctxt => METHOD (fn facts => |
|
30551 | 390 |
if null insts then quant (Method.insert_tac facts THEN' tac ctxt thms) |
30515 | 391 |
else |
392 |
(case thms of |
|
393 |
[thm] => quant (Method.insert_tac facts THEN' inst_tac ctxt insts thm) |
|
394 |
| _ => error "Cannot have instantiations with multiple rules"))); |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
395 |
|
30551 | 396 |
val res_inst_meth = method res_inst_tac (K Tactic.resolve_tac); |
397 |
val eres_inst_meth = method eres_inst_tac (K Tactic.eresolve_tac); |
|
398 |
val cut_inst_meth = method cut_inst_tac (K Tactic.cut_rules_tac); |
|
399 |
val dres_inst_meth = method dres_inst_tac (K Tactic.dresolve_tac); |
|
400 |
val forw_inst_meth = method forw_inst_tac (K Tactic.forward_tac); |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
401 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
402 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
403 |
(* setup *) |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
404 |
|
26463 | 405 |
val _ = Context.>> (Context.map_theory |
30515 | 406 |
(Method.setup (Binding.name "rule_tac") res_inst_meth "apply rule (dynamic instantiation)" #> |
407 |
Method.setup (Binding.name "erule_tac") eres_inst_meth |
|
408 |
"apply rule in elimination manner (dynamic instantiation)" #> |
|
409 |
Method.setup (Binding.name "drule_tac") dres_inst_meth |
|
410 |
"apply rule in destruct manner (dynamic instantiation)" #> |
|
411 |
Method.setup (Binding.name "frule_tac") forw_inst_meth |
|
412 |
"apply rule in forward manner (dynamic instantiation)" #> |
|
413 |
Method.setup (Binding.name "cut_tac") cut_inst_meth "cut rule (dynamic instantiation)" #> |
|
414 |
Method.setup (Binding.name "subgoal_tac") |
|
415 |
(Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name_source) >> |
|
416 |
(fn (quant, props) => fn ctxt => SIMPLE_METHOD'' quant (subgoals_tac ctxt props))) |
|
417 |
"insert subgoal (dynamic instantiation)" #> |
|
418 |
Method.setup (Binding.name "thin_tac") |
|
419 |
(Args.goal_spec -- Scan.lift Args.name_source >> |
|
420 |
(fn (quant, prop) => fn ctxt => SIMPLE_METHOD'' quant (thin_tac ctxt prop))) |
|
421 |
"remove premise (dynamic instantiation)")); |
|
20336
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
422 |
|
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
423 |
end; |
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
wenzelm
parents:
diff
changeset
|
424 |
|
36950 | 425 |
structure Basic_Rule_Insts: BASIC_RULE_INSTS = RuleInsts; |
426 |
open Basic_Rule_Insts; |