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