author | huffman |
Sun, 28 Feb 2010 20:56:28 -0800 | |
changeset 35481 | 7bb9157507a9 |
parent 35468 | 09bc6a2e2296 |
child 35482 | d756837b708d |
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 |
|
35444 | 12 |
val theorems: |
13 |
Domain_Library.eq * Domain_Library.eq list |
|
14 |
-> typ * (binding * (bool * binding option * typ) list * mixfix) list |
|
15 |
-> theory -> thm list * theory; |
|
16 |
||
31005 | 17 |
val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory; |
32740 | 18 |
val quiet_mode: bool Unsynchronized.ref; |
19 |
val trace_domain: bool Unsynchronized.ref; |
|
31005 | 20 |
end; |
21 |
||
31023 | 22 |
structure Domain_Theorems :> DOMAIN_THEOREMS = |
31005 | 23 |
struct |
23152 | 24 |
|
32740 | 25 |
val quiet_mode = Unsynchronized.ref false; |
26 |
val trace_domain = Unsynchronized.ref false; |
|
29402 | 27 |
|
28 |
fun message s = if !quiet_mode then () else writeln s; |
|
29 |
fun trace s = if !trace_domain then tracing s else (); |
|
30 |
||
25805 | 31 |
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
|
32 |
val adm_all = @{thm adm_all}; |
25805 | 33 |
val adm_conj = @{thm adm_conj}; |
34 |
val adm_subst = @{thm adm_subst}; |
|
33396 | 35 |
val ch2ch_fst = @{thm ch2ch_fst}; |
36 |
val ch2ch_snd = @{thm ch2ch_snd}; |
|
25805 | 37 |
val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL}; |
38 |
val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR}; |
|
39 |
val chain_iterate = @{thm chain_iterate}; |
|
40 |
val contlub_cfun_fun = @{thm contlub_cfun_fun}; |
|
33396 | 41 |
val contlub_fst = @{thm contlub_fst}; |
42 |
val contlub_snd = @{thm contlub_snd}; |
|
43 |
val contlubE = @{thm contlubE}; |
|
44 |
val cont_const = @{thm cont_const}; |
|
45 |
val cont_id = @{thm cont_id}; |
|
46 |
val cont2cont_fst = @{thm cont2cont_fst}; |
|
47 |
val cont2cont_snd = @{thm cont2cont_snd}; |
|
48 |
val cont2cont_Rep_CFun = @{thm cont2cont_Rep_CFun}; |
|
25805 | 49 |
val fix_def2 = @{thm fix_def2}; |
50 |
val lub_equal = @{thm lub_equal}; |
|
51 |
val retraction_strict = @{thm retraction_strict}; |
|
52 |
val wfix_ind = @{thm wfix_ind}; |
|
35464 | 53 |
val iso_intro = @{thm iso.intro}; |
23152 | 54 |
|
55 |
open Domain_Library; |
|
56 |
infixr 0 ===>; |
|
57 |
infixr 0 ==>; |
|
58 |
infix 0 == ; |
|
59 |
infix 1 ===; |
|
60 |
infix 1 ~= ; |
|
61 |
infix 1 <<; |
|
62 |
infix 1 ~<<; |
|
63 |
infix 9 ` ; |
|
64 |
infix 9 `% ; |
|
65 |
infix 9 `%%; |
|
66 |
infixr 9 oo; |
|
67 |
||
68 |
(* ----- general proof facilities ------------------------------------------- *) |
|
69 |
||
24503 | 70 |
fun legacy_infer_term thy t = |
71 |
let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy) |
|
72 |
in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end; |
|
73 |
||
23152 | 74 |
fun pg'' thy defs t tacs = |
75 |
let |
|
24503 | 76 |
val t' = legacy_infer_term thy t; |
23152 | 77 |
val asms = Logic.strip_imp_prems t'; |
78 |
val prop = Logic.strip_imp_concl t'; |
|
26711 | 79 |
fun tac {prems, context} = |
23152 | 80 |
rewrite_goals_tac defs THEN |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
81 |
EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context}) |
23152 | 82 |
in Goal.prove_global thy [] asms prop tac end; |
83 |
||
84 |
fun pg' thy defs t tacsf = |
|
85 |
let |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
86 |
fun tacs {prems, context} = |
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
87 |
if null prems then tacsf context |
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
88 |
else cut_facts_tac prems 1 :: tacsf context; |
23152 | 89 |
in pg'' thy defs t tacs end; |
90 |
||
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
91 |
(* FIXME!!!!!!!!! *) |
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
92 |
(* We should NEVER re-parse variable names as strings! *) |
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
93 |
(* The names can conflict with existing constants or other syntax! *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
94 |
fun case_UU_tac ctxt rews i v = |
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
95 |
InductTacs.case_tac ctxt (v^"=UU") i THEN |
23152 | 96 |
asm_simp_tac (HOLCF_ss addsimps rews) i; |
97 |
||
98 |
val chain_tac = |
|
99 |
REPEAT_DETERM o resolve_tac |
|
33396 | 100 |
[chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL, ch2ch_fst, ch2ch_snd]; |
23152 | 101 |
|
102 |
(* ----- general proofs ----------------------------------------------------- *) |
|
103 |
||
29064 | 104 |
val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp} |
23152 | 105 |
|
35444 | 106 |
fun theorems |
107 |
(((dname, _), cons) : eq, eqs : eq list) |
|
108 |
(dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list) |
|
109 |
(thy : theory) = |
|
23152 | 110 |
let |
111 |
||
29402 | 112 |
val _ = message ("Proving isomorphism properties of domain "^dname^" ..."); |
33801 | 113 |
val map_tab = Domain_Isomorphism.get_map_tab thy; |
114 |
||
23152 | 115 |
|
116 |
(* ----- getting the axioms and definitions --------------------------------- *) |
|
117 |
||
118 |
local |
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
119 |
fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); |
23152 | 120 |
in |
121 |
val ax_abs_iso = ga "abs_iso" dname; |
|
122 |
val ax_rep_iso = ga "rep_iso" dname; |
|
123 |
val ax_when_def = ga "when_def" dname; |
|
124 |
val ax_copy_def = ga "copy_def" dname; |
|
125 |
end; (* local *) |
|
126 |
||
35444 | 127 |
(* ----- define constructors ------------------------------------------------ *) |
128 |
||
129 |
val lhsT = fst dom_eqn; |
|
130 |
||
131 |
val rhsT = |
|
132 |
let |
|
133 |
fun mk_arg_typ (lazy, sel, T) = if lazy then mk_uT T else T; |
|
134 |
fun mk_con_typ (bind, args, mx) = |
|
135 |
if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args); |
|
136 |
fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons); |
|
137 |
in |
|
138 |
mk_eq_typ dom_eqn |
|
139 |
end; |
|
140 |
||
141 |
val rep_const = Const(dname^"_rep", lhsT ->> rhsT); |
|
142 |
||
143 |
val abs_const = Const(dname^"_abs", rhsT ->> lhsT); |
|
144 |
||
35481
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
145 |
val iso_info : Domain_Isomorphism.iso_info = |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
146 |
{ |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
147 |
absT = lhsT, |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
148 |
repT = rhsT, |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
149 |
abs_const = abs_const, |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
150 |
rep_const = rep_const, |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
151 |
abs_inverse = ax_abs_iso, |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
152 |
rep_inverse = ax_rep_iso |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
153 |
}; |
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
154 |
|
35444 | 155 |
val (result, thy) = |
156 |
Domain_Constructors.add_domain_constructors |
|
35481
7bb9157507a9
add_domain_constructors takes iso_info record as argument
huffman
parents:
35468
diff
changeset
|
157 |
(Long_Name.base_name dname) (snd dom_eqn) iso_info ax_when_def thy; |
35444 | 158 |
|
35451 | 159 |
val con_appls = #con_betas result; |
35457
d63655b88369
move proofs of casedist into domain_constructors.ML
huffman
parents:
35456
diff
changeset
|
160 |
val {exhaust, casedist, ...} = result; |
35458
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents:
35457
diff
changeset
|
161 |
val {con_compacts, con_rews, inverts, injects, dist_les, dist_eqs, ...} = result; |
35457
d63655b88369
move proofs of casedist into domain_constructors.ML
huffman
parents:
35456
diff
changeset
|
162 |
val {sel_rews, ...} = result; |
35459
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
huffman
parents:
35458
diff
changeset
|
163 |
val when_rews = #cases result; |
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
huffman
parents:
35458
diff
changeset
|
164 |
val when_strict = hd when_rews; |
35461
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents:
35460
diff
changeset
|
165 |
val dis_rews = #dis_rews result; |
35466
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
huffman
parents:
35464
diff
changeset
|
166 |
val mat_rews = #match_rews result; |
35468
09bc6a2e2296
move definition and syntax of pattern combinators into domain_constructors.ML
huffman
parents:
35466
diff
changeset
|
167 |
val axs_pat_def = #pat_rews result; |
35444 | 168 |
|
23152 | 169 |
(* ----- theorems concerning the isomorphism -------------------------------- *) |
170 |
||
35444 | 171 |
val pg = pg' thy; |
172 |
||
23152 | 173 |
val dc_copy = %%:(dname^"_copy"); |
174 |
||
175 |
val abs_strict = ax_rep_iso RS (allI RS retraction_strict); |
|
176 |
val rep_strict = ax_abs_iso RS (allI RS retraction_strict); |
|
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
34974
diff
changeset
|
177 |
val iso_rews = map Drule.export_without_context [ax_abs_iso, ax_rep_iso, abs_strict, rep_strict]; |
23152 | 178 |
|
179 |
(* ----- theorems concerning the constructors, discriminators and selectors - *) |
|
180 |
||
181 |
local |
|
182 |
fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args; |
|
183 |
||
35288 | 184 |
fun pat_lhs (con,_,args) = mk_branch (list_comb (%%:(pat_name con), ps args)); |
23152 | 185 |
|
35288 | 186 |
fun pat_rhs (con,_,[]) = mk_return ((%:"rhs") ` HOLogic.unit) |
187 |
| pat_rhs (con,_,args) = |
|
26012 | 188 |
(mk_branch (mk_ctuple_pat (ps args))) |
23152 | 189 |
`(%:"rhs")`(mk_ctuple (map %# args)); |
190 |
||
191 |
fun pat_strict c = |
|
192 |
let |
|
25132 | 193 |
val axs = @{thm branch_def} :: axs_pat_def; |
23152 | 194 |
val goal = mk_trp (strict (pat_lhs c ` (%:"rhs"))); |
195 |
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
|
196 |
in pg axs goal (K tacs) end; |
23152 | 197 |
|
35288 | 198 |
fun pat_app c (con, _, args) = |
23152 | 199 |
let |
25132 | 200 |
val axs = @{thm branch_def} :: axs_pat_def; |
23152 | 201 |
val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args); |
35288 | 202 |
val rhs = if con = first c then pat_rhs c else mk_fail; |
23152 | 203 |
val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); |
204 |
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
|
205 |
in pg axs goal (K tacs) end; |
23152 | 206 |
|
29402 | 207 |
val _ = trace " Proving pat_stricts..."; |
23152 | 208 |
val pat_stricts = map pat_strict cons; |
29402 | 209 |
val _ = trace " Proving pat_apps..."; |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
210 |
val pat_apps = maps (fn c => map (pat_app c) cons) cons; |
23152 | 211 |
in |
212 |
val pat_rews = pat_stricts @ pat_apps; |
|
213 |
end; |
|
214 |
||
215 |
(* ----- theorems concerning one induction step ----------------------------- *) |
|
216 |
||
217 |
val copy_strict = |
|
218 |
let |
|
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
219 |
val _ = trace " Proving copy_strict..."; |
23152 | 220 |
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
|
221 |
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
|
222 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps rules) 1]; |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
223 |
in |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
224 |
SOME (pg [ax_copy_def] goal (K tacs)) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
225 |
handle |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
226 |
THM (s, _, _) => (trace s; NONE) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
227 |
| ERROR s => (trace s; NONE) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
228 |
end; |
23152 | 229 |
|
230 |
local |
|
35288 | 231 |
fun copy_app (con, _, args) = |
23152 | 232 |
let |
233 |
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
|
234 |
fun one_rhs arg = |
33971 | 235 |
if Datatype_Aux.is_rec_type (dtyp_of arg) |
33801 | 236 |
then Domain_Axioms.copy_of_dtyp map_tab |
237 |
(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
|
238 |
else (%# arg); |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
239 |
val rhs = con_app2 con one_rhs args; |
35059
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
240 |
fun is_rec arg = Datatype_Aux.is_rec_type (dtyp_of arg); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
241 |
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
242 |
fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); |
23152 | 243 |
val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); |
33317 | 244 |
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
|
245 |
val stricts = abs_strict :: rep_strict :: @{thms domain_map_stricts}; |
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
246 |
(* FIXME! case_UU_tac *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
247 |
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
|
248 |
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
|
249 |
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
|
250 |
in pg (ax_copy_def::con_appls) goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; |
23152 | 251 |
in |
29402 | 252 |
val _ = trace " Proving copy_apps..."; |
23152 | 253 |
val copy_apps = map copy_app cons; |
254 |
end; |
|
255 |
||
256 |
local |
|
35288 | 257 |
fun one_strict (con, _, args) = |
23152 | 258 |
let |
259 |
val goal = mk_trp (dc_copy`UU`(con_app con args) === UU); |
|
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
260 |
val rews = the_list copy_strict @ copy_apps @ con_rews; |
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
261 |
(* FIXME! case_UU_tac *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
262 |
fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @ |
23152 | 263 |
[asm_simp_tac (HOLCF_ss addsimps rews) 1]; |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
264 |
in |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
265 |
SOME (pg [] goal tacs) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
266 |
handle |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
267 |
THM (s, _, _) => (trace s; NONE) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
268 |
| ERROR s => (trace s; NONE) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
269 |
end; |
23152 | 270 |
|
35288 | 271 |
fun has_nonlazy_rec (_, _, args) = exists is_nonlazy_rec args; |
23152 | 272 |
in |
29402 | 273 |
val _ = trace " Proving copy_stricts..."; |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
274 |
val copy_stricts = map_filter one_strict (filter has_nonlazy_rec cons); |
23152 | 275 |
end; |
276 |
||
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
277 |
val copy_rews = the_list copy_strict @ copy_apps @ copy_stricts; |
23152 | 278 |
|
279 |
in |
|
280 |
thy |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
281 |
|> Sign.add_path (Long_Name.base_name dname) |
31004 | 282 |
|> snd o PureThy.add_thmss [ |
283 |
((Binding.name "iso_rews" , iso_rews ), [Simplifier.simp_add]), |
|
284 |
((Binding.name "exhaust" , [exhaust] ), []), |
|
285 |
((Binding.name "casedist" , [casedist] ), [Induct.cases_type dname]), |
|
286 |
((Binding.name "when_rews" , when_rews ), [Simplifier.simp_add]), |
|
287 |
((Binding.name "compacts" , con_compacts), [Simplifier.simp_add]), |
|
33427 | 288 |
((Binding.name "con_rews" , con_rews ), |
289 |
[Simplifier.simp_add, Fixrec.fixrec_simp_add]), |
|
31004 | 290 |
((Binding.name "sel_rews" , sel_rews ), [Simplifier.simp_add]), |
291 |
((Binding.name "dis_rews" , dis_rews ), [Simplifier.simp_add]), |
|
292 |
((Binding.name "pat_rews" , pat_rews ), [Simplifier.simp_add]), |
|
293 |
((Binding.name "dist_les" , dist_les ), [Simplifier.simp_add]), |
|
294 |
((Binding.name "dist_eqs" , dist_eqs ), [Simplifier.simp_add]), |
|
295 |
((Binding.name "inverts" , inverts ), [Simplifier.simp_add]), |
|
296 |
((Binding.name "injects" , injects ), [Simplifier.simp_add]), |
|
297 |
((Binding.name "copy_rews" , copy_rews ), [Simplifier.simp_add]), |
|
33427 | 298 |
((Binding.name "match_rews", mat_rews ), |
299 |
[Simplifier.simp_add, Fixrec.fixrec_simp_add])] |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
300 |
|> Sign.parent_path |
28536 | 301 |
|> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ |
23152 | 302 |
pat_rews @ dist_les @ dist_eqs @ copy_rews) |
303 |
end; (* let *) |
|
304 |
||
305 |
fun comp_theorems (comp_dnam, eqs: eq list) thy = |
|
306 |
let |
|
27232 | 307 |
val global_ctxt = ProofContext.init thy; |
33801 | 308 |
val map_tab = Domain_Isomorphism.get_map_tab thy; |
27232 | 309 |
|
23152 | 310 |
val dnames = map (fst o fst) eqs; |
311 |
val conss = map snd eqs; |
|
28965 | 312 |
val comp_dname = Sign.full_bname thy comp_dnam; |
23152 | 313 |
|
29402 | 314 |
val _ = message ("Proving induction properties of domain "^comp_dname^" ..."); |
23152 | 315 |
val pg = pg' thy; |
316 |
||
317 |
(* ----- getting the composite axiom and definitions ------------------------ *) |
|
318 |
||
319 |
local |
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
320 |
fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); |
23152 | 321 |
in |
322 |
val axs_reach = map (ga "reach" ) dnames; |
|
323 |
val axs_take_def = map (ga "take_def" ) dnames; |
|
324 |
val axs_finite_def = map (ga "finite_def") dnames; |
|
325 |
val ax_copy2_def = ga "copy_def" comp_dnam; |
|
35444 | 326 |
(* TEMPORARILY DISABLED |
23152 | 327 |
val ax_bisim_def = ga "bisim_def" comp_dnam; |
35444 | 328 |
TEMPORARILY DISABLED *) |
23152 | 329 |
end; |
330 |
||
331 |
local |
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26342
diff
changeset
|
332 |
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
|
333 |
fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s); |
23152 | 334 |
in |
335 |
val cases = map (gt "casedist" ) dnames; |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
336 |
val con_rews = maps (gts "con_rews" ) dnames; |
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
337 |
val copy_rews = maps (gts "copy_rews") dnames; |
23152 | 338 |
end; |
339 |
||
340 |
fun dc_take dn = %%:(dn^"_take"); |
|
341 |
val x_name = idx_name dnames "x"; |
|
342 |
val P_name = idx_name dnames "P"; |
|
343 |
val n_eqs = length eqs; |
|
344 |
||
345 |
(* ----- theorems concerning finite approximation and finite induction ------ *) |
|
346 |
||
347 |
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
|
348 |
val iterate_Cprod_ss = global_simpset_of @{theory Fix}; |
23152 | 349 |
val copy_con_rews = copy_rews @ con_rews; |
350 |
val copy_take_defs = |
|
351 |
(if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def; |
|
29402 | 352 |
val _ = trace " Proving take_stricts..."; |
35057 | 353 |
fun one_take_strict ((dn, args), _) = |
23152 | 354 |
let |
35057 | 355 |
val goal = mk_trp (strict (dc_take dn $ %:"n")); |
356 |
val rules = [ |
|
357 |
@{thm monofun_fst [THEN monofunE]}, |
|
358 |
@{thm monofun_snd [THEN monofunE]}]; |
|
359 |
val tacs = [ |
|
360 |
rtac @{thm UU_I} 1, |
|
361 |
rtac @{thm below_eq_trans} 1, |
|
362 |
resolve_tac axs_reach 2, |
|
363 |
rtac @{thm monofun_cfun_fun} 1, |
|
364 |
REPEAT (resolve_tac rules 1), |
|
365 |
rtac @{thm iterate_below_fix} 1]; |
|
366 |
in pg axs_take_def goal (K tacs) end; |
|
367 |
val take_stricts = map one_take_strict eqs; |
|
23152 | 368 |
fun take_0 n dn = |
369 |
let |
|
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
370 |
val goal = mk_trp ((dc_take dn $ @{term "0::nat"}) `% x_name n === UU); |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
371 |
in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end; |
23152 | 372 |
val take_0s = mapn take_0 1 dnames; |
29402 | 373 |
val _ = trace " Proving take_apps..."; |
35288 | 374 |
fun one_take_app dn (con, _, args) = |
23152 | 375 |
let |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
376 |
fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n"; |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
377 |
fun one_rhs arg = |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
378 |
if Datatype_Aux.is_rec_type (dtyp_of arg) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
379 |
then Domain_Axioms.copy_of_dtyp map_tab |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
380 |
mk_take (dtyp_of arg) ` (%# arg) |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
381 |
else (%# arg); |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
382 |
val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args); |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
383 |
val rhs = con_app2 con one_rhs args; |
35059
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
384 |
fun is_rec arg = Datatype_Aux.is_rec_type (dtyp_of arg); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
385 |
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
386 |
fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
387 |
val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
388 |
val tacs = [asm_simp_tac (HOLCF_ss addsimps copy_con_rews) 1]; |
acbc346e5310
correct definedness side conditions for copy_apps and take_apps
huffman
parents:
35058
diff
changeset
|
389 |
in pg copy_take_defs goal (K tacs) end; |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
390 |
fun one_take_apps ((dn, _), cons) = map (one_take_app dn) cons; |
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
391 |
val take_apps = maps one_take_apps eqs; |
23152 | 392 |
in |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
34974
diff
changeset
|
393 |
val take_rews = map Drule.export_without_context |
35058
d0cc1650b378
handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents:
35057
diff
changeset
|
394 |
(take_stricts @ take_0s @ take_apps); |
23152 | 395 |
end; (* local *) |
396 |
||
397 |
local |
|
35288 | 398 |
fun one_con p (con, _, args) = |
23152 | 399 |
let |
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
400 |
val P_names = map P_name (1 upto (length dnames)); |
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
401 |
val vns = Name.variant_list P_names (map vname args); |
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
402 |
val nonlazy_vns = map snd (filter_out (is_lazy o fst) (args ~~ vns)); |
23152 | 403 |
fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; |
404 |
val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); |
|
33317 | 405 |
val t2 = lift ind_hyp (filter is_rec args, t1); |
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
406 |
val t3 = lift_defined (bound_arg vns) (nonlazy_vns, t2); |
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
407 |
in Library.foldr mk_All (vns, t3) end; |
23152 | 408 |
|
409 |
fun one_eq ((p, cons), concl) = |
|
410 |
mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl); |
|
411 |
||
412 |
fun ind_term concf = Library.foldr one_eq |
|
413 |
(mapn (fn n => fn x => (P_name n, x)) 1 conss, |
|
414 |
mk_trp (foldr1 mk_conj (mapn concf 1 dnames))); |
|
415 |
val take_ss = HOL_ss addsimps take_rews; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
416 |
fun quant_tac ctxt i = EVERY |
27239 | 417 |
(mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames); |
23152 | 418 |
|
419 |
fun ind_prems_tac prems = EVERY |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
420 |
(maps (fn cons => |
23152 | 421 |
(resolve_tac prems 1 :: |
35288 | 422 |
maps (fn (_,_,args) => |
23152 | 423 |
resolve_tac prems 1 :: |
424 |
map (K(atac 1)) (nonlazy args) @ |
|
33317 | 425 |
map (K(atac 1)) (filter is_rec args)) |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
426 |
cons)) |
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
427 |
conss); |
23152 | 428 |
local |
429 |
(* check whether every/exists constructor of the n-th part of the equation: |
|
430 |
it has a possibly indirectly recursive argument that isn't/is possibly |
|
431 |
indirectly lazy *) |
|
432 |
fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => |
|
433 |
is_rec arg andalso not(rec_of arg mem ns) andalso |
|
434 |
((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse |
|
435 |
rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) |
|
436 |
(lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) |
|
35288 | 437 |
) o third) cons; |
23152 | 438 |
fun all_rec_to ns = rec_to forall not all_rec_to ns; |
439 |
fun warn (n,cons) = |
|
440 |
if all_rec_to [] false (n,cons) |
|
441 |
then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true) |
|
442 |
else false; |
|
443 |
fun lazy_rec_to ns = rec_to exists I lazy_rec_to ns; |
|
444 |
||
445 |
in |
|
446 |
val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; |
|
447 |
val is_emptys = map warn n__eqs; |
|
448 |
val is_finite = forall (not o lazy_rec_to [] false) n__eqs; |
|
449 |
end; |
|
450 |
in (* local *) |
|
29402 | 451 |
val _ = trace " Proving finite_ind..."; |
23152 | 452 |
val finite_ind = |
453 |
let |
|
454 |
fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n)); |
|
455 |
val goal = ind_term concf; |
|
456 |
||
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
457 |
fun tacf {prems, context} = |
23152 | 458 |
let |
459 |
val tacs1 = [ |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
460 |
quant_tac context 1, |
23152 | 461 |
simp_tac HOL_ss 1, |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
462 |
InductTacs.induct_tac context [[SOME "n"]] 1, |
23152 | 463 |
simp_tac (take_ss addsimps prems) 1, |
464 |
TRY (safe_tac HOL_cs)]; |
|
465 |
fun arg_tac arg = |
|
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
466 |
(* FIXME! case_UU_tac *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
467 |
case_UU_tac context (prems @ con_rews) 1 |
23152 | 468 |
(List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); |
35288 | 469 |
fun con_tacs (con, _, args) = |
23152 | 470 |
asm_simp_tac take_ss 1 :: |
33317 | 471 |
map arg_tac (filter is_nonlazy_rec args) @ |
23152 | 472 |
[resolve_tac prems 1] @ |
33317 | 473 |
map (K (atac 1)) (nonlazy args) @ |
474 |
map (K (etac spec 1)) (filter is_rec args); |
|
23152 | 475 |
fun cases_tacs (cons, cases) = |
27239 | 476 |
res_inst_tac context [(("x", 0), "x")] cases 1 :: |
23152 | 477 |
asm_simp_tac (take_ss addsimps prems) 1 :: |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
478 |
maps con_tacs cons; |
23152 | 479 |
in |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
480 |
tacs1 @ maps cases_tacs (conss ~~ cases) |
23152 | 481 |
end; |
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
482 |
in pg'' thy [] goal tacf |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
483 |
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
|
484 |
end; |
23152 | 485 |
|
29402 | 486 |
val _ = trace " Proving take_lemmas..."; |
23152 | 487 |
val take_lemmas = |
488 |
let |
|
489 |
fun take_lemma n (dn, ax_reach) = |
|
490 |
let |
|
491 |
val lhs = dc_take dn $ Bound 0 `%(x_name n); |
|
492 |
val rhs = dc_take dn $ Bound 0 `%(x_name n^"'"); |
|
493 |
val concl = mk_trp (%:(x_name n) === %:(x_name n^"'")); |
|
494 |
val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl; |
|
33396 | 495 |
val rules = [contlub_fst RS contlubE RS ssubst, |
496 |
contlub_snd RS contlubE RS ssubst]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
497 |
fun tacf {prems, context} = [ |
27239 | 498 |
res_inst_tac context [(("t", 0), x_name n )] (ax_reach RS subst) 1, |
499 |
res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1, |
|
23152 | 500 |
stac fix_def2 1, |
501 |
REPEAT (CHANGED |
|
33396 | 502 |
(resolve_tac rules 1 THEN chain_tac 1)), |
23152 | 503 |
stac contlub_cfun_fun 1, |
504 |
stac contlub_cfun_fun 2, |
|
505 |
rtac lub_equal 3, |
|
506 |
chain_tac 1, |
|
507 |
rtac allI 1, |
|
508 |
resolve_tac prems 1]; |
|
509 |
in pg'' thy axs_take_def goal tacf end; |
|
510 |
in mapn take_lemma 1 (dnames ~~ axs_reach) end; |
|
511 |
||
512 |
(* ----- theorems concerning finiteness and induction ----------------------- *) |
|
513 |
||
29402 | 514 |
val _ = trace " Proving finites, ind..."; |
23152 | 515 |
val (finites, ind) = |
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
516 |
( |
23152 | 517 |
if is_finite |
518 |
then (* finite case *) |
|
519 |
let |
|
520 |
fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x"); |
|
521 |
fun dname_lemma dn = |
|
522 |
let |
|
523 |
val prem1 = mk_trp (defined (%:"x")); |
|
524 |
val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU); |
|
525 |
val prem2 = mk_trp (mk_disj (disj1, take_enough dn)); |
|
526 |
val concl = mk_trp (take_enough dn); |
|
527 |
val goal = prem1 ===> prem2 ===> concl; |
|
528 |
val tacs = [ |
|
529 |
etac disjE 1, |
|
530 |
etac notE 1, |
|
531 |
resolve_tac take_lemmas 1, |
|
532 |
asm_simp_tac take_ss 1, |
|
533 |
atac 1]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
534 |
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
|
535 |
val _ = trace " Proving finite_lemmas1a"; |
23152 | 536 |
val finite_lemmas1a = map dname_lemma dnames; |
537 |
||
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
538 |
val _ = trace " Proving finite_lemma1b"; |
23152 | 539 |
val finite_lemma1b = |
540 |
let |
|
541 |
fun mk_eqn n ((dn, args), _) = |
|
542 |
let |
|
543 |
val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU; |
|
544 |
val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0; |
|
545 |
in |
|
546 |
mk_constrainall |
|
547 |
(x_name n, Type (dn,args), mk_disj (disj1, disj2)) |
|
548 |
end; |
|
549 |
val goal = |
|
550 |
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
|
551 |
fun arg_tacs ctxt vn = [ |
27239 | 552 |
eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1, |
23152 | 553 |
etac disjE 1, |
554 |
asm_simp_tac (HOL_ss addsimps con_rews) 1, |
|
555 |
asm_simp_tac take_ss 1]; |
|
35288 | 556 |
fun con_tacs ctxt (con, _, args) = |
23152 | 557 |
asm_simp_tac take_ss 1 :: |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
558 |
maps (arg_tacs ctxt) (nonlazy_rec args); |
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
559 |
fun foo_tacs ctxt n (cons, cases) = |
23152 | 560 |
simp_tac take_ss 1 :: |
561 |
rtac allI 1 :: |
|
27239 | 562 |
res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 :: |
23152 | 563 |
asm_simp_tac take_ss 1 :: |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
564 |
maps (con_tacs ctxt) cons; |
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
565 |
fun tacs ctxt = |
23152 | 566 |
rtac allI 1 :: |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
567 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: |
23152 | 568 |
simp_tac take_ss 1 :: |
569 |
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
|
570 |
flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases)); |
23152 | 571 |
in pg [] goal tacs end; |
572 |
||
573 |
fun one_finite (dn, l1b) = |
|
574 |
let |
|
575 |
val goal = mk_trp (%%:(dn^"_finite") $ %:"x"); |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
576 |
fun tacs ctxt = [ |
35443
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents:
35288
diff
changeset
|
577 |
(* FIXME! case_UU_tac *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
578 |
case_UU_tac ctxt take_rews 1 "x", |
23152 | 579 |
eresolve_tac finite_lemmas1a 1, |
580 |
step_tac HOL_cs 1, |
|
581 |
step_tac HOL_cs 1, |
|
582 |
cut_facts_tac [l1b] 1, |
|
583 |
fast_tac HOL_cs 1]; |
|
584 |
in pg axs_finite_def goal tacs end; |
|
585 |
||
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
586 |
val _ = trace " Proving finites"; |
27232 | 587 |
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
|
588 |
val _ = trace " Proving ind"; |
23152 | 589 |
val ind = |
590 |
let |
|
591 |
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
|
592 |
fun tacf {prems, context} = |
23152 | 593 |
let |
594 |
fun finite_tacs (finite, fin_ind) = [ |
|
595 |
rtac(rewrite_rule axs_finite_def finite RS exE)1, |
|
596 |
etac subst 1, |
|
597 |
rtac fin_ind 1, |
|
598 |
ind_prems_tac prems]; |
|
599 |
in |
|
600 |
TRY (safe_tac HOL_cs) :: |
|
27232 | 601 |
maps finite_tacs (finites ~~ atomize global_ctxt finite_ind) |
23152 | 602 |
end; |
603 |
in pg'' thy [] (ind_term concf) tacf end; |
|
604 |
in (finites, ind) end (* let *) |
|
605 |
||
606 |
else (* infinite case *) |
|
607 |
let |
|
608 |
fun one_finite n dn = |
|
27239 | 609 |
read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle; |
23152 | 610 |
val finites = mapn one_finite 1 dnames; |
611 |
||
612 |
val goal = |
|
613 |
let |
|
26012 | 614 |
fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); |
23152 | 615 |
fun concf n dn = %:(P_name n) $ %:(x_name n); |
616 |
in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; |
|
33396 | 617 |
val cont_rules = |
618 |
[cont_id, cont_const, cont2cont_Rep_CFun, |
|
619 |
cont2cont_fst, cont2cont_snd]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
620 |
fun tacf {prems, context} = |
23152 | 621 |
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
|
622 |
quant_tac context 1, |
23152 | 623 |
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
|
624 |
REPEAT_DETERM (rtac adm_all 1), |
23152 | 625 |
REPEAT_DETERM ( |
626 |
TRY (rtac adm_conj 1) THEN |
|
627 |
rtac adm_subst 1 THEN |
|
33396 | 628 |
REPEAT (resolve_tac cont_rules 1) THEN |
629 |
resolve_tac prems 1), |
|
23152 | 630 |
strip_tac 1, |
631 |
rtac (rewrite_rule axs_take_def finite_ind) 1, |
|
632 |
ind_prems_tac prems]; |
|
633 |
val ind = (pg'' thy [] goal tacf |
|
634 |
handle ERROR _ => |
|
33396 | 635 |
(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
|
636 |
in (finites, ind) end |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
637 |
) |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
638 |
handle THM _ => |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
639 |
(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
|
640 |
| ERROR _ => |
33810
38375b16ffd9
nicer warning message for indirect-recursive domain definitions
huffman
parents:
33801
diff
changeset
|
641 |
(warning "Cannot prove induction rule"; ([], TrueI)); |
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
642 |
|
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
643 |
|
23152 | 644 |
end; (* local *) |
645 |
||
646 |
(* ----- theorem concerning coinduction ------------------------------------- *) |
|
647 |
||
35444 | 648 |
(* COINDUCTION TEMPORARILY DISABLED |
23152 | 649 |
local |
650 |
val xs = mapn (fn n => K (x_name n)) 1 dnames; |
|
651 |
fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); |
|
652 |
val take_ss = HOL_ss addsimps take_rews; |
|
653 |
val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")")); |
|
29402 | 654 |
val _ = trace " Proving coind_lemma..."; |
23152 | 655 |
val coind_lemma = |
656 |
let |
|
657 |
fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1; |
|
658 |
fun mk_eqn n dn = |
|
659 |
(dc_take dn $ %:"n" ` bnd_arg n 0) === |
|
660 |
(dc_take dn $ %:"n" ` bnd_arg n 1); |
|
661 |
fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t)); |
|
662 |
val goal = |
|
663 |
mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R", |
|
664 |
Library.foldr mk_all2 (xs, |
|
665 |
Library.foldr mk_imp (mapn mk_prj 0 dnames, |
|
666 |
foldr1 mk_conj (mapn mk_eqn 0 dnames))))); |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
667 |
fun x_tacs ctxt n x = [ |
23152 | 668 |
rotate_tac (n+1) 1, |
669 |
etac all2E 1, |
|
27239 | 670 |
eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1, |
23152 | 671 |
TRY (safe_tac HOL_cs), |
672 |
REPEAT (CHANGED (asm_simp_tac take_ss 1))]; |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
673 |
fun tacs ctxt = [ |
23152 | 674 |
rtac impI 1, |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
675 |
InductTacs.induct_tac ctxt [[SOME "n"]] 1, |
23152 | 676 |
simp_tac take_ss 1, |
677 |
safe_tac HOL_cs] @ |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
678 |
flat (mapn (x_tacs ctxt) 0 xs); |
23152 | 679 |
in pg [ax_bisim_def] goal tacs end; |
680 |
in |
|
29402 | 681 |
val _ = trace " Proving coind..."; |
23152 | 682 |
val coind = |
683 |
let |
|
684 |
fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'")); |
|
685 |
fun mk_eqn x = %:x === %:(x^"'"); |
|
686 |
val goal = |
|
687 |
mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===> |
|
688 |
Logic.list_implies (mapn mk_prj 0 xs, |
|
689 |
mk_trp (foldr1 mk_conj (map mk_eqn xs))); |
|
690 |
val tacs = |
|
691 |
TRY (safe_tac HOL_cs) :: |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
692 |
maps (fn take_lemma => [ |
23152 | 693 |
rtac take_lemma 1, |
694 |
cut_facts_tac [coind_lemma] 1, |
|
695 |
fast_tac HOL_cs 1]) |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26012
diff
changeset
|
696 |
take_lemmas; |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27153
diff
changeset
|
697 |
in pg [] goal (K tacs) end; |
23152 | 698 |
end; (* local *) |
35444 | 699 |
COINDUCTION TEMPORARILY DISABLED *) |
23152 | 700 |
|
32172 | 701 |
val inducts = Project_Rule.projections (ProofContext.init thy) ind; |
30829 | 702 |
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
|
703 |
val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI); |
30829 | 704 |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24503
diff
changeset
|
705 |
in thy |> Sign.add_path comp_dnam |
31004 | 706 |
|> snd o PureThy.add_thmss [ |
707 |
((Binding.name "take_rews" , take_rews ), [Simplifier.simp_add]), |
|
708 |
((Binding.name "take_lemmas", take_lemmas ), []), |
|
709 |
((Binding.name "finites" , finites ), []), |
|
710 |
((Binding.name "finite_ind" , [finite_ind]), []), |
|
35444 | 711 |
((Binding.name "ind" , [ind] ), [])(*, |
712 |
((Binding.name "coind" , [coind] ), [])*)] |
|
31232
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
713 |
|> (if induct_failed then I |
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents:
31160
diff
changeset
|
714 |
else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))) |
28536 | 715 |
|> Sign.parent_path |> pair take_rews |
23152 | 716 |
end; (* let *) |
717 |
end; (* struct *) |