| author | haftmann | 
| Thu, 11 Mar 2010 14:38:09 +0100 | |
| changeset 35721 | f7bbee848403 | 
| parent 35663 | ada7bc39c6b1 | 
| child 35774 | 218e9766a848 | 
| 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 | |
| 35558 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 15 | -> Domain_Take_Proofs.iso_info | 
| 35444 | 16 | -> theory -> thm list * theory; | 
| 17 | ||
| 35657 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 18 | val comp_theorems : | 
| 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 19 | bstring * Domain_Library.eq list -> | 
| 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 20 | Domain_Take_Proofs.take_induct_info -> | 
| 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 21 | theory -> thm list * theory | 
| 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 22 | |
| 32740 | 23 | val quiet_mode: bool Unsynchronized.ref; | 
| 24 | val trace_domain: bool Unsynchronized.ref; | |
| 31005 | 25 | end; | 
| 26 | ||
| 31023 | 27 | structure Domain_Theorems :> DOMAIN_THEOREMS = | 
| 31005 | 28 | struct | 
| 23152 | 29 | |
| 32740 | 30 | val quiet_mode = Unsynchronized.ref false; | 
| 31 | val trace_domain = Unsynchronized.ref false; | |
| 29402 | 32 | |
| 33 | fun message s = if !quiet_mode then () else writeln s; | |
| 34 | fun trace s = if !trace_domain then tracing s else (); | |
| 35 | ||
| 23152 | 36 | open Domain_Library; | 
| 37 | infixr 0 ===>; | |
| 38 | infixr 0 ==>; | |
| 39 | infix 0 == ; | |
| 40 | infix 1 ===; | |
| 41 | infix 1 ~= ; | |
| 42 | infix 1 <<; | |
| 43 | infix 1 ~<<; | |
| 44 | infix 9 ` ; | |
| 45 | infix 9 `% ; | |
| 46 | infix 9 `%%; | |
| 47 | infixr 9 oo; | |
| 48 | ||
| 49 | (* ----- general proof facilities ------------------------------------------- *) | |
| 50 | ||
| 24503 | 51 | fun legacy_infer_term thy t = | 
| 52 | let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy) | |
| 53 | in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end; | |
| 54 | ||
| 23152 | 55 | fun pg'' thy defs t tacs = | 
| 56 | let | |
| 24503 | 57 | val t' = legacy_infer_term thy t; | 
| 23152 | 58 | val asms = Logic.strip_imp_prems t'; | 
| 59 | val prop = Logic.strip_imp_concl t'; | |
| 26711 | 60 |     fun tac {prems, context} =
 | 
| 23152 | 61 | rewrite_goals_tac defs THEN | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 62 |       EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
 | 
| 23152 | 63 | in Goal.prove_global thy [] asms prop tac end; | 
| 64 | ||
| 65 | fun pg' thy defs t tacsf = | |
| 66 | let | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 67 |     fun tacs {prems, context} =
 | 
| 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 68 | if null prems then tacsf context | 
| 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 69 | else cut_facts_tac prems 1 :: tacsf context; | 
| 23152 | 70 | in pg'' thy defs t tacs end; | 
| 71 | ||
| 35443 
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
 huffman parents: 
35288diff
changeset | 72 | (* FIXME!!!!!!!!! *) | 
| 
2e0f9516947e
change domain package's treatment of variable names in theorems to be like datatype package
 huffman parents: 
35288diff
changeset | 73 | (* 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: 
35288diff
changeset | 74 | (* The names can conflict with existing constants or other syntax! *) | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 75 | fun case_UU_tac ctxt rews i v = | 
| 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27153diff
changeset | 76 | InductTacs.case_tac ctxt (v^"=UU") i THEN | 
| 23152 | 77 | asm_simp_tac (HOLCF_ss addsimps rews) i; | 
| 78 | ||
| 79 | (* ----- general proofs ----------------------------------------------------- *) | |
| 80 | ||
| 29064 | 81 | val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp}
 | 
| 23152 | 82 | |
| 35444 | 83 | fun theorems | 
| 84 | (((dname, _), cons) : eq, eqs : eq list) | |
| 85 | (dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list) | |
| 35558 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 86 | (iso_info : Domain_Take_Proofs.iso_info) | 
| 35444 | 87 | (thy : theory) = | 
| 23152 | 88 | let | 
| 89 | ||
| 29402 | 90 | val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
 | 
| 35514 
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
 huffman parents: 
35512diff
changeset | 91 | val map_tab = Domain_Take_Proofs.get_map_tab thy; | 
| 33801 | 92 | |
| 23152 | 93 | |
| 94 | (* ----- getting the axioms and definitions --------------------------------- *) | |
| 95 | ||
| 35558 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 96 | val ax_abs_iso = #abs_inverse iso_info; | 
| 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 97 | val ax_rep_iso = #rep_inverse iso_info; | 
| 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 98 | |
| 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 99 | val abs_const = #abs_const iso_info; | 
| 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 100 | val rep_const = #rep_const iso_info; | 
| 
bb088a6fafbc
add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
 huffman parents: 
35557diff
changeset | 101 | |
| 23152 | 102 | local | 
| 26343 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26342diff
changeset | 103 | fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); | 
| 23152 | 104 | in | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 105 | val ax_take_0 = ga "take_0" dname; | 
| 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 106 | val ax_take_Suc = ga "take_Suc" dname; | 
| 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 107 | val ax_take_strict = ga "take_strict" dname; | 
| 23152 | 108 | end; (* local *) | 
| 109 | ||
| 35444 | 110 | (* ----- define constructors ------------------------------------------------ *) | 
| 111 | ||
| 112 | val (result, thy) = | |
| 113 | Domain_Constructors.add_domain_constructors | |
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35482diff
changeset | 114 | (Long_Name.base_name dname) (snd dom_eqn) iso_info thy; | 
| 35444 | 115 | |
| 35451 | 116 | val con_appls = #con_betas result; | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 117 | val {exhaust, casedist, ...} = result;
 | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 118 | val {con_compacts, con_rews, inverts, injects, dist_les, dist_eqs, ...} = result;
 | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 119 | val {sel_rews, ...} = result;
 | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 120 | val when_rews = #cases result; | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 121 | 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: 
35460diff
changeset | 122 | val dis_rews = #dis_rews result; | 
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35464diff
changeset | 123 | val mat_rews = #match_rews result; | 
| 35482 
d756837b708d
move proofs of pat_rews to domain_constructors.ML
 huffman parents: 
35481diff
changeset | 124 | val pat_rews = #pat_rews result; | 
| 35444 | 125 | |
| 23152 | 126 | (* ----- theorems concerning the isomorphism -------------------------------- *) | 
| 127 | ||
| 35444 | 128 | val pg = pg' thy; | 
| 129 | ||
| 35560 | 130 | val retraction_strict = @{thm retraction_strict};
 | 
| 23152 | 131 | val abs_strict = ax_rep_iso RS (allI RS retraction_strict); | 
| 132 | 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: 
34974diff
changeset | 133 | val iso_rews = map Drule.export_without_context [ax_abs_iso, ax_rep_iso, abs_strict, rep_strict]; | 
| 23152 | 134 | |
| 135 | (* ----- theorems concerning one induction step ----------------------------- *) | |
| 136 | ||
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 137 | local | 
| 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 138 | fun dc_take dn = %%:(dn^"_take"); | 
| 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 139 | val dnames = map (fst o fst) eqs; | 
| 35523 
cc57f4a274a3
fix proof script for take_apps so it works with indirect recursion
 huffman parents: 
35521diff
changeset | 140 | val deflation_thms = Domain_Take_Proofs.get_deflation_thms thy; | 
| 
cc57f4a274a3
fix proof script for take_apps so it works with indirect recursion
 huffman parents: 
35521diff
changeset | 141 | fun get_deflation_take dn = PureThy.get_thm thy (dn ^ ".deflation_take"); | 
| 
cc57f4a274a3
fix proof script for take_apps so it works with indirect recursion
 huffman parents: 
35521diff
changeset | 142 | val axs_deflation_take = map get_deflation_take dnames; | 
| 23152 | 143 | |
| 35559 | 144 | fun copy_of_dtyp tab r dt = | 
| 145 | if Datatype_Aux.is_rec_type dt then copy tab r dt else ID | |
| 146 | and copy tab r (Datatype_Aux.DtRec i) = r i | |
| 147 | | copy tab r (Datatype_Aux.DtTFree a) = ID | |
| 148 | | copy tab r (Datatype_Aux.DtType (c, ds)) = | |
| 149 | case Symtab.lookup tab c of | |
| 150 | SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds) | |
| 151 |       | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID);
 | |
| 152 | ||
| 35521 | 153 | fun one_take_app (con, args) = | 
| 23152 | 154 | let | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 155 | 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: 
31160diff
changeset | 156 | fun one_rhs arg = | 
| 33971 | 157 | if Datatype_Aux.is_rec_type (dtyp_of arg) | 
| 35559 | 158 | then copy_of_dtyp map_tab | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 159 | mk_take (dtyp_of arg) ` (%# arg) | 
| 31232 
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
 huffman parents: 
31160diff
changeset | 160 | else (%# arg); | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 161 | val lhs = (dc_take dname $ (%%:"Suc" $ %:"n"))`(con_app con args); | 
| 31232 
689aa7da48cc
define copy functions using combinators; add checking for failed proofs of induction rules
 huffman parents: 
31160diff
changeset | 162 | val rhs = con_app2 con one_rhs args; | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 163 | val goal = mk_trp (lhs === rhs); | 
| 35590 | 164 | val rules = | 
| 165 |           [ax_take_Suc, ax_abs_iso, @{thm cfcomp2}]
 | |
| 166 |           @ @{thms take_con_rules ID1 deflation_strict}
 | |
| 35523 
cc57f4a274a3
fix proof script for take_apps so it works with indirect recursion
 huffman parents: 
35521diff
changeset | 167 | @ deflation_thms @ axs_deflation_take; | 
| 35590 | 168 | val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]; | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 169 | in pg con_appls goal (K tacs) end; | 
| 35557 
5da670d57118
uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
 huffman parents: 
35528diff
changeset | 170 | val take_apps = map one_take_app cons; | 
| 23152 | 171 | in | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 172 | val take_rews = ax_take_0 :: ax_take_strict :: take_apps; | 
| 23152 | 173 | end; | 
| 174 | ||
| 35630 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 175 | val case_ns = | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 176 | "bottom" :: map (fn (b,_,_) => Binding.name_of b) (snd dom_eqn); | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 177 | |
| 23152 | 178 | in | 
| 179 | thy | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 180 | |> Sign.add_path (Long_Name.base_name dname) | 
| 31004 | 181 | |> snd o PureThy.add_thmss [ | 
| 182 | ((Binding.name "iso_rews" , iso_rews ), [Simplifier.simp_add]), | |
| 183 | ((Binding.name "exhaust" , [exhaust] ), []), | |
| 35630 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 184 | ((Binding.name "casedist" , [casedist] ), | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 185 | [Rule_Cases.case_names case_ns, Induct.cases_type dname]), | 
| 31004 | 186 | ((Binding.name "when_rews" , when_rews ), [Simplifier.simp_add]), | 
| 187 | ((Binding.name "compacts" , con_compacts), [Simplifier.simp_add]), | |
| 33427 | 188 | ((Binding.name "con_rews" , con_rews ), | 
| 189 | [Simplifier.simp_add, Fixrec.fixrec_simp_add]), | |
| 31004 | 190 | ((Binding.name "sel_rews" , sel_rews ), [Simplifier.simp_add]), | 
| 191 | ((Binding.name "dis_rews" , dis_rews ), [Simplifier.simp_add]), | |
| 192 | ((Binding.name "pat_rews" , pat_rews ), [Simplifier.simp_add]), | |
| 193 | ((Binding.name "dist_les" , dist_les ), [Simplifier.simp_add]), | |
| 194 | ((Binding.name "dist_eqs" , dist_eqs ), [Simplifier.simp_add]), | |
| 195 | ((Binding.name "inverts" , inverts ), [Simplifier.simp_add]), | |
| 196 | ((Binding.name "injects" , injects ), [Simplifier.simp_add]), | |
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 197 | ((Binding.name "take_rews" , take_rews ), [Simplifier.simp_add]), | 
| 33427 | 198 | ((Binding.name "match_rews", mat_rews ), | 
| 199 | [Simplifier.simp_add, Fixrec.fixrec_simp_add])] | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24503diff
changeset | 200 | |> Sign.parent_path | 
| 28536 | 201 | |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35486diff
changeset | 202 | pat_rews @ dist_les @ dist_eqs) | 
| 23152 | 203 | end; (* let *) | 
| 204 | ||
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 205 | (******************************************************************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 206 | (****************************** induction rules *******************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 207 | (******************************************************************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 208 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 209 | fun prove_induction | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 210 | (comp_dnam, eqs : eq list) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 211 | (take_rews : thm list) | 
| 35659 | 212 | (take_info : Domain_Take_Proofs.take_induct_info) | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 213 | (thy : theory) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 214 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 215 | val dnames = map (fst o fst) eqs; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 216 | val conss = map snd eqs; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 217 | fun dc_take dn = %%:(dn^"_take"); | 
| 35662 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 218 | val x_name = idx_name dnames "x"; | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 219 | val P_name = idx_name dnames "P"; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 220 | val pg = pg' thy; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 221 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 222 | local | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 223 | fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 224 | fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 225 | in | 
| 35597 
e4331b99b03f
introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
 huffman parents: 
35590diff
changeset | 226 | val axs_rep_iso = map (ga "rep_iso") dnames; | 
| 
e4331b99b03f
introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
 huffman parents: 
35590diff
changeset | 227 | val axs_abs_iso = map (ga "abs_iso") dnames; | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 228 | val cases = map (ga "casedist" ) dnames; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 229 | val con_rews = maps (gts "con_rews" ) dnames; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 230 | end; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 231 | |
| 35662 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 232 |   val {take_consts, ...} = take_info;
 | 
| 35659 | 233 |   val {take_0_thms, take_Suc_thms, chain_take_thms, ...} = take_info;
 | 
| 35660 | 234 |   val {lub_take_thms, finite_defs, reach_thms, ...} = take_info;
 | 
| 35661 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 235 |   val {take_induct_thms, ...} = take_info;
 | 
| 35658 | 236 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 237 | fun one_con p (con, args) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 238 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 239 | val P_names = map P_name (1 upto (length dnames)); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 240 | val vns = Name.variant_list P_names (map vname args); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 241 | val nonlazy_vns = map snd (filter_out (is_lazy o fst) (args ~~ vns)); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 242 | fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 243 | val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 244 | val t2 = lift ind_hyp (filter is_rec args, t1); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 245 | val t3 = lift_defined (bound_arg vns) (nonlazy_vns, t2); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 246 | in Library.foldr mk_All (vns, t3) end; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 247 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 248 | fun one_eq ((p, cons), concl) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 249 | mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 250 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 251 | fun ind_term concf = Library.foldr one_eq | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 252 | (mapn (fn n => fn x => (P_name n, x)) 1 conss, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 253 | mk_trp (foldr1 mk_conj (mapn concf 1 dnames))); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 254 |   val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 255 | fun quant_tac ctxt i = EVERY | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 256 |     (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames);
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 257 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 258 | fun ind_prems_tac prems = EVERY | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 259 | (maps (fn cons => | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 260 | (resolve_tac prems 1 :: | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 261 | maps (fn (_,args) => | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 262 | resolve_tac prems 1 :: | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 263 | map (K(atac 1)) (nonlazy args) @ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 264 | map (K(atac 1)) (filter is_rec args)) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 265 | cons)) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 266 | conss); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 267 | local | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 268 | (* check whether every/exists constructor of the n-th part of the equation: | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 269 | it has a possibly indirectly recursive argument that isn't/is possibly | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 270 | indirectly lazy *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 271 | fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 272 | is_rec arg andalso not(rec_of arg mem ns) andalso | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 273 | ((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 274 | rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 275 | (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 276 | ) o snd) cons; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 277 | fun all_rec_to ns = rec_to forall not all_rec_to ns; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 278 | fun warn (n,cons) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 279 | if all_rec_to [] false (n,cons) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 280 |       then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 281 | else false; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 282 | fun lazy_rec_to ns = rec_to exists I lazy_rec_to ns; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 283 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 284 | in | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 285 | val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 286 | val is_emptys = map warn n__eqs; | 
| 35661 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 287 | val is_finite = #is_finite take_info; | 
| 35601 
50ba5010b876
print message when finiteness of domain definition is detected
 huffman parents: 
35599diff
changeset | 288 | val _ = if is_finite | 
| 
50ba5010b876
print message when finiteness of domain definition is detected
 huffman parents: 
35599diff
changeset | 289 |             then message ("Proving finiteness rule for domain "^comp_dnam^" ...")
 | 
| 
50ba5010b876
print message when finiteness of domain definition is detected
 huffman parents: 
35599diff
changeset | 290 | else (); | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 291 | end; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 292 | val _ = trace " Proving finite_ind..."; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 293 | val finite_ind = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 294 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 295 | fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n)); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 296 | val goal = ind_term concf; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 297 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 298 |       fun tacf {prems, context} =
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 299 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 300 | val tacs1 = [ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 301 | quant_tac context 1, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 302 | simp_tac HOL_ss 1, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 303 | InductTacs.induct_tac context [[SOME "n"]] 1, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 304 | simp_tac (take_ss addsimps prems) 1, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 305 | TRY (safe_tac HOL_cs)]; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 306 | fun arg_tac arg = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 307 | (* FIXME! case_UU_tac *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 308 | case_UU_tac context (prems @ con_rews) 1 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 309 | (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 310 | fun con_tacs (con, args) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 311 | asm_simp_tac take_ss 1 :: | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 312 | map arg_tac (filter is_nonlazy_rec args) @ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 313 | [resolve_tac prems 1] @ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 314 | map (K (atac 1)) (nonlazy args) @ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 315 | map (K (etac spec 1)) (filter is_rec args); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 316 | fun cases_tacs (cons, cases) = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 317 |             res_inst_tac context [(("y", 0), "x")] cases 1 ::
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 318 | asm_simp_tac (take_ss addsimps prems) 1 :: | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 319 | maps con_tacs cons; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 320 | in | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 321 | tacs1 @ maps cases_tacs (conss ~~ cases) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 322 | end; | 
| 35663 | 323 | in pg'' thy [] goal tacf end; | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 324 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 325 | (* ----- theorems concerning finiteness and induction ----------------------- *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 326 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 327 | val global_ctxt = ProofContext.init thy; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 328 | |
| 35661 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 329 | val _ = trace " Proving ind..."; | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 330 | val ind = | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 331 | if is_finite | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 332 | then (* finite case *) | 
| 35597 
e4331b99b03f
introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
 huffman parents: 
35590diff
changeset | 333 | let | 
| 35661 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 334 | fun concf n dn = %:(P_name n) $ %:(x_name n); | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 335 |         fun tacf {prems, context} =
 | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 336 | let | 
| 35661 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 337 | fun finite_tacs (take_induct, fin_ind) = [ | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 338 | rtac take_induct 1, | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 339 | rtac fin_ind 1, | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 340 | ind_prems_tac prems]; | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 341 | in | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 342 | TRY (safe_tac HOL_cs) :: | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 343 | maps finite_tacs (take_induct_thms ~~ atomize global_ctxt finite_ind) | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 344 | end; | 
| 
ede27eb8e94b
don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
 huffman parents: 
35660diff
changeset | 345 | in pg'' thy [] (ind_term concf) tacf end | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 346 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 347 | else (* infinite case *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 348 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 349 | val goal = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 350 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 351 | fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 352 | fun concf n dn = %:(P_name n) $ %:(x_name n); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 353 | in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 354 | val cont_rules = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 355 |             @{thms cont_id cont_const cont2cont_Rep_CFun
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 356 | cont2cont_fst cont2cont_snd}; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 357 | val subgoal = | 
| 35662 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 358 | let | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 359 | val Ts = map (Type o fst) eqs; | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 360 | val P_names = Datatype_Prop.indexify_names (map (K "P") dnames); | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 361 | val x_names = Datatype_Prop.indexify_names (map (K "x") dnames); | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 362 | val P_types = map (fn T => T --> HOLogic.boolT) Ts; | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 363 | val Ps = map Free (P_names ~~ P_types); | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 364 | val xs = map Free (x_names ~~ Ts); | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 365 |             val n = Free ("n", HOLogic.natT);
 | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 366 | val goals = | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 367 | map (fn ((P,t),x) => P $ HOLCF_Library.mk_capply (t $ n, x)) | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 368 | (Ps ~~ take_consts ~~ xs); | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 369 | in | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 370 | HOLogic.mk_Trueprop | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 371 |             (HOLogic.mk_all ("n", HOLogic.natT, foldr1 HOLogic.mk_conj goals))
 | 
| 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 372 | end; | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 373 |         fun tacf {prems, context} =
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 374 | let | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 375 | val subtac = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 376 | EVERY [rtac allI 1, rtac finite_ind 1, ind_prems_tac prems]; | 
| 35662 
44d7aafdddb9
construct fully typed goal in proof of induction rule
 huffman parents: 
35661diff
changeset | 377 | val subthm = Goal.prove context [] [] subgoal (K subtac); | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 378 | in | 
| 35660 | 379 | map (fn ax_reach => rtac (ax_reach RS subst) 1) reach_thms @ [ | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 380 | cut_facts_tac (subthm :: take (length dnames) prems) 1, | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 381 |             REPEAT (rtac @{thm conjI} 1 ORELSE
 | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 382 |                     EVERY [etac @{thm admD [OF _ ch2ch_Rep_CFunL]} 1,
 | 
| 35659 | 383 | resolve_tac chain_take_thms 1, | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 384 | asm_simp_tac HOL_basic_ss 1]) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 385 | ] | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 386 | end; | 
| 35663 | 387 | in pg'' thy [] goal tacf end; | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 388 | |
| 35630 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 389 | val case_ns = | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 390 | let | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 391 | val bottoms = | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 392 | if length dnames = 1 then ["bottom"] else | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 393 | map (fn s => "bottom_" ^ Long_Name.base_name s) dnames; | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 394 | fun one_eq bot (_,cons) = | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 395 | bot :: map (fn (c,_) => Long_Name.base_name c) cons; | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 396 | in flat (map2 one_eq bottoms eqs) end; | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 397 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 398 | val inducts = Project_Rule.projections (ProofContext.init thy) ind; | 
| 35630 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 399 | fun ind_rule (dname, rule) = | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 400 | ((Binding.empty, [rule]), | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 401 | [Rule_Cases.case_names case_ns, Induct.induct_type dname]); | 
| 
8e562d56d404
add case_names attribute to casedist and ind rules
 huffman parents: 
35601diff
changeset | 402 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 403 | in thy |> Sign.add_path comp_dnam | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 404 | |> snd o PureThy.add_thmss [ | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 405 | ((Binding.name "finite_ind" , [finite_ind]), []), | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 406 | ((Binding.name "ind" , [ind] ), [])] | 
| 35663 | 407 | |> (snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))) | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 408 | |> Sign.parent_path | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 409 | end; (* prove_induction *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 410 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 411 | (******************************************************************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 412 | (************************ bisimulation and coinduction ************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 413 | (******************************************************************************) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 414 | |
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 415 | fun prove_coinduction | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 416 | (comp_dnam, eqs : eq list) | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 417 | (take_lemmas : thm list) | 
| 35599 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 418 | (thy : theory) : theory = | 
| 23152 | 419 | let | 
| 27232 | 420 | |
| 23152 | 421 | val dnames = map (fst o fst) eqs; | 
| 28965 | 422 | val comp_dname = Sign.full_bname thy comp_dnam; | 
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 423 | fun dc_take dn = %%:(dn^"_take"); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 424 | val x_name = idx_name dnames "x"; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 425 | val n_eqs = length eqs; | 
| 23152 | 426 | |
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 427 | val take_rews = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 428 | maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames; | 
| 35497 | 429 | |
| 430 | (* ----- define bisimulation predicate -------------------------------------- *) | |
| 431 | ||
| 432 | local | |
| 433 | open HOLCF_Library | |
| 434 | val dtypes = map (Type o fst) eqs; | |
| 435 | val relprod = mk_tupleT (map (fn tp => tp --> tp --> boolT) dtypes); | |
| 436 | val bisim_bind = Binding.name (comp_dnam ^ "_bisim"); | |
| 437 | val bisim_type = relprod --> boolT; | |
| 438 | in | |
| 439 | val (bisim_const, thy) = | |
| 440 | Sign.declare_const ((bisim_bind, bisim_type), NoSyn) thy; | |
| 441 | end; | |
| 442 | ||
| 443 | local | |
| 444 | ||
| 445 | fun legacy_infer_term thy t = | |
| 446 | singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t); | |
| 447 | fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t); | |
| 448 | fun infer_props thy = map (apsnd (legacy_infer_prop thy)); | |
| 449 | fun add_defs_i x = PureThy.add_defs false (map Thm.no_attributes x); | |
| 450 | fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy; | |
| 451 | ||
| 452 | val comp_dname = Sign.full_bname thy comp_dnam; | |
| 453 | val dnames = map (fst o fst) eqs; | |
| 454 | val x_name = idx_name dnames "x"; | |
| 455 | ||
| 35521 | 456 | fun one_con (con, args) = | 
| 35497 | 457 | let | 
| 458 | val nonrec_args = filter_out is_rec args; | |
| 459 | val rec_args = filter is_rec args; | |
| 460 | val recs_cnt = length rec_args; | |
| 461 | val allargs = nonrec_args @ rec_args | |
| 462 | @ map (upd_vname (fn s=> s^"'")) rec_args; | |
| 463 | val allvns = map vname allargs; | |
| 464 | fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; | |
| 465 | val vns1 = map (vname_arg "" ) args; | |
| 466 | val vns2 = map (vname_arg "'") args; | |
| 467 | val allargs_cnt = length nonrec_args + 2*recs_cnt; | |
| 468 | val rec_idxs = (recs_cnt-1) downto 0; | |
| 469 | val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) | |
| 470 | (allargs~~((allargs_cnt-1) downto 0))); | |
| 471 | fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ | |
| 472 | Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); | |
| 473 | val capps = | |
| 474 | List.foldr | |
| 475 | mk_conj | |
| 476 | (mk_conj( | |
| 477 | Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1), | |
| 478 | Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2))) | |
| 479 | (mapn rel_app 1 rec_args); | |
| 480 | in | |
| 481 | List.foldr | |
| 482 | mk_ex | |
| 483 | (Library.foldr mk_conj | |
| 484 | (map (defined o Bound) nonlazy_idxs,capps)) allvns | |
| 485 | end; | |
| 486 | fun one_comp n (_,cons) = | |
| 487 | mk_all (x_name(n+1), | |
| 488 | mk_all (x_name(n+1)^"'", | |
| 489 | mk_imp (proj (Bound 2) eqs n $ Bound 1 $ Bound 0, | |
| 490 | foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU) | |
| 491 | ::map one_con cons)))); | |
| 492 | val bisim_eqn = | |
| 493 | %%:(comp_dname^"_bisim") == | |
| 494 |          mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs));
 | |
| 495 | ||
| 496 | in | |
| 497 | val ([ax_bisim_def], thy) = | |
| 498 | thy | |
| 499 | |> Sign.add_path comp_dnam | |
| 500 | |> add_defs_infer [(Binding.name "bisim_def", bisim_eqn)] | |
| 501 | ||> Sign.parent_path; | |
| 502 | end; (* local *) | |
| 503 | ||
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 504 | (* ----- theorem concerning coinduction ------------------------------------- *) | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 505 | |
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 506 | local | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 507 | val pg = pg' thy; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 508 | val xs = mapn (fn n => K (x_name n)) 1 dnames; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 509 | fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 510 |   val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
 | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 511 |   val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
 | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 512 | val _ = trace " Proving coind_lemma..."; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 513 | val coind_lemma = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 514 | let | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 515 | fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 516 | fun mk_eqn n dn = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 517 | (dc_take dn $ %:"n" ` bnd_arg n 0) === | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 518 | (dc_take dn $ %:"n" ` bnd_arg n 1); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 519 | fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t)); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 520 | val goal = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 521 | mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R", | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 522 | Library.foldr mk_all2 (xs, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 523 | Library.foldr mk_imp (mapn mk_prj 0 dnames, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 524 | foldr1 mk_conj (mapn mk_eqn 0 dnames))))); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 525 | fun x_tacs ctxt n x = [ | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 526 | rotate_tac (n+1) 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 527 | etac all2E 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 528 |         eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
 | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 529 | TRY (safe_tac HOL_cs), | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 530 | REPEAT (CHANGED (asm_simp_tac take_ss 1))]; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 531 | fun tacs ctxt = [ | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 532 | rtac impI 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 533 | InductTacs.induct_tac ctxt [[SOME "n"]] 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 534 | simp_tac take_ss 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 535 | safe_tac HOL_cs] @ | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 536 | flat (mapn (x_tacs ctxt) 0 xs); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 537 | in pg [ax_bisim_def] goal tacs end; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 538 | in | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 539 | val _ = trace " Proving coind..."; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 540 | val coind = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 541 | let | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 542 | fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'")); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 543 | fun mk_eqn x = %:x === %:(x^"'"); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 544 | val goal = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 545 | mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===> | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 546 | Logic.list_implies (mapn mk_prj 0 xs, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 547 | mk_trp (foldr1 mk_conj (map mk_eqn xs))); | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 548 | val tacs = | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 549 | TRY (safe_tac HOL_cs) :: | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 550 | maps (fn take_lemma => [ | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 551 | rtac take_lemma 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 552 | cut_facts_tac [coind_lemma] 1, | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 553 | fast_tac HOL_cs 1]) | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 554 | take_lemmas; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 555 | in pg [] goal (K tacs) end; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 556 | end; (* local *) | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 557 | |
| 35599 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 558 | in thy |> Sign.add_path comp_dnam | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 559 | |> snd o PureThy.add_thmss [((Binding.name "coind", [coind]), [])] | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 560 | |> Sign.parent_path | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 561 | end; (* let *) | 
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 562 | |
| 35657 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 563 | fun comp_theorems | 
| 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 564 | (comp_dnam : string, eqs : eq list) | 
| 35659 | 565 | (take_info : Domain_Take_Proofs.take_induct_info) | 
| 35657 
0537c34c6067
pass take_induct_info as an argument to comp_theorems
 huffman parents: 
35654diff
changeset | 566 | (thy : theory) = | 
| 35574 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 567 | let | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 568 | val map_tab = Domain_Take_Proofs.get_map_tab thy; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 569 | |
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 570 | val dnames = map (fst o fst) eqs; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 571 | val comp_dname = Sign.full_bname thy comp_dnam; | 
| 
ee5df989b7c4
move coinduction-related stuff into function prove_coindunction
 huffman parents: 
35560diff
changeset | 572 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 573 | (* ----- getting the composite axiom and definitions ------------------------ *) | 
| 23152 | 574 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 575 | (* Test for indirect recursion *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 576 | local | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 577 | fun indirect_arg arg = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 578 | rec_of arg = ~1 andalso Datatype_Aux.is_rec_type (dtyp_of arg); | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 579 | fun indirect_con (_, args) = exists indirect_arg args; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 580 | fun indirect_eq (_, cons) = exists indirect_con cons; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 581 | in | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 582 | val is_indirect = exists indirect_eq eqs; | 
| 35599 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 583 | val _ = | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 584 | if is_indirect | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 585 | then message "Indirect recursion detected, skipping proofs of (co)induction rules" | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 586 |       else message ("Proving induction properties of domain "^comp_dname^" ...");
 | 
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 587 | end; | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 588 | |
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 589 | (* theorems about take *) | 
| 23152 | 590 | |
| 35659 | 591 | val take_lemmas = #take_lemma_thms take_info; | 
| 23152 | 592 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 593 | val take_rews = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 594 | maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames; | 
| 23152 | 595 | |
| 35585 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 596 | (* prove induction rules, unless definition is indirect recursive *) | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 597 | val thy = | 
| 
555f26f00e47
skip proof of induction rule for indirect-recursive domain definitions
 huffman parents: 
35574diff
changeset | 598 | if is_indirect then thy else | 
| 35660 | 599 | prove_induction (comp_dnam, eqs) take_rews take_info thy; | 
| 23152 | 600 | |
| 35599 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 601 | val thy = | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 602 | if is_indirect then thy else | 
| 
20670f5564e9
skip coinduction proofs for indirect-recursive domain definitions
 huffman parents: 
35597diff
changeset | 603 | prove_coinduction (comp_dnam, eqs) take_lemmas thy; | 
| 23152 | 604 | |
| 35642 
f478d5a9d238
generate separate qualified theorem name for each type's reach and take_lemma
 huffman parents: 
35630diff
changeset | 605 | in | 
| 
f478d5a9d238
generate separate qualified theorem name for each type's reach and take_lemma
 huffman parents: 
35630diff
changeset | 606 | (take_rews, thy) | 
| 23152 | 607 | end; (* let *) | 
| 608 | end; (* struct *) |