| author | urbanc |
| Fri, 27 Jun 2008 00:37:30 +0200 | |
| changeset 27374 | 2a3c22fd95ab |
| parent 27239 | f2f42f9fa09d |
| child 28536 | 8dccb6035d0f |
| permissions | -rw-r--r-- |
| 23152 | 1 |
(* Title: HOLCF/Tools/domain/domain_theorems.ML |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
New proofs/tactics by Brian Huffman |
|
5 |
||
6 |
Proof generator for domain command. |
|
7 |
*) |
|
8 |
||
| 26342 | 9 |
val HOLCF_ss = @{simpset};
|
| 23152 | 10 |
|
11 |
structure Domain_Theorems = struct |
|
12 |
||
13 |
local |
|
14 |
||
| 25805 | 15 |
val adm_impl_admw = @{thm adm_impl_admw};
|
|
25895
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
huffman
parents:
25805
diff
changeset
|
16 |
val adm_all = @{thm adm_all};
|
| 25805 | 17 |
val adm_conj = @{thm adm_conj};
|
18 |
val adm_subst = @{thm adm_subst};
|
|
19 |
val antisym_less_inverse = @{thm antisym_less_inverse};
|
|
20 |
val beta_cfun = @{thm beta_cfun};
|
|
21 |
val cfun_arg_cong = @{thm cfun_arg_cong};
|
|
22 |
val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL};
|
|
23 |
val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR};
|
|
24 |
val chain_iterate = @{thm chain_iterate};
|
|
25 |
val compact_ONE = @{thm compact_ONE};
|
|
26 |
val compact_sinl = @{thm compact_sinl};
|
|
27 |
val compact_sinr = @{thm compact_sinr};
|
|
28 |
val compact_spair = @{thm compact_spair};
|
|
29 |
val compact_up = @{thm compact_up};
|
|
30 |
val contlub_cfun_arg = @{thm contlub_cfun_arg};
|
|
31 |
val contlub_cfun_fun = @{thm contlub_cfun_fun};
|
|
32 |
val fix_def2 = @{thm fix_def2};
|
|
33 |
val injection_eq = @{thm injection_eq};
|
|
34 |
val injection_less = @{thm injection_less};
|
|
35 |
val lub_equal = @{thm lub_equal};
|
|
36 |
val monofun_cfun_arg = @{thm monofun_cfun_arg};
|
|
37 |
val retraction_strict = @{thm retraction_strict};
|
|
38 |
val spair_eq = @{thm spair_eq};
|
|
39 |
val spair_less = @{thm spair_less};
|
|
40 |
val sscase1 = @{thm sscase1};
|
|
41 |
val ssplit1 = @{thm ssplit1};
|
|
42 |
val strictify1 = @{thm strictify1};
|
|
43 |
val wfix_ind = @{thm wfix_ind};
|
|
44 |
||
45 |
val iso_intro = @{thm iso.intro};
|
|
46 |
val iso_abs_iso = @{thm iso.abs_iso};
|
|
47 |
val iso_rep_iso = @{thm iso.rep_iso};
|
|
48 |
val iso_abs_strict = @{thm iso.abs_strict};
|
|
49 |
val iso_rep_strict = @{thm iso.rep_strict};
|
|
50 |
val iso_abs_defin' = @{thm iso.abs_defin'};
|
|
51 |
val iso_rep_defin' = @{thm iso.rep_defin'};
|
|
52 |
val iso_abs_defined = @{thm iso.abs_defined};
|
|
53 |
val iso_rep_defined = @{thm iso.rep_defined};
|
|
54 |
val iso_compact_abs = @{thm iso.compact_abs};
|
|
55 |
val iso_compact_rep = @{thm iso.compact_rep};
|
|
56 |
val iso_iso_swap = @{thm iso.iso_swap};
|
|
57 |
||
58 |
val exh_start = @{thm exh_start};
|
|
59 |
val ex_defined_iffs = @{thms ex_defined_iffs};
|
|
60 |
val exh_casedist0 = @{thm exh_casedist0};
|
|
61 |
val exh_casedists = @{thms exh_casedists};
|
|
| 23152 | 62 |
|
63 |
open Domain_Library; |
|
64 |
infixr 0 ===>; |
|
65 |
infixr 0 ==>; |
|
66 |
infix 0 == ; |
|
67 |
infix 1 ===; |
|
68 |
infix 1 ~= ; |
|
69 |
infix 1 <<; |
|
70 |
infix 1 ~<<; |
|
71 |
infix 9 ` ; |
|
72 |
infix 9 `% ; |
|
73 |
infix 9 `%%; |
|
74 |
infixr 9 oo; |
|
75 |
||
76 |
(* ----- general proof facilities ------------------------------------------- *) |
|
77 |
||
| 24503 | 78 |
fun legacy_infer_term thy t = |
79 |
let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy) |
|
80 |
in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end; |
|
81 |
||
| 23152 | 82 |
fun pg'' thy defs t tacs = |
83 |
let |
|
| 24503 | 84 |
val t' = legacy_infer_term thy t; |
| 23152 | 85 |
val asms = Logic.strip_imp_prems t'; |
86 |
val prop = Logic.strip_imp_concl t'; |
|
| 26711 | 87 |
fun tac {prems, context} =
|
| 23152 | 88 |
rewrite_goals_tac defs THEN |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
89 |
EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
|
| 23152 | 90 |
in Goal.prove_global thy [] asms prop tac end; |
91 |
||
92 |
fun pg' thy defs t tacsf = |
|
93 |
let |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
94 |
fun tacs {prems, context} =
|
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
95 |
if null prems then tacsf context |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
96 |
else cut_facts_tac prems 1 :: tacsf context; |
| 23152 | 97 |
in pg'' thy defs t tacs end; |
98 |
||
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
99 |
fun case_UU_tac ctxt rews i v = |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
100 |
InductTacs.case_tac ctxt (v^"=UU") i THEN |
| 23152 | 101 |
asm_simp_tac (HOLCF_ss addsimps rews) i; |
102 |
||
103 |
val chain_tac = |
|
104 |
REPEAT_DETERM o resolve_tac |
|
105 |
[chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL]; |
|
106 |
||
107 |
(* ----- general proofs ----------------------------------------------------- *) |
|
108 |
||
109 |
val all2E = prove_goal HOL.thy "[| !x y . P x y; P x y ==> R |] ==> R" |
|
110 |
(fn prems =>[ |
|
111 |
resolve_tac prems 1, |
|
112 |
cut_facts_tac prems 1, |
|
113 |
fast_tac HOL_cs 1]); |
|
114 |
||
115 |
val dist_eqI = prove_goal (the_context ()) "!!x::'a::po. ~ x << y ==> x ~= y" |
|
116 |
(fn prems => |
|
| 26342 | 117 |
[blast_tac (@{claset} addDs [antisym_less_inverse]) 1]);
|
| 23152 | 118 |
|
119 |
in |
|
120 |
||
121 |
fun theorems (((dname, _), cons) : eq, eqs : eq list) thy = |
|
122 |
let |
|
123 |
||
124 |
val dummy = writeln ("Proving isomorphism properties of domain "^dname^" ...");
|
|
125 |
val pg = pg' thy; |
|
126 |
||
127 |
(* ----- getting the axioms and definitions --------------------------------- *) |
|
128 |
||
129 |
local |
|
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
130 |
fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); |
| 23152 | 131 |
in |
132 |
val ax_abs_iso = ga "abs_iso" dname; |
|
133 |
val ax_rep_iso = ga "rep_iso" dname; |
|
134 |
val ax_when_def = ga "when_def" dname; |
|
135 |
fun get_def mk_name (con,_) = ga (mk_name con^"_def") dname; |
|
136 |
val axs_con_def = map (get_def extern_name) cons; |
|
137 |
val axs_dis_def = map (get_def dis_name) cons; |
|
138 |
val axs_mat_def = map (get_def mat_name) cons; |
|
139 |
val axs_pat_def = map (get_def pat_name) cons; |
|
140 |
val axs_sel_def = |
|
141 |
let |
|
142 |
fun def_of_sel sel = ga (sel^"_def") dname; |
|
143 |
fun def_of_arg arg = Option.map def_of_sel (sel_of arg); |
|
144 |
fun defs_of_con (_, args) = List.mapPartial def_of_arg args; |
|
145 |
in |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
146 |
maps defs_of_con cons |
| 23152 | 147 |
end; |
148 |
val ax_copy_def = ga "copy_def" dname; |
|
149 |
end; (* local *) |
|
150 |
||
151 |
(* ----- theorems concerning the isomorphism -------------------------------- *) |
|
152 |
||
153 |
val dc_abs = %%:(dname^"_abs"); |
|
154 |
val dc_rep = %%:(dname^"_rep"); |
|
155 |
val dc_copy = %%:(dname^"_copy"); |
|
156 |
val x_name = "x"; |
|
157 |
||
158 |
val iso_locale = iso_intro OF [ax_abs_iso, ax_rep_iso]; |
|
159 |
val abs_strict = ax_rep_iso RS (allI RS retraction_strict); |
|
160 |
val rep_strict = ax_abs_iso RS (allI RS retraction_strict); |
|
161 |
val abs_defin' = iso_locale RS iso_abs_defin'; |
|
162 |
val rep_defin' = iso_locale RS iso_rep_defin'; |
|
163 |
val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict]; |
|
164 |
||
165 |
(* ----- generating beta reduction rules from definitions-------------------- *) |
|
166 |
||
167 |
local |
|
168 |
fun arglist (Const _ $ Abs (s, _, t)) = |
|
169 |
let |
|
170 |
val (vars,body) = arglist t; |
|
171 |
in (s :: vars, body) end |
|
172 |
| arglist t = ([], t); |
|
173 |
fun bind_fun vars t = Library.foldr mk_All (vars, t); |
|
174 |
fun bound_vars 0 = [] |
|
175 |
| bound_vars i = Bound (i-1) :: bound_vars (i - 1); |
|
176 |
in |
|
177 |
fun appl_of_def def = |
|
178 |
let |
|
179 |
val (_ $ con $ lam) = concl_of def; |
|
180 |
val (vars, rhs) = arglist lam; |
|
181 |
val lhs = list_ccomb (con, bound_vars (length vars)); |
|
182 |
val appl = bind_fun vars (lhs == rhs); |
|
183 |
val cs = ContProc.cont_thms lam; |
|
184 |
val betas = map (fn c => mk_meta_eq (c RS beta_cfun)) cs; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
185 |
in pg (def::betas) appl (K [rtac reflexive_thm 1]) end; |
| 23152 | 186 |
end; |
187 |
||
188 |
val when_appl = appl_of_def ax_when_def; |
|
189 |
val con_appls = map appl_of_def axs_con_def; |
|
190 |
||
191 |
local |
|
192 |
fun arg2typ n arg = |
|
193 |
let val t = TVar (("'a", n), pcpoS)
|
|
194 |
in (n + 1, if is_lazy arg then mk_uT t else t) end; |
|
195 |
||
196 |
fun args2typ n [] = (n, oneT) |
|
197 |
| args2typ n [arg] = arg2typ n arg |
|
198 |
| args2typ n (arg::args) = |
|
199 |
let |
|
200 |
val (n1, t1) = arg2typ n arg; |
|
201 |
val (n2, t2) = args2typ n1 args |
|
202 |
in (n2, mk_sprodT (t1, t2)) end; |
|
203 |
||
204 |
fun cons2typ n [] = (n,oneT) |
|
205 |
| cons2typ n [con] = args2typ n (snd con) |
|
206 |
| cons2typ n (con::cons) = |
|
207 |
let |
|
208 |
val (n1, t1) = args2typ n (snd con); |
|
209 |
val (n2, t2) = cons2typ n1 cons |
|
210 |
in (n2, mk_ssumT (t1, t2)) end; |
|
211 |
in |
|
212 |
fun cons2ctyp cons = ctyp_of thy (snd (cons2typ 1 cons)); |
|
213 |
end; |
|
214 |
||
215 |
local |
|
216 |
val iso_swap = iso_locale RS iso_iso_swap; |
|
217 |
fun one_con (con, args) = |
|
218 |
let |
|
219 |
val vns = map vname args; |
|
220 |
val eqn = %:x_name === con_app2 con %: vns; |
|
221 |
val conj = foldr1 mk_conj (eqn :: map (defined o %:) (nonlazy args)); |
|
222 |
in Library.foldr mk_ex (vns, conj) end; |
|
223 |
||
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23152
diff
changeset
|
224 |
val conj_assoc = @{thm conj_assoc};
|
| 23152 | 225 |
val exh = foldr1 mk_disj ((%:x_name === UU) :: map one_con cons); |
226 |
val thm1 = instantiate' [SOME (cons2ctyp cons)] [] exh_start; |
|
227 |
val thm2 = rewrite_rule (map mk_meta_eq ex_defined_iffs) thm1; |
|
| 25805 | 228 |
val thm3 = rewrite_rule [mk_meta_eq @{thm conj_assoc}] thm2;
|
| 23152 | 229 |
|
230 |
(* first 3 rules replace "x = UU \/ P" with "rep$x = UU \/ P" *) |
|
231 |
val tacs = [ |
|
232 |
rtac disjE 1, |
|
233 |
etac (rep_defin' RS disjI1) 2, |
|
234 |
etac disjI2 2, |
|
235 |
rewrite_goals_tac [mk_meta_eq iso_swap], |
|
236 |
rtac thm3 1]; |
|
237 |
in |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
238 |
val exhaust = pg con_appls (mk_trp exh) (K tacs); |
| 23152 | 239 |
val casedist = |
240 |
standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0)); |
|
241 |
end; |
|
242 |
||
243 |
local |
|
244 |
fun bind_fun t = Library.foldr mk_All (when_funs cons, t); |
|
245 |
fun bound_fun i _ = Bound (length cons - i); |
|
246 |
val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons); |
|
247 |
in |
|
248 |
val when_strict = |
|
249 |
let |
|
250 |
val axs = [when_appl, mk_meta_eq rep_strict]; |
|
251 |
val goal = bind_fun (mk_trp (strict when_app)); |
|
252 |
val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
253 |
in pg axs goal (K tacs) end; |
| 23152 | 254 |
|
255 |
val when_apps = |
|
256 |
let |
|
257 |
fun one_when n (con,args) = |
|
258 |
let |
|
259 |
val axs = when_appl :: con_appls; |
|
260 |
val goal = bind_fun (lift_defined %: (nonlazy args, |
|
261 |
mk_trp (when_app`(con_app con args) === |
|
262 |
list_ccomb (bound_fun n 0, map %# args)))); |
|
263 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
264 |
in pg axs goal (K tacs) end; |
| 23152 | 265 |
in mapn one_when 1 cons end; |
266 |
end; |
|
267 |
val when_rews = when_strict :: when_apps; |
|
268 |
||
269 |
(* ----- theorems concerning the constructors, discriminators and selectors - *) |
|
270 |
||
271 |
local |
|
272 |
fun dis_strict (con, _) = |
|
273 |
let |
|
274 |
val goal = mk_trp (strict (%%:(dis_name con))); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
275 |
in pg axs_dis_def goal (K [rtac when_strict 1]) end; |
| 23152 | 276 |
|
277 |
fun dis_app c (con, args) = |
|
278 |
let |
|
279 |
val lhs = %%:(dis_name c) ` con_app con args; |
|
| 26012 | 280 |
val rhs = if con = c then TT else FF; |
| 23152 | 281 |
val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); |
282 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
283 |
in pg axs_dis_def goal (K tacs) end; |
| 23152 | 284 |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
285 |
val dis_apps = maps (fn (c,_) => map (dis_app c) cons) cons; |
| 23152 | 286 |
|
287 |
fun dis_defin (con, args) = |
|
288 |
let |
|
289 |
val goal = defined (%:x_name) ==> defined (%%:(dis_name con) `% x_name); |
|
290 |
val tacs = |
|
291 |
[rtac casedist 1, |
|
292 |
contr_tac 1, |
|
293 |
DETERM_UNTIL_SOLVED (CHANGED |
|
294 |
(asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
295 |
in pg [] goal (K tacs) end; |
| 23152 | 296 |
|
297 |
val dis_stricts = map dis_strict cons; |
|
298 |
val dis_defins = map dis_defin cons; |
|
299 |
in |
|
300 |
val dis_rews = dis_stricts @ dis_defins @ dis_apps; |
|
301 |
end; |
|
302 |
||
303 |
local |
|
304 |
fun mat_strict (con, _) = |
|
305 |
let |
|
306 |
val goal = mk_trp (strict (%%:(mat_name con))); |
|
307 |
val tacs = [rtac when_strict 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
308 |
in pg axs_mat_def goal (K tacs) end; |
| 23152 | 309 |
|
310 |
val mat_stricts = map mat_strict cons; |
|
311 |
||
312 |
fun one_mat c (con, args) = |
|
313 |
let |
|
314 |
val lhs = %%:(mat_name c) ` con_app con args; |
|
315 |
val rhs = |
|
316 |
if con = c |
|
| 26012 | 317 |
then mk_return (mk_ctuple (map %# args)) |
318 |
else mk_fail; |
|
| 23152 | 319 |
val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); |
320 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
321 |
in pg axs_mat_def goal (K tacs) end; |
| 23152 | 322 |
|
323 |
val mat_apps = |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
324 |
maps (fn (c,_) => map (one_mat c) cons) cons; |
| 23152 | 325 |
in |
326 |
val mat_rews = mat_stricts @ mat_apps; |
|
327 |
end; |
|
328 |
||
329 |
local |
|
330 |
fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
|
|
331 |
||
| 26012 | 332 |
fun pat_lhs (con,args) = mk_branch (list_comb (%%:(pat_name con), ps args)); |
| 23152 | 333 |
|
| 26012 | 334 |
fun pat_rhs (con,[]) = mk_return ((%:"rhs") ` HOLogic.unit) |
| 23152 | 335 |
| pat_rhs (con,args) = |
| 26012 | 336 |
(mk_branch (mk_ctuple_pat (ps args))) |
| 23152 | 337 |
`(%:"rhs")`(mk_ctuple (map %# args)); |
338 |
||
339 |
fun pat_strict c = |
|
340 |
let |
|
| 25132 | 341 |
val axs = @{thm branch_def} :: axs_pat_def;
|
| 23152 | 342 |
val goal = mk_trp (strict (pat_lhs c ` (%:"rhs"))); |
343 |
val tacs = [simp_tac (HOLCF_ss addsimps [when_strict]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
344 |
in pg axs goal (K tacs) end; |
| 23152 | 345 |
|
346 |
fun pat_app c (con, args) = |
|
347 |
let |
|
| 25132 | 348 |
val axs = @{thm branch_def} :: axs_pat_def;
|
| 23152 | 349 |
val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args); |
| 26012 | 350 |
val rhs = if con = fst c then pat_rhs c else mk_fail; |
| 23152 | 351 |
val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); |
352 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
353 |
in pg axs goal (K tacs) end; |
| 23152 | 354 |
|
355 |
val pat_stricts = map pat_strict cons; |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
356 |
val pat_apps = maps (fn c => map (pat_app c) cons) cons; |
| 23152 | 357 |
in |
358 |
val pat_rews = pat_stricts @ pat_apps; |
|
359 |
end; |
|
360 |
||
361 |
local |
|
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23152
diff
changeset
|
362 |
val rev_contrapos = @{thm rev_contrapos};
|
| 23152 | 363 |
fun con_strict (con, args) = |
364 |
let |
|
365 |
fun one_strict vn = |
|
366 |
let |
|
367 |
fun f arg = if vname arg = vn then UU else %# arg; |
|
368 |
val goal = mk_trp (con_app2 con f args === UU); |
|
369 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
370 |
in pg con_appls goal (K tacs) end; |
| 23152 | 371 |
in map one_strict (nonlazy args) end; |
372 |
||
373 |
fun con_defin (con, args) = |
|
374 |
let |
|
375 |
val concl = mk_trp (defined (con_app con args)); |
|
376 |
val goal = lift_defined %: (nonlazy args, concl); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
377 |
fun tacs ctxt = [ |
| 25805 | 378 |
rtac @{thm rev_contrapos} 1,
|
| 27239 | 379 |
eres_inst_tac ctxt [(("f", 0), dis_name con)] cfun_arg_cong 1,
|
| 23152 | 380 |
asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]; |
381 |
in pg [] goal tacs end; |
|
382 |
in |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
383 |
val con_stricts = maps con_strict cons; |
| 23152 | 384 |
val con_defins = map con_defin cons; |
385 |
val con_rews = con_stricts @ con_defins; |
|
386 |
end; |
|
387 |
||
388 |
local |
|
389 |
val rules = |
|
390 |
[compact_sinl, compact_sinr, compact_spair, compact_up, compact_ONE]; |
|
391 |
fun con_compact (con, args) = |
|
392 |
let |
|
| 26012 | 393 |
val concl = mk_trp (mk_compact (con_app con args)); |
394 |
val goal = lift (fn x => mk_compact (%#x)) (args, concl); |
|
| 23152 | 395 |
val tacs = [ |
396 |
rtac (iso_locale RS iso_compact_abs) 1, |
|
397 |
REPEAT (resolve_tac rules 1 ORELSE atac 1)]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
398 |
in pg con_appls goal (K tacs) end; |
| 23152 | 399 |
in |
400 |
val con_compacts = map con_compact cons; |
|
401 |
end; |
|
402 |
||
403 |
local |
|
404 |
fun one_sel sel = |
|
405 |
pg axs_sel_def (mk_trp (strict (%%:sel))) |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
406 |
(K [simp_tac (HOLCF_ss addsimps when_rews) 1]); |
| 23152 | 407 |
|
408 |
fun sel_strict (_, args) = |
|
409 |
List.mapPartial (Option.map one_sel o sel_of) args; |
|
410 |
in |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
411 |
val sel_stricts = maps sel_strict cons; |
| 23152 | 412 |
end; |
413 |
||
414 |
local |
|
415 |
fun sel_app_same c n sel (con, args) = |
|
416 |
let |
|
417 |
val nlas = nonlazy args; |
|
418 |
val vns = map vname args; |
|
419 |
val vnn = List.nth (vns, n); |
|
420 |
val nlas' = List.filter (fn v => v <> vnn) nlas; |
|
421 |
val lhs = (%%:sel)`(con_app con args); |
|
422 |
val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn)); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
423 |
fun tacs1 ctxt = |
| 23152 | 424 |
if vnn mem nlas |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
425 |
then [case_UU_tac ctxt (when_rews @ con_stricts) 1 vnn] |
| 23152 | 426 |
else []; |
427 |
val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
428 |
in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; |
| 23152 | 429 |
|
430 |
fun sel_app_diff c n sel (con, args) = |
|
431 |
let |
|
432 |
val nlas = nonlazy args; |
|
433 |
val goal = mk_trp (%%:sel ` con_app con args === UU); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
434 |
fun tacs1 ctxt = map (case_UU_tac ctxt (when_rews @ con_stricts) 1) nlas; |
| 23152 | 435 |
val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
436 |
in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; |
| 23152 | 437 |
|
438 |
fun sel_app c n sel (con, args) = |
|
439 |
if con = c |
|
440 |
then sel_app_same c n sel (con, args) |
|
441 |
else sel_app_diff c n sel (con, args); |
|
442 |
||
443 |
fun one_sel c n sel = map (sel_app c n sel) cons; |
|
444 |
fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg); |
|
445 |
fun one_con (c, args) = |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
446 |
flat (map_filter I (mapn (one_sel' c) 0 args)); |
| 23152 | 447 |
in |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
448 |
val sel_apps = maps one_con cons; |
| 23152 | 449 |
end; |
450 |
||
451 |
local |
|
452 |
fun sel_defin sel = |
|
453 |
let |
|
454 |
val goal = defined (%:x_name) ==> defined (%%:sel`%x_name); |
|
455 |
val tacs = [ |
|
456 |
rtac casedist 1, |
|
457 |
contr_tac 1, |
|
458 |
DETERM_UNTIL_SOLVED (CHANGED |
|
459 |
(asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
460 |
in pg [] goal (K tacs) end; |
| 23152 | 461 |
in |
462 |
val sel_defins = |
|
463 |
if length cons = 1 |
|
464 |
then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg)) |
|
465 |
(filter_out is_lazy (snd (hd cons))) |
|
466 |
else []; |
|
467 |
end; |
|
468 |
||
469 |
val sel_rews = sel_stricts @ sel_defins @ sel_apps; |
|
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23152
diff
changeset
|
470 |
val rev_contrapos = @{thm rev_contrapos};
|
| 23152 | 471 |
|
472 |
val distincts_le = |
|
473 |
let |
|
474 |
fun dist (con1, args1) (con2, args2) = |
|
475 |
let |
|
476 |
val goal = lift_defined %: (nonlazy args1, |
|
477 |
mk_trp (con_app con1 args1 ~<< con_app con2 args2)); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
478 |
fun tacs ctxt = [ |
| 25805 | 479 |
rtac @{thm rev_contrapos} 1,
|
| 27239 | 480 |
eres_inst_tac ctxt [(("f", 0), dis_name con1)] monofun_cfun_arg 1]
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
481 |
@ map (case_UU_tac ctxt (con_stricts @ dis_rews) 1) (nonlazy args2) |
| 23152 | 482 |
@ [asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]; |
483 |
in pg [] goal tacs end; |
|
484 |
||
485 |
fun distinct (con1, args1) (con2, args2) = |
|
486 |
let |
|
487 |
val arg1 = (con1, args1); |
|
488 |
val arg2 = |
|
489 |
(con2, ListPair.map (fn (arg,vn) => upd_vname (K vn) arg) |
|
490 |
(args2, Name.variant_list (map vname args1) (map vname args2))); |
|
491 |
in [dist arg1 arg2, dist arg2 arg1] end; |
|
492 |
fun distincts [] = [] |
|
493 |
| distincts (c::cs) = (map (distinct c) cs) :: distincts cs; |
|
494 |
in distincts cons end; |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
495 |
val dist_les = flat (flat distincts_le); |
| 23152 | 496 |
val dist_eqs = |
497 |
let |
|
498 |
fun distinct (_,args1) ((_,args2), leqs) = |
|
499 |
let |
|
500 |
val (le1,le2) = (hd leqs, hd(tl leqs)); |
|
501 |
val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) |
|
502 |
in |
|
503 |
if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else |
|
504 |
if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else |
|
505 |
[eq1, eq2] |
|
506 |
end; |
|
507 |
fun distincts [] = [] |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
508 |
| distincts ((c,leqs)::cs) = flat |
| 23152 | 509 |
(ListPair.map (distinct c) ((map #1 cs),leqs)) @ |
510 |
distincts cs; |
|
511 |
in map standard (distincts (cons ~~ distincts_le)) end; |
|
512 |
||
513 |
local |
|
514 |
fun pgterm rel con args = |
|
515 |
let |
|
516 |
fun append s = upd_vname (fn v => v^s); |
|
517 |
val (largs, rargs) = (args, map (append "'") args); |
|
518 |
val concl = |
|
519 |
foldr1 mk_conj (ListPair.map rel (map %# largs, map %# rargs)); |
|
520 |
val prem = rel (con_app con largs, con_app con rargs); |
|
521 |
val sargs = case largs of [_] => [] | _ => nonlazy args; |
|
522 |
val prop = lift_defined %: (sargs, mk_trp (prem === concl)); |
|
523 |
in pg con_appls prop end; |
|
524 |
val cons' = List.filter (fn (_,args) => args<>[]) cons; |
|
525 |
in |
|
526 |
val inverts = |
|
527 |
let |
|
528 |
val abs_less = ax_abs_iso RS (allI RS injection_less); |
|
529 |
val tacs = |
|
530 |
[asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
531 |
in map (fn (con, args) => pgterm (op <<) con args (K tacs)) cons' end; |
| 23152 | 532 |
|
533 |
val injects = |
|
534 |
let |
|
535 |
val abs_eq = ax_abs_iso RS (allI RS injection_eq); |
|
536 |
val tacs = [asm_full_simp_tac (HOLCF_ss addsimps [abs_eq, spair_eq]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
537 |
in map (fn (con, args) => pgterm (op ===) con args (K tacs)) cons' end; |
| 23152 | 538 |
end; |
539 |
||
540 |
(* ----- theorems concerning one induction step ----------------------------- *) |
|
541 |
||
542 |
val copy_strict = |
|
543 |
let |
|
544 |
val goal = mk_trp (strict (dc_copy `% "f")); |
|
545 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps [abs_strict, when_strict]) 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
546 |
in pg [ax_copy_def] goal (K tacs) end; |
| 23152 | 547 |
|
548 |
local |
|
549 |
fun copy_app (con, args) = |
|
550 |
let |
|
551 |
val lhs = dc_copy`%"f"`(con_app con args); |
|
552 |
val rhs = con_app2 con (app_rec_arg (cproj (%:"f") eqs)) args; |
|
553 |
val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); |
|
554 |
val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args; |
|
555 |
val stricts = abs_strict::when_strict::con_stricts; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
556 |
fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args'; |
| 23152 | 557 |
val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_apps) 1]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
558 |
in pg [ax_copy_def] goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; |
| 23152 | 559 |
in |
560 |
val copy_apps = map copy_app cons; |
|
561 |
end; |
|
562 |
||
563 |
local |
|
564 |
fun one_strict (con, args) = |
|
565 |
let |
|
566 |
val goal = mk_trp (dc_copy`UU`(con_app con args) === UU); |
|
567 |
val rews = copy_strict :: copy_apps @ con_rews; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
568 |
fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @ |
| 23152 | 569 |
[asm_simp_tac (HOLCF_ss addsimps rews) 1]; |
570 |
in pg [] goal tacs end; |
|
571 |
||
572 |
fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args; |
|
573 |
in |
|
574 |
val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons); |
|
575 |
end; |
|
576 |
||
577 |
val copy_rews = copy_strict :: copy_apps @ copy_stricts; |
|
578 |
||
579 |
in |
|
580 |
thy |
|
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
581 |
|> Sign.add_path (Sign.base_name dname) |
| 23152 | 582 |
|> (snd o (PureThy.add_thmss (map Thm.no_attributes [ |
583 |
("iso_rews" , iso_rews ),
|
|
584 |
("exhaust" , [exhaust] ),
|
|
585 |
("casedist" , [casedist]),
|
|
586 |
("when_rews", when_rews ),
|
|
587 |
("compacts", con_compacts),
|
|
588 |
("con_rews", con_rews),
|
|
589 |
("sel_rews", sel_rews),
|
|
590 |
("dis_rews", dis_rews),
|
|
591 |
("pat_rews", pat_rews),
|
|
592 |
("dist_les", dist_les),
|
|
593 |
("dist_eqs", dist_eqs),
|
|
594 |
("inverts" , inverts ),
|
|
595 |
("injects" , injects ),
|
|
596 |
("copy_rews", copy_rews)])))
|
|
597 |
|> (snd o PureThy.add_thmss |
|
598 |
[(("match_rews", mat_rews), [Simplifier.simp_add])])
|
|
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
599 |
|> Sign.parent_path |
| 23152 | 600 |
|> rpair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ |
601 |
pat_rews @ dist_les @ dist_eqs @ copy_rews) |
|
602 |
end; (* let *) |
|
603 |
||
604 |
fun comp_theorems (comp_dnam, eqs: eq list) thy = |
|
605 |
let |
|
| 27232 | 606 |
val global_ctxt = ProofContext.init thy; |
607 |
||
| 23152 | 608 |
val dnames = map (fst o fst) eqs; |
609 |
val conss = map snd eqs; |
|
610 |
val comp_dname = Sign.full_name thy comp_dnam; |
|
611 |
||
612 |
val d = writeln("Proving induction properties of domain "^comp_dname^" ...");
|
|
613 |
val pg = pg' thy; |
|
614 |
||
615 |
(* ----- getting the composite axiom and definitions ------------------------ *) |
|
616 |
||
617 |
local |
|
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
618 |
fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); |
| 23152 | 619 |
in |
620 |
val axs_reach = map (ga "reach" ) dnames; |
|
621 |
val axs_take_def = map (ga "take_def" ) dnames; |
|
622 |
val axs_finite_def = map (ga "finite_def") dnames; |
|
623 |
val ax_copy2_def = ga "copy_def" comp_dnam; |
|
624 |
val ax_bisim_def = ga "bisim_def" comp_dnam; |
|
625 |
end; |
|
626 |
||
627 |
local |
|
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
628 |
fun gt s dn = PureThy.get_thm thy (dn ^ "." ^ s); |
|
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
629 |
fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s); |
| 23152 | 630 |
in |
631 |
val cases = map (gt "casedist" ) dnames; |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
632 |
val con_rews = maps (gts "con_rews" ) dnames; |
|
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
633 |
val copy_rews = maps (gts "copy_rews") dnames; |
| 23152 | 634 |
end; |
635 |
||
636 |
fun dc_take dn = %%:(dn^"_take"); |
|
637 |
val x_name = idx_name dnames "x"; |
|
638 |
val P_name = idx_name dnames "P"; |
|
639 |
val n_eqs = length eqs; |
|
640 |
||
641 |
(* ----- theorems concerning finite approximation and finite induction ------ *) |
|
642 |
||
643 |
local |
|
| 26342 | 644 |
val iterate_Cprod_ss = simpset_of @{theory Fix};
|
| 23152 | 645 |
val copy_con_rews = copy_rews @ con_rews; |
646 |
val copy_take_defs = |
|
647 |
(if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def; |
|
648 |
val take_stricts = |
|
649 |
let |
|
650 |
fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n"); |
|
651 |
val goal = mk_trp (foldr1 mk_conj (map one_eq eqs)); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
652 |
fun tacs ctxt = [ |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
653 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1, |
| 23152 | 654 |
simp_tac iterate_Cprod_ss 1, |
655 |
asm_simp_tac (iterate_Cprod_ss addsimps copy_rews) 1]; |
|
656 |
in pg copy_take_defs goal tacs end; |
|
657 |
||
658 |
val take_stricts' = rewrite_rule copy_take_defs take_stricts; |
|
659 |
fun take_0 n dn = |
|
660 |
let |
|
661 |
val goal = mk_trp ((dc_take dn $ %%:"HOL.zero") `% x_name n === UU); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
662 |
in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end; |
| 23152 | 663 |
val take_0s = mapn take_0 1 dnames; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
664 |
fun c_UU_tac ctxt = case_UU_tac ctxt (take_stricts'::copy_con_rews) 1; |
| 23152 | 665 |
val take_apps = |
666 |
let |
|
667 |
fun mk_eqn dn (con, args) = |
|
668 |
let |
|
669 |
fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n"; |
|
670 |
val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args); |
|
671 |
val rhs = con_app2 con (app_rec_arg mk_take) args; |
|
672 |
in Library.foldr mk_all (map vname args, lhs === rhs) end; |
|
673 |
fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons; |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
674 |
val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs)); |
| 23152 | 675 |
val simps = List.filter (has_fewer_prems 1) copy_rews; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
676 |
fun con_tac ctxt (con, args) = |
| 23152 | 677 |
if nonlazy_rec args = [] |
678 |
then all_tac |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
679 |
else EVERY (map (c_UU_tac ctxt) (nonlazy_rec args)) THEN |
| 23152 | 680 |
asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
681 |
fun eq_tacs ctxt ((dn, _), cons) = map (con_tac ctxt) cons; |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
682 |
fun tacs ctxt = |
| 23152 | 683 |
simp_tac iterate_Cprod_ss 1 :: |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
684 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: |
| 23152 | 685 |
simp_tac (iterate_Cprod_ss addsimps copy_con_rews) 1 :: |
686 |
asm_full_simp_tac (HOLCF_ss addsimps simps) 1 :: |
|
687 |
TRY (safe_tac HOL_cs) :: |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
688 |
maps (eq_tacs ctxt) eqs; |
| 23152 | 689 |
in pg copy_take_defs goal tacs end; |
690 |
in |
|
691 |
val take_rews = map standard |
|
| 27232 | 692 |
(atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps); |
| 23152 | 693 |
end; (* local *) |
694 |
||
695 |
local |
|
696 |
fun one_con p (con,args) = |
|
697 |
let |
|
698 |
fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; |
|
699 |
val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); |
|
700 |
val t2 = lift ind_hyp (List.filter is_rec args, t1); |
|
701 |
val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2); |
|
702 |
in Library.foldr mk_All (map vname args, t3) end; |
|
703 |
||
704 |
fun one_eq ((p, cons), concl) = |
|
705 |
mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl); |
|
706 |
||
707 |
fun ind_term concf = Library.foldr one_eq |
|
708 |
(mapn (fn n => fn x => (P_name n, x)) 1 conss, |
|
709 |
mk_trp (foldr1 mk_conj (mapn concf 1 dnames))); |
|
710 |
val take_ss = HOL_ss addsimps take_rews; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
711 |
fun quant_tac ctxt i = EVERY |
| 27239 | 712 |
(mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames);
|
| 23152 | 713 |
|
714 |
fun ind_prems_tac prems = EVERY |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
715 |
(maps (fn cons => |
| 23152 | 716 |
(resolve_tac prems 1 :: |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
717 |
maps (fn (_,args) => |
| 23152 | 718 |
resolve_tac prems 1 :: |
719 |
map (K(atac 1)) (nonlazy args) @ |
|
720 |
map (K(atac 1)) (List.filter is_rec args)) |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
721 |
cons)) |
|
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
722 |
conss); |
| 23152 | 723 |
local |
724 |
(* check whether every/exists constructor of the n-th part of the equation: |
|
725 |
it has a possibly indirectly recursive argument that isn't/is possibly |
|
726 |
indirectly lazy *) |
|
727 |
fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => |
|
728 |
is_rec arg andalso not(rec_of arg mem ns) andalso |
|
729 |
((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse |
|
730 |
rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) |
|
731 |
(lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) |
|
732 |
) o snd) cons; |
|
733 |
fun all_rec_to ns = rec_to forall not all_rec_to ns; |
|
734 |
fun warn (n,cons) = |
|
735 |
if all_rec_to [] false (n,cons) |
|
736 |
then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
|
|
737 |
else false; |
|
738 |
fun lazy_rec_to ns = rec_to exists I lazy_rec_to ns; |
|
739 |
||
740 |
in |
|
741 |
val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; |
|
742 |
val is_emptys = map warn n__eqs; |
|
743 |
val is_finite = forall (not o lazy_rec_to [] false) n__eqs; |
|
744 |
end; |
|
745 |
in (* local *) |
|
746 |
val finite_ind = |
|
747 |
let |
|
748 |
fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n)); |
|
749 |
val goal = ind_term concf; |
|
750 |
||
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
751 |
fun tacf {prems, context} =
|
| 23152 | 752 |
let |
753 |
val tacs1 = [ |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
754 |
quant_tac context 1, |
| 23152 | 755 |
simp_tac HOL_ss 1, |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
756 |
InductTacs.induct_tac context [[SOME "n"]] 1, |
| 23152 | 757 |
simp_tac (take_ss addsimps prems) 1, |
758 |
TRY (safe_tac HOL_cs)]; |
|
759 |
fun arg_tac arg = |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
760 |
case_UU_tac context (prems @ con_rews) 1 |
| 23152 | 761 |
(List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); |
762 |
fun con_tacs (con, args) = |
|
763 |
asm_simp_tac take_ss 1 :: |
|
764 |
map arg_tac (List.filter is_nonlazy_rec args) @ |
|
765 |
[resolve_tac prems 1] @ |
|
766 |
map (K (atac 1)) (nonlazy args) @ |
|
767 |
map (K (etac spec 1)) (List.filter is_rec args); |
|
768 |
fun cases_tacs (cons, cases) = |
|
| 27239 | 769 |
res_inst_tac context [(("x", 0), "x")] cases 1 ::
|
| 23152 | 770 |
asm_simp_tac (take_ss addsimps prems) 1 :: |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
771 |
maps con_tacs cons; |
| 23152 | 772 |
in |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
773 |
tacs1 @ maps cases_tacs (conss ~~ cases) |
| 23152 | 774 |
end; |
775 |
in pg'' thy [] goal tacf end; |
|
776 |
||
777 |
val take_lemmas = |
|
778 |
let |
|
779 |
fun take_lemma n (dn, ax_reach) = |
|
780 |
let |
|
781 |
val lhs = dc_take dn $ Bound 0 `%(x_name n); |
|
782 |
val rhs = dc_take dn $ Bound 0 `%(x_name n^"'"); |
|
783 |
val concl = mk_trp (%:(x_name n) === %:(x_name n^"'")); |
|
784 |
val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl;
|
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
785 |
fun tacf {prems, context} = [
|
| 27239 | 786 |
res_inst_tac context [(("t", 0), x_name n )] (ax_reach RS subst) 1,
|
787 |
res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1,
|
|
| 23152 | 788 |
stac fix_def2 1, |
789 |
REPEAT (CHANGED |
|
790 |
(rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)), |
|
791 |
stac contlub_cfun_fun 1, |
|
792 |
stac contlub_cfun_fun 2, |
|
793 |
rtac lub_equal 3, |
|
794 |
chain_tac 1, |
|
795 |
rtac allI 1, |
|
796 |
resolve_tac prems 1]; |
|
797 |
in pg'' thy axs_take_def goal tacf end; |
|
798 |
in mapn take_lemma 1 (dnames ~~ axs_reach) end; |
|
799 |
||
800 |
(* ----- theorems concerning finiteness and induction ----------------------- *) |
|
801 |
||
802 |
val (finites, ind) = |
|
803 |
if is_finite |
|
804 |
then (* finite case *) |
|
805 |
let |
|
806 |
fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x");
|
|
807 |
fun dname_lemma dn = |
|
808 |
let |
|
809 |
val prem1 = mk_trp (defined (%:"x")); |
|
810 |
val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
|
|
811 |
val prem2 = mk_trp (mk_disj (disj1, take_enough dn)); |
|
812 |
val concl = mk_trp (take_enough dn); |
|
813 |
val goal = prem1 ===> prem2 ===> concl; |
|
814 |
val tacs = [ |
|
815 |
etac disjE 1, |
|
816 |
etac notE 1, |
|
817 |
resolve_tac take_lemmas 1, |
|
818 |
asm_simp_tac take_ss 1, |
|
819 |
atac 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
820 |
in pg [] goal (K tacs) end; |
| 23152 | 821 |
val finite_lemmas1a = map dname_lemma dnames; |
822 |
||
823 |
val finite_lemma1b = |
|
824 |
let |
|
825 |
fun mk_eqn n ((dn, args), _) = |
|
826 |
let |
|
827 |
val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU; |
|
828 |
val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0; |
|
829 |
in |
|
830 |
mk_constrainall |
|
831 |
(x_name n, Type (dn,args), mk_disj (disj1, disj2)) |
|
832 |
end; |
|
833 |
val goal = |
|
834 |
mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs)));
|
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
835 |
fun arg_tacs ctxt vn = [ |
| 27239 | 836 |
eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1,
|
| 23152 | 837 |
etac disjE 1, |
838 |
asm_simp_tac (HOL_ss addsimps con_rews) 1, |
|
839 |
asm_simp_tac take_ss 1]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
840 |
fun con_tacs ctxt (con, args) = |
| 23152 | 841 |
asm_simp_tac take_ss 1 :: |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
842 |
maps (arg_tacs ctxt) (nonlazy_rec args); |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
843 |
fun foo_tacs ctxt n (cons, cases) = |
| 23152 | 844 |
simp_tac take_ss 1 :: |
845 |
rtac allI 1 :: |
|
| 27239 | 846 |
res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 ::
|
| 23152 | 847 |
asm_simp_tac take_ss 1 :: |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
848 |
maps (con_tacs ctxt) cons; |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
849 |
fun tacs ctxt = |
| 23152 | 850 |
rtac allI 1 :: |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
851 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: |
| 23152 | 852 |
simp_tac take_ss 1 :: |
853 |
TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) :: |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
854 |
flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases)); |
| 23152 | 855 |
in pg [] goal tacs end; |
856 |
||
857 |
fun one_finite (dn, l1b) = |
|
858 |
let |
|
859 |
val goal = mk_trp (%%:(dn^"_finite") $ %:"x"); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
860 |
fun tacs ctxt = [ |
|
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
861 |
case_UU_tac ctxt take_rews 1 "x", |
| 23152 | 862 |
eresolve_tac finite_lemmas1a 1, |
863 |
step_tac HOL_cs 1, |
|
864 |
step_tac HOL_cs 1, |
|
865 |
cut_facts_tac [l1b] 1, |
|
866 |
fast_tac HOL_cs 1]; |
|
867 |
in pg axs_finite_def goal tacs end; |
|
868 |
||
| 27232 | 869 |
val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b); |
| 23152 | 870 |
val ind = |
871 |
let |
|
872 |
fun concf n dn = %:(P_name n) $ %:(x_name n); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
873 |
fun tacf {prems, context} =
|
| 23152 | 874 |
let |
875 |
fun finite_tacs (finite, fin_ind) = [ |
|
876 |
rtac(rewrite_rule axs_finite_def finite RS exE)1, |
|
877 |
etac subst 1, |
|
878 |
rtac fin_ind 1, |
|
879 |
ind_prems_tac prems]; |
|
880 |
in |
|
881 |
TRY (safe_tac HOL_cs) :: |
|
| 27232 | 882 |
maps finite_tacs (finites ~~ atomize global_ctxt finite_ind) |
| 23152 | 883 |
end; |
884 |
in pg'' thy [] (ind_term concf) tacf end; |
|
885 |
in (finites, ind) end (* let *) |
|
886 |
||
887 |
else (* infinite case *) |
|
888 |
let |
|
889 |
fun one_finite n dn = |
|
| 27239 | 890 |
read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle;
|
| 23152 | 891 |
val finites = mapn one_finite 1 dnames; |
892 |
||
893 |
val goal = |
|
894 |
let |
|
| 26012 | 895 |
fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); |
| 23152 | 896 |
fun concf n dn = %:(P_name n) $ %:(x_name n); |
897 |
in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
898 |
fun tacf {prems, context} =
|
| 23152 | 899 |
map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [ |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
900 |
quant_tac context 1, |
| 23152 | 901 |
rtac (adm_impl_admw RS wfix_ind) 1, |
|
25895
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
huffman
parents:
25805
diff
changeset
|
902 |
REPEAT_DETERM (rtac adm_all 1), |
| 23152 | 903 |
REPEAT_DETERM ( |
904 |
TRY (rtac adm_conj 1) THEN |
|
905 |
rtac adm_subst 1 THEN |
|
906 |
cont_tacR 1 THEN resolve_tac prems 1), |
|
907 |
strip_tac 1, |
|
908 |
rtac (rewrite_rule axs_take_def finite_ind) 1, |
|
909 |
ind_prems_tac prems]; |
|
910 |
val ind = (pg'' thy [] goal tacf |
|
911 |
handle ERROR _ => |
|
912 |
(warning "Cannot prove infinite induction rule"; refl)); |
|
913 |
in (finites, ind) end; |
|
914 |
end; (* local *) |
|
915 |
||
916 |
(* ----- theorem concerning coinduction ------------------------------------- *) |
|
917 |
||
918 |
local |
|
919 |
val xs = mapn (fn n => K (x_name n)) 1 dnames; |
|
920 |
fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); |
|
921 |
val take_ss = HOL_ss addsimps take_rews; |
|
922 |
val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
|
|
923 |
val coind_lemma = |
|
924 |
let |
|
925 |
fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1; |
|
926 |
fun mk_eqn n dn = |
|
927 |
(dc_take dn $ %:"n" ` bnd_arg n 0) === |
|
928 |
(dc_take dn $ %:"n" ` bnd_arg n 1); |
|
929 |
fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t)); |
|
930 |
val goal = |
|
931 |
mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R", |
|
932 |
Library.foldr mk_all2 (xs, |
|
933 |
Library.foldr mk_imp (mapn mk_prj 0 dnames, |
|
934 |
foldr1 mk_conj (mapn mk_eqn 0 dnames))))); |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
935 |
fun x_tacs ctxt n x = [ |
| 23152 | 936 |
rotate_tac (n+1) 1, |
937 |
etac all2E 1, |
|
| 27239 | 938 |
eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
|
| 23152 | 939 |
TRY (safe_tac HOL_cs), |
940 |
REPEAT (CHANGED (asm_simp_tac take_ss 1))]; |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
941 |
fun tacs ctxt = [ |
| 23152 | 942 |
rtac impI 1, |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
943 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1, |
| 23152 | 944 |
simp_tac take_ss 1, |
945 |
safe_tac HOL_cs] @ |
|
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
946 |
flat (mapn (x_tacs ctxt) 0 xs); |
| 23152 | 947 |
in pg [ax_bisim_def] goal tacs end; |
948 |
in |
|
949 |
val coind = |
|
950 |
let |
|
951 |
fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'")); |
|
952 |
fun mk_eqn x = %:x === %:(x^"'"); |
|
953 |
val goal = |
|
954 |
mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===> |
|
955 |
Logic.list_implies (mapn mk_prj 0 xs, |
|
956 |
mk_trp (foldr1 mk_conj (map mk_eqn xs))); |
|
957 |
val tacs = |
|
958 |
TRY (safe_tac HOL_cs) :: |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
959 |
maps (fn take_lemma => [ |
| 23152 | 960 |
rtac take_lemma 1, |
961 |
cut_facts_tac [coind_lemma] 1, |
|
962 |
fast_tac HOL_cs 1]) |
|
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
963 |
take_lemmas; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
964 |
in pg [] goal (K tacs) end; |
| 23152 | 965 |
end; (* local *) |
966 |
||
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
967 |
in thy |> Sign.add_path comp_dnam |
| 23152 | 968 |
|> (snd o (PureThy.add_thmss (map Thm.no_attributes [ |
969 |
("take_rews" , take_rews ),
|
|
970 |
("take_lemmas", take_lemmas),
|
|
971 |
("finites" , finites ),
|
|
972 |
("finite_ind", [finite_ind]),
|
|
973 |
("ind" , [ind ]),
|
|
974 |
("coind" , [coind ])])))
|
|
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
975 |
|> Sign.parent_path |> rpair take_rews |
| 23152 | 976 |
end; (* let *) |
977 |
end; (* local *) |
|
978 |
end; (* struct *) |