|
59739
|
1 |
infix 1 then_cconv
|
|
|
2 |
infix 0 else_cconv
|
|
|
3 |
|
|
|
4 |
type cconv = conv
|
|
|
5 |
|
|
|
6 |
signature BASIC_CCONV =
|
|
|
7 |
sig
|
|
|
8 |
val then_cconv: cconv * cconv -> cconv
|
|
|
9 |
val else_cconv: cconv * cconv -> cconv
|
|
|
10 |
val CCONVERSION: cconv -> int -> tactic
|
|
|
11 |
end
|
|
|
12 |
|
|
|
13 |
signature CCONV =
|
|
|
14 |
sig
|
|
|
15 |
include BASIC_CCONV
|
|
|
16 |
val no_cconv: cconv
|
|
|
17 |
val all_cconv: cconv
|
|
|
18 |
val first_cconv: cconv list -> cconv
|
|
|
19 |
val abs_cconv: (cterm * Proof.context -> cconv) -> Proof.context -> cconv
|
|
|
20 |
val combination_cconv: cconv -> cconv -> cconv
|
|
|
21 |
val comb_cconv: cconv -> cconv
|
|
|
22 |
val arg_cconv: cconv -> cconv
|
|
|
23 |
val fun_cconv: cconv -> cconv
|
|
|
24 |
val arg1_cconv: cconv -> cconv
|
|
|
25 |
val fun2_cconv: cconv -> cconv
|
|
|
26 |
val rewr_cconv: thm -> cconv
|
|
|
27 |
val rewrs_cconv: thm list -> cconv
|
|
|
28 |
val params_cconv: int -> (Proof.context -> cconv) -> Proof.context -> cconv
|
|
|
29 |
val prems_cconv: int -> cconv -> cconv
|
|
|
30 |
val concl_cconv: int -> cconv -> cconv
|
|
|
31 |
val fconv_rule: cconv -> thm -> thm
|
|
|
32 |
val gconv_rule: cconv -> int -> thm -> thm
|
|
|
33 |
end
|
|
|
34 |
|
|
|
35 |
structure CConv : CCONV =
|
|
|
36 |
struct
|
|
|
37 |
|
|
|
38 |
val concl_lhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_lhs
|
|
|
39 |
val concl_rhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_rhs
|
|
|
40 |
|
|
|
41 |
fun transitive th1 th2 = Drule.transitive_thm OF [th1, th2]
|
|
|
42 |
|
|
|
43 |
val combination_thm =
|
|
|
44 |
let
|
|
|
45 |
val fg = @{cprop "f :: 'a :: {} \<Rightarrow> 'b :: {} \<equiv> g"}
|
|
|
46 |
val st = @{cprop "s :: 'a :: {} \<equiv> t"}
|
|
|
47 |
val thm = Thm.combination (Thm.assume fg) (Thm.assume st)
|
|
|
48 |
|> Thm.implies_intr st
|
|
|
49 |
|> Thm.implies_intr fg
|
|
|
50 |
in Drule.export_without_context thm end
|
|
|
51 |
|
|
|
52 |
fun abstract_rule_thm n =
|
|
|
53 |
let
|
|
|
54 |
val eq = @{cprop "\<And>x :: 'a :: {}. (s :: 'a \<Rightarrow> 'b :: {}) x \<equiv> t x"}
|
|
|
55 |
val x = @{cterm "x :: 'a :: {}"}
|
|
|
56 |
val thm = eq
|
|
|
57 |
|> Thm.assume
|
|
|
58 |
|> Thm.forall_elim x
|
|
|
59 |
|> Thm.abstract_rule n x
|
|
|
60 |
|> Thm.implies_intr eq
|
|
|
61 |
in Drule.export_without_context thm end
|
|
|
62 |
|
|
|
63 |
val no_cconv = Conv.no_conv
|
|
|
64 |
val all_cconv = Conv.all_conv
|
|
|
65 |
|
|
|
66 |
fun (cv1 else_cconv cv2) ct =
|
|
|
67 |
(cv1 ct
|
|
|
68 |
handle THM _ => cv2 ct
|
|
|
69 |
| CTERM _ => cv2 ct
|
|
|
70 |
| TERM _ => cv2 ct
|
|
|
71 |
| TYPE _ => cv2 ct)
|
|
|
72 |
|
|
|
73 |
fun (cv1 then_cconv cv2) ct =
|
|
|
74 |
let
|
|
|
75 |
val eq1 = cv1 ct
|
|
|
76 |
val eq2 = cv2 (concl_rhs_of eq1)
|
|
|
77 |
in
|
|
|
78 |
if Thm.is_reflexive eq1 then eq2
|
|
|
79 |
else if Thm.is_reflexive eq2 then eq1
|
|
|
80 |
else transitive eq1 eq2
|
|
|
81 |
end
|
|
|
82 |
|
|
|
83 |
fun first_cconv cvs = fold_rev (curry op else_cconv) cvs no_cconv
|
|
|
84 |
|
|
|
85 |
fun rewr_cconv rule ct =
|
|
|
86 |
let
|
|
|
87 |
val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule
|
|
|
88 |
val lhs = concl_lhs_of rule1
|
|
|
89 |
val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1
|
|
|
90 |
val rule3 = Thm.instantiate (Thm.match (lhs, ct)) rule2
|
|
|
91 |
handle Pattern.MATCH => raise CTERM ("rewr_cconv", [lhs, ct])
|
|
|
92 |
val rule4 =
|
|
|
93 |
if concl_lhs_of rule3 aconvc ct then rule3
|
|
|
94 |
else let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3)
|
|
|
95 |
in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (concl_rhs_of rule3)) end
|
|
|
96 |
in
|
|
|
97 |
transitive rule4 (Thm.beta_conversion true (concl_rhs_of rule4))
|
|
|
98 |
end
|
|
|
99 |
|
|
|
100 |
fun rewrs_cconv rules = first_cconv (map rewr_cconv rules)
|
|
|
101 |
|
|
|
102 |
fun combination_cconv cv1 cv2 cterm =
|
|
|
103 |
let val (l, r) = Thm.dest_comb cterm
|
|
|
104 |
in combination_thm OF [cv1 l, cv2 r] end
|
|
|
105 |
|
|
|
106 |
fun comb_cconv cv = combination_cconv cv cv
|
|
|
107 |
|
|
|
108 |
fun fun_cconv conversion =
|
|
|
109 |
combination_cconv conversion all_cconv
|
|
|
110 |
|
|
|
111 |
fun arg_cconv conversion =
|
|
|
112 |
combination_cconv all_cconv conversion
|
|
|
113 |
|
|
|
114 |
fun abs_cconv cv ctxt ct =
|
|
|
115 |
(case Thm.term_of ct of
|
|
|
116 |
Abs (x, _, _) =>
|
|
|
117 |
let
|
|
|
118 |
(* Instantiate the rule properly and apply it to the eq theorem. *)
|
|
|
119 |
fun abstract_rule u v eq =
|
|
|
120 |
let
|
|
|
121 |
(* Take a variable v and an equality theorem of form:
|
|
|
122 |
P1 Pure.imp ... Pure.imp Pn Pure.imp L v == R v
|
|
|
123 |
And build a term of form:
|
|
|
124 |
!!v. (%x. L x) v == (%x. R x) v *)
|
|
|
125 |
fun mk_concl var eq =
|
|
|
126 |
let
|
|
|
127 |
val certify = Thm.cterm_of ctxt
|
|
|
128 |
fun abs term = (Term.lambda var term) $ var
|
|
|
129 |
fun equals_cong f t =
|
|
|
130 |
Logic.dest_equals t
|
|
|
131 |
|> (fn (a, b) => (f a, f b))
|
|
|
132 |
|> Logic.mk_equals
|
|
|
133 |
in
|
|
|
134 |
Thm.concl_of eq
|
|
|
135 |
|> equals_cong abs
|
|
|
136 |
|> Logic.all var |> certify
|
|
|
137 |
end
|
|
|
138 |
val rule = abstract_rule_thm x
|
|
|
139 |
val inst = Thm.match (Drule.cprems_of rule |> hd, mk_concl (Thm.term_of v) eq)
|
|
|
140 |
in
|
|
|
141 |
(Drule.instantiate_normalize inst rule OF [Drule.generalize ([], [u]) eq])
|
|
|
142 |
|> Drule.zero_var_indexes
|
|
|
143 |
end
|
|
|
144 |
|
|
|
145 |
(* Destruct the abstraction and apply the conversion. *)
|
|
|
146 |
val (u, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt
|
|
|
147 |
val (v, ct') = Thm.dest_abs (SOME u) ct
|
|
|
148 |
val eq = cv (v, ctxt') ct'
|
|
|
149 |
in
|
|
|
150 |
if Thm.is_reflexive eq
|
|
|
151 |
then all_cconv ct
|
|
|
152 |
else abstract_rule u v eq
|
|
|
153 |
end
|
|
|
154 |
| _ => raise CTERM ("abs_cconv", [ct]))
|
|
|
155 |
|
|
|
156 |
val arg1_cconv = fun_cconv o arg_cconv
|
|
|
157 |
val fun2_cconv = fun_cconv o fun_cconv
|
|
|
158 |
|
|
|
159 |
(* conversions on HHF rules *)
|
|
|
160 |
|
|
|
161 |
(*rewrite B in !!x1 ... xn. B*)
|
|
|
162 |
fun params_cconv n cv ctxt ct =
|
|
|
163 |
if n <> 0 andalso Logic.is_all (Thm.term_of ct)
|
|
|
164 |
then arg_cconv (abs_cconv (params_cconv (n - 1) cv o #2) ctxt) ct
|
|
|
165 |
else cv ctxt ct
|
|
|
166 |
|
|
|
167 |
(* TODO: This code behaves not exactly like Conv.prems_cconv does.
|
|
|
168 |
Fix this! *)
|
|
|
169 |
(*rewrite the A's in A1 Pure.imp ... Pure.imp An Pure.imp B*)
|
|
|
170 |
fun prems_cconv 0 cv ct = cv ct
|
|
|
171 |
| prems_cconv n cv ct =
|
|
|
172 |
(case ct |> Thm.term_of of
|
|
|
173 |
(Const (@{const_name "Pure.imp"}, _) $ _) $ _ => ((if n = 1 then fun_cconv else I) o arg_cconv) (prems_cconv (n-1) cv) ct
|
|
|
174 |
| _ => cv ct)
|
|
|
175 |
|
|
|
176 |
(*rewrite B in A1 Pure.imp ... Pure.imp An Pure.imp B*)
|
|
|
177 |
fun concl_cconv 0 cv ct = cv ct
|
|
|
178 |
| concl_cconv n cv ct =
|
|
|
179 |
(case ct |> Thm.term_of of
|
|
|
180 |
(Const (@{const_name "Pure.imp"}, _) $ _) $ _ => arg_cconv (concl_cconv (n-1) cv) ct
|
|
|
181 |
| _ => cv ct)
|
|
|
182 |
|
|
|
183 |
(*forward conversion, cf. FCONV_RULE in LCF*)
|
|
|
184 |
fun fconv_rule cv th =
|
|
|
185 |
let
|
|
|
186 |
val eq = cv (Thm.cprop_of th)
|
|
|
187 |
in
|
|
|
188 |
if Thm.is_reflexive eq then th
|
|
|
189 |
else th COMP (Thm.permute_prems 0 (Thm.nprems_of eq) (eq RS Drule.equal_elim_rule1))
|
|
|
190 |
end
|
|
|
191 |
|
|
|
192 |
(*goal conversion*)
|
|
|
193 |
fun gconv_rule cv i th =
|
|
|
194 |
(case try (Thm.cprem_of th) i of
|
|
|
195 |
SOME ct =>
|
|
|
196 |
let
|
|
|
197 |
val eq = cv ct
|
|
|
198 |
|
|
|
199 |
(* Drule.with_subgoal assumes that there are no new premises generated
|
|
|
200 |
and thus rotates the premises wrongly. *)
|
|
|
201 |
fun with_subgoal i f thm =
|
|
|
202 |
let
|
|
|
203 |
val num_prems = Thm.nprems_of thm
|
|
|
204 |
val rotate_to_front = rotate_prems (i - 1)
|
|
|
205 |
fun rotate_back thm = rotate_prems (1 - i + num_prems - Thm.nprems_of thm) thm
|
|
|
206 |
in
|
|
|
207 |
thm |> rotate_to_front |> f |> rotate_back
|
|
|
208 |
end
|
|
|
209 |
in
|
|
|
210 |
if Thm.is_reflexive eq then th
|
|
|
211 |
else with_subgoal i (fconv_rule (arg1_cconv (K eq))) th
|
|
|
212 |
end
|
|
|
213 |
| NONE => raise THM ("gconv_rule", i, [th]))
|
|
|
214 |
|
|
|
215 |
(* Conditional conversions as tactics. *)
|
|
|
216 |
fun CCONVERSION cv i st = Seq.single (gconv_rule cv i st)
|
|
|
217 |
handle THM _ => Seq.empty
|
|
|
218 |
| CTERM _ => Seq.empty
|
|
|
219 |
| TERM _ => Seq.empty
|
|
|
220 |
| TYPE _ => Seq.empty
|
|
|
221 |
|
|
|
222 |
end
|
|
|
223 |
|
|
|
224 |
structure Basic_CConv: BASIC_CCONV = CConv
|
|
|
225 |
open Basic_CConv
|