| author | paulson <lp15@cam.ac.uk> | 
| Thu, 26 Apr 2018 22:47:04 +0100 | |
| changeset 68048 | 0b4fb9fd91b1 | 
| parent 63003 | bf5fcc65586b | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Tools/Domain/domain_constructors.ML | 
| 35444 | 2 | Author: Brian Huffman | 
| 3 | ||
| 4 | Defines constructor functions for a given domain isomorphism | |
| 5 | and proves related theorems. | |
| 6 | *) | |
| 7 | ||
| 8 | signature DOMAIN_CONSTRUCTORS = | |
| 9 | sig | |
| 40014 | 10 | type constr_info = | 
| 11 |     {
 | |
| 40017 | 12 | iso_info : Domain_Take_Proofs.iso_info, | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 13 | con_specs : (term * (bool * typ) list) list, | 
| 40014 | 14 | con_betas : thm list, | 
| 15 | nchotomy : thm, | |
| 16 | exhaust : thm, | |
| 17 | compacts : thm list, | |
| 18 | con_rews : thm list, | |
| 19 | inverts : thm list, | |
| 20 | injects : thm list, | |
| 21 | dist_les : thm list, | |
| 22 | dist_eqs : thm list, | |
| 23 | cases : thm list, | |
| 24 | sel_rews : thm list, | |
| 25 | dis_rews : thm list, | |
| 26 | match_rews : thm list | |
| 27 | } | |
| 35444 | 28 | val add_domain_constructors : | 
| 35777 
bcc77916b7b9
pass binding as argument to add_domain_constructors; proper binding for case combinator
 huffman parents: 
35561diff
changeset | 29 | binding | 
| 35481 
7bb9157507a9
add_domain_constructors takes iso_info record as argument
 huffman parents: 
35476diff
changeset | 30 | -> (binding * (bool * binding option * typ) list * mixfix) list | 
| 35514 
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
 huffman parents: 
35496diff
changeset | 31 | -> Domain_Take_Proofs.iso_info | 
| 35444 | 32 | -> theory | 
| 40832 | 33 | -> constr_info * theory | 
| 34 | end | |
| 35444 | 35 | |
| 36 | ||
| 41296 
6aaf80ea9715
switch to transparent ascription, to avoid warning messages
 huffman parents: 
40853diff
changeset | 37 | structure Domain_Constructors : DOMAIN_CONSTRUCTORS = | 
| 35444 | 38 | struct | 
| 39 | ||
| 40832 | 40 | open HOLCF_Library | 
| 35561 | 41 | |
| 40832 | 42 | infixr 6 ->> | 
| 43 | infix -->> | |
| 44 | infix 9 ` | |
| 35444 | 45 | |
| 40014 | 46 | type constr_info = | 
| 47 |   {
 | |
| 40017 | 48 | iso_info : Domain_Take_Proofs.iso_info, | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 49 | con_specs : (term * (bool * typ) list) list, | 
| 40014 | 50 | con_betas : thm list, | 
| 51 | nchotomy : thm, | |
| 52 | exhaust : thm, | |
| 53 | compacts : thm list, | |
| 54 | con_rews : thm list, | |
| 55 | inverts : thm list, | |
| 56 | injects : thm list, | |
| 57 | dist_les : thm list, | |
| 58 | dist_eqs : thm list, | |
| 59 | cases : thm list, | |
| 60 | sel_rews : thm list, | |
| 61 | dis_rews : thm list, | |
| 62 | match_rews : thm list | |
| 63 | } | |
| 64 | ||
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 65 | (************************** miscellaneous functions ***************************) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 66 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48563diff
changeset | 67 | val simple_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48563diff
changeset | 68 |   simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms simp_thms})
 | 
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 69 | |
| 37078 
a1656804fcad
domain package internal proofs use fixed set of continuity rules, rather than taking cont2cont rules from context
 huffman parents: 
36998diff
changeset | 70 | val beta_rules = | 
| 40326 
73d45866dbda
renamed lemma cont2cont_Rep_CFun to cont2cont_APP
 huffman parents: 
40321diff
changeset | 71 |   @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @
 | 
| 40832 | 72 |   @{thms cont2cont_fst cont2cont_snd cont2cont_Pair}
 | 
| 37078 
a1656804fcad
domain package internal proofs use fixed set of continuity rules, rather than taking cont2cont rules from context
 huffman parents: 
36998diff
changeset | 73 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48563diff
changeset | 74 | val beta_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48563diff
changeset | 75 |   simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms simp_thms} @ beta_rules))
 | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 76 | |
| 35444 | 77 | fun define_consts | 
| 78 | (specs : (binding * term * mixfix) list) | |
| 79 | (thy : theory) | |
| 80 | : (term list * thm list) * theory = | |
| 81 | let | |
| 40832 | 82 | fun mk_decl (b, t, mx) = (b, fastype_of t, mx) | 
| 83 | val decls = map mk_decl specs | |
| 84 | val thy = Cont_Consts.add_consts decls thy | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 85 | fun mk_const (b, T, _) = Const (Sign.full_name thy b, T) | 
| 40832 | 86 | val consts = map mk_const decls | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 87 | fun mk_def c (b, t, _) = | 
| 46909 | 88 | (Thm.def_binding b, Logic.mk_equals (c, t)) | 
| 40832 | 89 | val defs = map2 mk_def consts specs | 
| 35444 | 90 | val (def_thms, thy) = | 
| 40832 | 91 | Global_Theory.add_defs false (map Thm.no_attributes defs) thy | 
| 35444 | 92 | in | 
| 93 | ((consts, def_thms), thy) | |
| 40832 | 94 | end | 
| 35444 | 95 | |
| 35449 | 96 | fun prove | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 97 | (thy : theory) | 
| 35449 | 98 | (defs : thm list) | 
| 35447 | 99 | (goal : term) | 
| 35449 | 100 |     (tacs : {prems: thm list, context: Proof.context} -> tactic list)
 | 
| 35447 | 101 | : thm = | 
| 35449 | 102 | let | 
| 103 |     fun tac {prems, context} =
 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 104 | rewrite_goals_tac context defs THEN | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 105 |       EVERY (tacs {prems = map (rewrite_rule context defs) prems, context = context})
 | 
| 35449 | 106 | in | 
| 107 | Goal.prove_global thy [] [] goal tac | |
| 40832 | 108 | end | 
| 35445 | 109 | |
| 35483 | 110 | fun get_vars_avoiding | 
| 111 | (taken : string list) | |
| 112 | (args : (bool * typ) list) | |
| 113 | : (term list * term list) = | |
| 114 | let | |
| 40832 | 115 | val Ts = map snd args | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
54895diff
changeset | 116 | val ns = Name.variant_list taken (Old_Datatype_Prop.make_tnames Ts) | 
| 40832 | 117 | val vs = map Free (ns ~~ Ts) | 
| 118 | val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs)) | |
| 35483 | 119 | in | 
| 120 | (vs, nonlazy) | |
| 40832 | 121 | end | 
| 35483 | 122 | |
| 40832 | 123 | fun get_vars args = get_vars_avoiding [] args | 
| 35483 | 124 | |
| 35445 | 125 | (************** generating beta reduction rules from definitions **************) | 
| 35444 | 126 | |
| 35445 | 127 | local | 
| 128 | fun arglist (Const _ $ Abs (s, T, t)) = | |
| 129 | let | |
| 40832 | 130 | val arg = Free (s, T) | 
| 131 | val (args, body) = arglist (subst_bound (arg, t)) | |
| 35445 | 132 | in (arg :: args, body) end | 
| 40832 | 133 | | arglist t = ([], t) | 
| 35445 | 134 | in | 
| 135 | fun beta_of_def thy def_thm = | |
| 136 | let | |
| 48563 
04e129931181
unvarify thm statement stemming from old-style definition, to avoid schematic type variables in subsequent goal;
 wenzelm parents: 
46909diff
changeset | 137 | val (con, lam) = | 
| 59582 | 138 | Logic.dest_equals (Logic.unvarify_global (Thm.concl_of def_thm)) | 
| 40832 | 139 | val (args, rhs) = arglist lam | 
| 140 | val lhs = list_ccomb (con, args) | |
| 141 | val goal = mk_equals (lhs, rhs) | |
| 142 | val cs = ContProc.cont_thms lam | |
| 143 |         val betas = map (fn c => mk_meta_eq (c RS @{thm beta_cfun})) cs
 | |
| 35445 | 144 | in | 
| 60754 | 145 | prove thy (def_thm::betas) goal | 
| 146 |           (fn {context = ctxt, ...} => [resolve_tac ctxt [reflexive_thm] 1])
 | |
| 40832 | 147 | end | 
| 148 | end | |
| 35444 | 149 | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 150 | (******************************************************************************) | 
| 35453 | 151 | (************* definitions and theorems for constructor functions *************) | 
| 152 | (******************************************************************************) | |
| 153 | ||
| 154 | fun add_constructors | |
| 35454 
cf6ba350b7ca
remove unnecessary stuff from argument to add_constructors function
 huffman parents: 
35453diff
changeset | 155 | (spec : (binding * (bool * typ) list * mixfix) list) | 
| 35453 | 156 | (abs_const : term) | 
| 157 | (iso_locale : thm) | |
| 158 | (thy : theory) | |
| 159 | = | |
| 160 | let | |
| 161 | ||
| 162 | (* get theorems about rep and abs *) | |
| 40832 | 163 |     val abs_strict = iso_locale RS @{thm iso.abs_strict}
 | 
| 35453 | 164 | |
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 165 | (* get types of type isomorphism *) | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 166 | val (_, lhsT) = dest_cfunT (fastype_of abs_const) | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 167 | |
| 35455 | 168 | fun vars_of args = | 
| 169 | let | |
| 40832 | 170 | val Ts = map snd args | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
54895diff
changeset | 171 | val ns = Old_Datatype_Prop.make_tnames Ts | 
| 35455 | 172 | in | 
| 173 | map Free (ns ~~ Ts) | |
| 40832 | 174 | end | 
| 35455 | 175 | |
| 35453 | 176 | (* define constructor functions *) | 
| 177 | val ((con_consts, con_defs), thy) = | |
| 178 | let | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 179 | fun one_arg (lazy, _) var = if lazy then mk_up var else var | 
| 40832 | 180 | fun one_con (_,args,_) = mk_stuple (map2 one_arg args (vars_of args)) | 
| 181 | fun mk_abs t = abs_const ` t | |
| 182 | val rhss = map mk_abs (mk_sinjects (map one_con spec)) | |
| 35453 | 183 | fun mk_def (bind, args, mx) rhs = | 
| 40832 | 184 | (bind, big_lambdas (vars_of args) rhs, mx) | 
| 35453 | 185 | in | 
| 186 | define_consts (map2 mk_def spec rhss) thy | |
| 40832 | 187 | end | 
| 35453 | 188 | |
| 189 | (* prove beta reduction rules for constructors *) | |
| 40832 | 190 | val con_betas = map (beta_of_def thy) con_defs | 
| 35453 | 191 | |
| 192 | (* replace bindings with terms in constructor spec *) | |
| 193 | val spec' : (term * (bool * typ) list) list = | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 194 | let fun one_con con (_, args, _) = (con, args) | 
| 40832 | 195 | in map2 one_con con_consts spec end | 
| 35453 | 196 | |
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 197 | (* prove exhaustiveness of constructors *) | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 198 | local | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 199 |       fun arg2typ n (true,  _) = (n+1, mk_upT (TVar (("'a", n), @{sort cpo})))
 | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 200 |         | arg2typ n (false, _) = (n+1, TVar (("'a", n), @{sort pcpo}))
 | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 201 | fun args2typ n [] = (n, oneT) | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 202 | | args2typ n [arg] = arg2typ n arg | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 203 | | args2typ n (arg::args) = | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 204 | let | 
| 40832 | 205 | val (n1, t1) = arg2typ n arg | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 206 | val (n2, t2) = args2typ n1 args | 
| 40832 | 207 | in (n2, mk_sprodT (t1, t2)) end | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 208 | fun cons2typ n [] = (n, oneT) | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 209 | | cons2typ n [con] = args2typ n (snd con) | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 210 | | cons2typ n (con::cons) = | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 211 | let | 
| 40832 | 212 | val (n1, t1) = args2typ n (snd con) | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 213 | val (n2, t2) = cons2typ n1 cons | 
| 40832 | 214 | in (n2, mk_ssumT (t1, t2)) end | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 215 | val ct = Thm.global_ctyp_of thy (snd (cons2typ 1 spec')) | 
| 60801 | 216 |       val thm1 = Thm.instantiate' [SOME ct] [] @{thm exh_start}
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 217 | val thm2 = rewrite_rule (Proof_Context.init_global thy) | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 218 |         (map mk_meta_eq @{thms ex_bottom_iffs}) thm1
 | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 219 | val thm3 = rewrite_rule (Proof_Context.init_global thy) | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 220 |         [mk_meta_eq @{thm conj_assoc}] thm2
 | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 221 | |
| 40832 | 222 |       val y = Free ("y", lhsT)
 | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 223 | fun one_con (con, args) = | 
| 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 224 | let | 
| 40832 | 225 | val (vs, nonlazy) = get_vars_avoiding ["y"] args | 
| 226 | val eqn = mk_eq (y, list_ccomb (con, vs)) | |
| 227 | val conj = foldr1 mk_conj (eqn :: map mk_defined nonlazy) | |
| 228 | in Library.foldr mk_ex (vs, conj) end | |
| 229 | val goal = mk_trp (foldr1 mk_disj (mk_undef y :: map one_con spec')) | |
| 41429 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 huffman parents: 
41296diff
changeset | 230 | (* first rules replace "y = bottom \/ P" with "rep$y = bottom \/ P" *) | 
| 54895 | 231 | fun tacs ctxt = [ | 
| 60754 | 232 |           resolve_tac ctxt [iso_locale RS @{thm iso.casedist_rule}] 1,
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 233 |           rewrite_goals_tac ctxt [mk_meta_eq (iso_locale RS @{thm iso.iso_swap})],
 | 
| 60754 | 234 | resolve_tac ctxt [thm3] 1] | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 235 | in | 
| 54895 | 236 | val nchotomy = prove thy con_betas goal (tacs o #context) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 237 | val exhaust = | 
| 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 238 |           (nchotomy RS @{thm exh_casedist0})
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 239 |           |> rewrite_rule (Proof_Context.init_global thy) @{thms exh_casedists}
 | 
| 40832 | 240 | |> Drule.zero_var_indexes | 
| 241 | end | |
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 242 | |
| 35453 | 243 | (* prove compactness rules for constructors *) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 244 | val compacts = | 
| 35453 | 245 | let | 
| 246 |         val rules = @{thms compact_sinl compact_sinr compact_spair
 | |
| 40832 | 247 | compact_up compact_ONE} | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 248 | fun tacs ctxt = | 
| 60754 | 249 |           [resolve_tac ctxt [iso_locale RS @{thm iso.compact_abs}] 1,
 | 
| 250 | REPEAT (resolve_tac ctxt rules 1 ORELSE assume_tac ctxt 1)] | |
| 35453 | 251 | fun con_compact (con, args) = | 
| 252 | let | |
| 40832 | 253 | val vs = vars_of args | 
| 254 | val con_app = list_ccomb (con, vs) | |
| 255 | val concl = mk_trp (mk_compact con_app) | |
| 256 | val assms = map (mk_trp o mk_compact) vs | |
| 257 | val goal = Logic.list_implies (assms, concl) | |
| 35453 | 258 | in | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 259 | prove thy con_betas goal (tacs o #context) | 
| 40832 | 260 | end | 
| 35453 | 261 | in | 
| 262 | map con_compact spec' | |
| 40832 | 263 | end | 
| 35453 | 264 | |
| 265 | (* prove strictness rules for constructors *) | |
| 266 | local | |
| 46490 
e4863ab5e09b
more basic simplification, eliminated slightly odd tactic style from 1995 (cf. ea0668a1c0ba);
 wenzelm parents: 
46125diff
changeset | 267 | fun con_strict (con, args) = | 
| 35453 | 268 | let | 
| 40832 | 269 |           val rules = abs_strict :: @{thms con_strict_rules}
 | 
| 270 | val (vs, nonlazy) = get_vars args | |
| 35453 | 271 | fun one_strict v' = | 
| 272 | let | |
| 41429 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 huffman parents: 
41296diff
changeset | 273 | val bottom = mk_bottom (fastype_of v') | 
| 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 huffman parents: 
41296diff
changeset | 274 | val vs' = map (fn v => if v = v' then bottom else v) vs | 
| 40832 | 275 | val goal = mk_trp (mk_undef (list_ccomb (con, vs'))) | 
| 54895 | 276 | fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1] | 
| 277 | in prove thy con_betas goal (tacs o #context) end | |
| 40832 | 278 | in map one_strict nonlazy end | 
| 35453 | 279 | |
| 280 | fun con_defin (con, args) = | |
| 281 | let | |
| 282 | fun iff_disj (t, []) = HOLogic.mk_not t | |
| 40832 | 283 | | iff_disj (t, ts) = mk_eq (t, foldr1 HOLogic.mk_disj ts) | 
| 284 | val (vs, nonlazy) = get_vars args | |
| 285 | val lhs = mk_undef (list_ccomb (con, vs)) | |
| 286 | val rhss = map mk_undef nonlazy | |
| 287 | val goal = mk_trp (iff_disj (lhs, rhss)) | |
| 288 |           val rule1 = iso_locale RS @{thm iso.abs_bottom_iff}
 | |
| 289 |           val rules = rule1 :: @{thms con_bottom_iff_rules}
 | |
| 54895 | 290 | fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1] | 
| 291 | in prove thy con_betas goal (tacs o #context) end | |
| 35453 | 292 | in | 
| 40832 | 293 | val con_stricts = maps con_strict spec' | 
| 294 | val con_defins = map con_defin spec' | |
| 295 | val con_rews = con_stricts @ con_defins | |
| 296 | end | |
| 35453 | 297 | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 298 | (* prove injectiveness of constructors *) | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 299 | local | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 300 | fun pgterm rel (con, args) = | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 301 | let | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 302 | fun prime (Free (n, T)) = Free (n^"'", T) | 
| 40832 | 303 | | prime t = t | 
| 304 | val (xs, nonlazy) = get_vars args | |
| 305 | val ys = map prime xs | |
| 306 | val lhs = rel (list_ccomb (con, xs), list_ccomb (con, ys)) | |
| 307 | val rhs = foldr1 mk_conj (ListPair.map rel (xs, ys)) | |
| 308 | val concl = mk_trp (mk_eq (lhs, rhs)) | |
| 309 | val zs = case args of [_] => [] | _ => nonlazy | |
| 310 | val assms = map (mk_trp o mk_defined) zs | |
| 311 | val goal = Logic.list_implies (assms, concl) | |
| 312 | in prove thy con_betas goal end | |
| 313 | val cons' = filter (fn (_, args) => not (null args)) spec' | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 314 | in | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 315 | val inverts = | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 316 | let | 
| 40832 | 317 |           val abs_below = iso_locale RS @{thm iso.abs_below}
 | 
| 318 |           val rules1 = abs_below :: @{thms sinl_below sinr_below spair_below up_below}
 | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 319 |           val rules2 = @{thms up_defined spair_defined ONE_defined}
 | 
| 40832 | 320 | val rules = rules1 @ rules2 | 
| 54895 | 321 | fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1] | 
| 322 | in map (fn c => pgterm mk_below c (tacs o #context)) cons' end | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 323 | val injects = | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 324 | let | 
| 40832 | 325 |           val abs_eq = iso_locale RS @{thm iso.abs_eq}
 | 
| 326 |           val rules1 = abs_eq :: @{thms sinl_eq sinr_eq spair_eq up_eq}
 | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 327 |           val rules2 = @{thms up_defined spair_defined ONE_defined}
 | 
| 40832 | 328 | val rules = rules1 @ rules2 | 
| 54895 | 329 | fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1] | 
| 330 | in map (fn c => pgterm mk_eq c (tacs o #context)) cons' end | |
| 40832 | 331 | end | 
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 332 | |
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 333 | (* prove distinctness of constructors *) | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 334 | local | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 335 | fun map_dist (f : 'a -> 'a -> 'b) (xs : 'a list) : 'b list = | 
| 40832 | 336 | flat (map_index (fn (i, x) => map (f x) (nth_drop i xs)) xs) | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 337 | fun prime (Free (n, T)) = Free (n^"'", T) | 
| 40832 | 338 | | prime t = t | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 339 | fun iff_disj (t, []) = mk_not t | 
| 40832 | 340 | | iff_disj (t, ts) = mk_eq (t, foldr1 mk_disj ts) | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 341 | fun iff_disj2 (t, [], _) = mk_not t | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 342 | | iff_disj2 (t, _, []) = mk_not t | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 343 | | iff_disj2 (t, ts, us) = | 
| 40832 | 344 | mk_eq (t, mk_conj (foldr1 mk_disj ts, foldr1 mk_disj us)) | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 345 | fun dist_le (con1, args1) (con2, args2) = | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 346 | let | 
| 40832 | 347 | val (vs1, zs1) = get_vars args1 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58112diff
changeset | 348 | val (vs2, _) = get_vars args2 |> apply2 (map prime) | 
| 40832 | 349 | val lhs = mk_below (list_ccomb (con1, vs1), list_ccomb (con2, vs2)) | 
| 350 | val rhss = map mk_undef zs1 | |
| 351 | val goal = mk_trp (iff_disj (lhs, rhss)) | |
| 352 |           val rule1 = iso_locale RS @{thm iso.abs_below}
 | |
| 353 |           val rules = rule1 :: @{thms con_below_iff_rules}
 | |
| 54895 | 354 | fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1] | 
| 355 | in prove thy con_betas goal (tacs o #context) end | |
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 356 | fun dist_eq (con1, args1) (con2, args2) = | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 357 | let | 
| 40832 | 358 | val (vs1, zs1) = get_vars args1 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58112diff
changeset | 359 | val (vs2, zs2) = get_vars args2 |> apply2 (map prime) | 
| 40832 | 360 | val lhs = mk_eq (list_ccomb (con1, vs1), list_ccomb (con2, vs2)) | 
| 361 | val rhss1 = map mk_undef zs1 | |
| 362 | val rhss2 = map mk_undef zs2 | |
| 363 | val goal = mk_trp (iff_disj2 (lhs, rhss1, rhss2)) | |
| 364 |           val rule1 = iso_locale RS @{thm iso.abs_eq}
 | |
| 365 |           val rules = rule1 :: @{thms con_eq_iff_rules}
 | |
| 54895 | 366 | fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1] | 
| 367 | in prove thy con_betas goal (tacs o #context) end | |
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 368 | in | 
| 40832 | 369 | val dist_les = map_dist dist_le spec' | 
| 370 | val dist_eqs = map_dist dist_eq spec' | |
| 371 | end | |
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 372 | |
| 35453 | 373 | val result = | 
| 374 |       {
 | |
| 375 | con_consts = con_consts, | |
| 376 | con_betas = con_betas, | |
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 377 | nchotomy = nchotomy, | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 378 | exhaust = exhaust, | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 379 | compacts = compacts, | 
| 35456 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 380 | con_rews = con_rews, | 
| 
5356534f880e
move proofs of injects and inverts to domain_constructors.ML
 huffman parents: 
35455diff
changeset | 381 | inverts = inverts, | 
| 35458 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 382 | injects = injects, | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 383 | dist_les = dist_les, | 
| 
deaf221c4a59
moved proofs of dist_les and dist_eqs to domain_constructors.ML
 huffman parents: 
35457diff
changeset | 384 | dist_eqs = dist_eqs | 
| 40832 | 385 | } | 
| 35453 | 386 | in | 
| 387 | (result, thy) | |
| 40832 | 388 | end | 
| 35453 | 389 | |
| 390 | (******************************************************************************) | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 391 | (**************** definition and theorems for case combinator *****************) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 392 | (******************************************************************************) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 393 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 394 | fun add_case_combinator | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 395 | (spec : (term * (bool * typ) list) list) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 396 | (lhsT : typ) | 
| 35777 
bcc77916b7b9
pass binding as argument to add_domain_constructors; proper binding for case combinator
 huffman parents: 
35561diff
changeset | 397 | (dbind : binding) | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 398 | (con_betas : thm list) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 399 | (iso_locale : thm) | 
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 400 | (rep_const : term) | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 401 | (thy : theory) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 402 | : ((typ -> term) * thm list) * theory = | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 403 | let | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 404 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 405 | (* prove rep/abs rules *) | 
| 40832 | 406 |     val rep_strict = iso_locale RS @{thm iso.rep_strict}
 | 
| 407 |     val abs_inverse = iso_locale RS @{thm iso.abs_iso}
 | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 408 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 409 | (* calculate function arguments of case combinator *) | 
| 40832 | 410 | val tns = map fst (Term.add_tfreesT lhsT []) | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42290diff
changeset | 411 |     val resultT = TFree (singleton (Name.variant_list tns) "'t", @{sort pcpo})
 | 
| 40832 | 412 | fun fTs T = map (fn (_, args) => map snd args -->> T) spec | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
54895diff
changeset | 413 | val fns = Old_Datatype_Prop.indexify_names (map (K "f") spec) | 
| 40832 | 414 | val fs = map Free (fns ~~ fTs resultT) | 
| 415 | fun caseT T = fTs T -->> (lhsT ->> T) | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 416 | |
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 417 | (* definition of case combinator *) | 
| 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 418 | local | 
| 40832 | 419 | val case_bind = Binding.suffix_name "_case" dbind | 
| 35784 | 420 | fun lambda_arg (lazy, v) t = | 
| 40832 | 421 | (if lazy then mk_fup else I) (big_lambda v t) | 
| 40212 | 422 | fun lambda_args [] t = mk_one_case t | 
| 35784 | 423 | | lambda_args (x::[]) t = lambda_arg x t | 
| 40832 | 424 | | lambda_args (x::xs) t = mk_ssplit (lambda_arg x (lambda_args xs t)) | 
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 425 | fun one_con f (_, args) = | 
| 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 426 | let | 
| 40832 | 427 | val Ts = map snd args | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
54895diff
changeset | 428 | val ns = Name.variant_list fns (Old_Datatype_Prop.make_tnames Ts) | 
| 40832 | 429 | val vs = map Free (ns ~~ Ts) | 
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 430 | in | 
| 35784 | 431 | lambda_args (map fst args ~~ vs) (list_ccomb (f, vs)) | 
| 40832 | 432 | end | 
| 35785 | 433 | fun mk_sscases [t] = mk_strictify t | 
| 40832 | 434 | | mk_sscases ts = foldr1 mk_sscase ts | 
| 435 | val body = mk_sscases (map2 one_con fs spec) | |
| 436 | val rhs = big_lambdas fs (mk_cfcomp (body, rep_const)) | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 437 | val ((_, case_defs), thy) = | 
| 40832 | 438 | define_consts [(case_bind, rhs, NoSyn)] thy | 
| 439 | val case_name = Sign.full_name thy case_bind | |
| 35486 
c91854705b1d
move definition of case combinator to domain_constructors.ML
 huffman parents: 
35485diff
changeset | 440 | in | 
| 40832 | 441 | val case_def = hd case_defs | 
| 442 | fun case_const T = Const (case_name, caseT T) | |
| 443 | val case_app = list_ccomb (case_const resultT, fs) | |
| 444 | val thy = thy | |
| 445 | end | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 446 | |
| 35472 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 447 | (* define syntax for case combinator *) | 
| 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 448 | (* TODO: re-implement case syntax using a parse translation *) | 
| 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 449 | local | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 450 | fun syntax c = Lexicon.mark_const (fst (dest_Const c)) | 
| 40832 | 451 | fun xconst c = Long_Name.base_name (fst (dest_Const c)) | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 452 | fun c_ast authentic con = Ast.Constant (if authentic then syntax con else xconst con) | 
| 40832 | 453 | fun showint n = string_of_int (n+1) | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 454 |       fun expvar n = Ast.Variable ("e" ^ showint n)
 | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 455 |       fun argvar n (m, _) = Ast.Variable ("a" ^ showint n ^ "_" ^ showint m)
 | 
| 40832 | 456 | fun argvars n args = map_index (argvar n) args | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 457 | fun app s (l, r) = Ast.mk_appl (Ast.Constant s) [l, r] | 
| 40832 | 458 | val cabs = app "_cabs" | 
| 459 |       val capp = app @{const_syntax Rep_cfun}
 | |
| 35472 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 460 | val capps = Library.foldl capp | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 461 | fun con1 authentic n (con, args) = | 
| 40832 | 462 | Library.foldl capp (c_ast authentic con, argvars n args) | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 463 | fun con1_constraint authentic n (con, args) = | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 464 | Library.foldl capp | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 465 | (Ast.Appl | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 466 |               [Ast.Constant @{syntax_const "_constrain"}, c_ast authentic con,
 | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 467 |                 Ast.Variable ("'a" ^ string_of_int n)],
 | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 468 | argvars n args) | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 469 | fun case1 constraint authentic (n, c) = | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 470 |         app @{syntax_const "_case1"}
 | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 471 | ((if constraint then con1_constraint else con1) authentic n c, expvar n) | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 472 | fun arg1 (n, (_, args)) = List.foldr cabs (expvar n) (argvars n args) | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 473 |       fun when1 n (m, c) = if n = m then arg1 (n, c) else Ast.Constant @{const_syntax bottom}
 | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 474 | val case_constant = Ast.Constant (syntax (case_const dummyT)) | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 475 | fun case_trans constraint authentic = | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 476 | (app "_case_syntax" | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 477 | (Ast.Variable "x", | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 478 |              foldr1 (app @{syntax_const "_case2"}) (map_index (case1 constraint authentic) spec)),
 | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 479 | capp (capps (case_constant, map_index arg1 spec), Ast.Variable "x")) | 
| 35485 | 480 | fun one_abscon_trans authentic (n, c) = | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 481 | (if authentic then Syntax.Parse_Print_Rule else Syntax.Parse_Rule) | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 482 | (cabs (con1 authentic n c, expvar n), | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 483 | capps (case_constant, map_index (when1 n) spec)) | 
| 35472 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 484 | fun abscon_trans authentic = | 
| 40832 | 485 | map_index (one_abscon_trans authentic) spec | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 486 | val trans_rules : Ast.ast Syntax.trrule list = | 
| 46125 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 487 | Syntax.Parse_Print_Rule (case_trans false true) :: | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 488 | Syntax.Parse_Rule (case_trans false false) :: | 
| 
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
 wenzelm parents: 
45654diff
changeset | 489 | Syntax.Parse_Rule (case_trans true false) :: | 
| 40832 | 490 | abscon_trans false @ abscon_trans true | 
| 35472 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 491 | in | 
| 42204 | 492 | val thy = Sign.add_trrules trans_rules thy | 
| 40832 | 493 | end | 
| 35472 
c23b42730b9b
move case combinator syntax to domain_constructors.ML
 huffman parents: 
35470diff
changeset | 494 | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 495 | (* prove beta reduction rule for case combinator *) | 
| 40832 | 496 | val case_beta = beta_of_def thy case_def | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 497 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 498 | (* prove strictness of case combinator *) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 499 | val case_strict = | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 500 | let | 
| 40832 | 501 |         val defs = case_beta :: map mk_meta_eq [rep_strict, @{thm cfcomp2}]
 | 
| 502 | val goal = mk_trp (mk_strict case_app) | |
| 503 |         val rules = @{thms sscase1 ssplit1 strictify1 one_case1}
 | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 504 | fun tacs ctxt = [resolve_tac ctxt rules 1] | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 505 | in prove thy defs goal (tacs o #context) end | 
| 46490 
e4863ab5e09b
more basic simplification, eliminated slightly odd tactic style from 1995 (cf. ea0668a1c0ba);
 wenzelm parents: 
46125diff
changeset | 506 | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 507 | (* prove rewrites for case combinator *) | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 508 | local | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 509 | fun one_case (con, args) f = | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 510 | let | 
| 40832 | 511 | val (vs, nonlazy) = get_vars args | 
| 512 | val assms = map (mk_trp o mk_defined) nonlazy | |
| 513 | val lhs = case_app ` list_ccomb (con, vs) | |
| 514 | val rhs = list_ccomb (f, vs) | |
| 515 | val concl = mk_trp (mk_eq (lhs, rhs)) | |
| 516 | val goal = Logic.list_implies (assms, concl) | |
| 517 | val defs = case_beta :: con_betas | |
| 518 |           val rules1 = @{thms strictify2 sscase2 sscase3 ssplit2 fup2 ID1}
 | |
| 519 |           val rules2 = @{thms con_bottom_iff_rules}
 | |
| 520 |           val rules3 = @{thms cfcomp2 one_case2}
 | |
| 521 | val rules = abs_inverse :: rules1 @ rules2 @ rules3 | |
| 54895 | 522 | fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps rules) 1] | 
| 523 | in prove thy defs goal (tacs o #context) end | |
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 524 | in | 
| 40832 | 525 | val case_apps = map2 one_case spec fs | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 526 | end | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 527 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 528 | in | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 529 | ((case_const, case_strict :: case_apps), thy) | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 530 | end | 
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 531 | |
| 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 532 | (******************************************************************************) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 533 | (************** definitions and theorems for selector functions ***************) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 534 | (******************************************************************************) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 535 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 536 | fun add_selectors | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 537 | (spec : (term * (bool * binding option * typ) list) list) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 538 | (rep_const : term) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 539 | (abs_inv : thm) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 540 | (rep_strict : thm) | 
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40214diff
changeset | 541 | (rep_bottom_iff : thm) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 542 | (con_betas : thm list) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 543 | (thy : theory) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 544 | : thm list * theory = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 545 | let | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 546 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 547 | (* define selector functions *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 548 | val ((sel_consts, sel_defs), thy) = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 549 | let | 
| 40832 | 550 | fun rangeT s = snd (dest_cfunT (fastype_of s)) | 
| 551 | fun mk_outl s = mk_cfcomp (from_sinl (dest_ssumT (rangeT s)), s) | |
| 552 | fun mk_outr s = mk_cfcomp (from_sinr (dest_ssumT (rangeT s)), s) | |
| 553 | fun mk_sfst s = mk_cfcomp (sfst_const (dest_sprodT (rangeT s)), s) | |
| 554 | fun mk_ssnd s = mk_cfcomp (ssnd_const (dest_sprodT (rangeT s)), s) | |
| 555 | fun mk_down s = mk_cfcomp (from_up (dest_upT (rangeT s)), s) | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 556 | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 557 | fun sels_of_arg _ (_, NONE, _) = [] | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 558 | | sels_of_arg s (lazy, SOME b, _) = | 
| 40832 | 559 | [(b, if lazy then mk_down s else s, NoSyn)] | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 560 | fun sels_of_args _ [] = [] | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 561 | | sels_of_args s (v :: []) = sels_of_arg s v | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 562 | | sels_of_args s (v :: vs) = | 
| 40832 | 563 | sels_of_arg (mk_sfst s) v @ sels_of_args (mk_ssnd s) vs | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 564 | fun sels_of_cons _ [] = [] | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 565 | | sels_of_cons s ((_, args) :: []) = sels_of_args s args | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 566 | | sels_of_cons s ((_, args) :: cs) = | 
| 40832 | 567 | sels_of_args (mk_outl s) args @ sels_of_cons (mk_outr s) cs | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 568 | val sel_eqns : (binding * term * mixfix) list = | 
| 40832 | 569 | sels_of_cons rep_const spec | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 570 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 571 | define_consts sel_eqns thy | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 572 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 573 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 574 | (* replace bindings with terms in constructor spec *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 575 | val spec2 : (term * (bool * term option * typ) list) list = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 576 | let | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 577 | fun prep_arg (lazy, NONE, T) sels = ((lazy, NONE, T), sels) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 578 | | prep_arg (lazy, SOME _, T) sels = | 
| 40832 | 579 | ((lazy, SOME (hd sels), T), tl sels) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 580 | fun prep_con (con, args) sels = | 
| 40832 | 581 | apfst (pair con) (fold_map prep_arg args sels) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 582 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 583 | fst (fold_map prep_con spec sel_consts) | 
| 40832 | 584 | end | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 585 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 586 | (* prove selector strictness rules *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 587 | val sel_stricts : thm list = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 588 | let | 
| 40832 | 589 |         val rules = rep_strict :: @{thms sel_strict_rules}
 | 
| 54895 | 590 | fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1] | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 591 | fun sel_strict sel = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 592 | let | 
| 40832 | 593 | val goal = mk_trp (mk_strict sel) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 594 | in | 
| 54895 | 595 | prove thy sel_defs goal (tacs o #context) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 596 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 597 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 598 | map sel_strict sel_consts | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 599 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 600 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 601 | (* prove selector application rules *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 602 | val sel_apps : thm list = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 603 | let | 
| 40832 | 604 | val defs = con_betas @ sel_defs | 
| 605 |         val rules = abs_inv :: @{thms sel_app_rules}
 | |
| 54895 | 606 | fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1] | 
| 37165 | 607 | fun sel_apps_of (i, (con, args: (bool * term option * typ) list)) = | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 608 | let | 
| 40832 | 609 | val Ts : typ list = map #3 args | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
54895diff
changeset | 610 | val ns : string list = Old_Datatype_Prop.make_tnames Ts | 
| 40832 | 611 | val vs : term list = map Free (ns ~~ Ts) | 
| 612 | val con_app : term = list_ccomb (con, vs) | |
| 613 | val vs' : (bool * term) list = map #1 args ~~ vs | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 614 | fun one_same (n, sel, _) = | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 615 | let | 
| 40832 | 616 | val xs = map snd (filter_out fst (nth_drop n vs')) | 
| 617 | val assms = map (mk_trp o mk_defined) xs | |
| 618 | val concl = mk_trp (mk_eq (sel ` con_app, nth vs n)) | |
| 619 | val goal = Logic.list_implies (assms, concl) | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 620 | in | 
| 54895 | 621 | prove thy defs goal (tacs o #context) | 
| 40832 | 622 | end | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 623 | fun one_diff (_, sel, T) = | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 624 | let | 
| 40832 | 625 | val goal = mk_trp (mk_eq (sel ` con_app, mk_bottom T)) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 626 | in | 
| 54895 | 627 | prove thy defs goal (tacs o #context) | 
| 40832 | 628 | end | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 629 | fun one_con (j, (_, args')) : thm list = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 630 | let | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 631 | fun prep (_, (_, NONE, _)) = NONE | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 632 | | prep (i, (_, SOME sel, T)) = SOME (i, sel, T) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 633 | val sels : (int * term * typ) list = | 
| 40832 | 634 | map_filter prep (map_index I args') | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 635 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 636 | if i = j | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 637 | then map one_same sels | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 638 | else map one_diff sels | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 639 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 640 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 641 | flat (map_index one_con spec2) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 642 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 643 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 644 | flat (map_index sel_apps_of spec2) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 645 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 646 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 647 | (* prove selector definedness rules *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 648 | val sel_defins : thm list = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 649 | let | 
| 40832 | 650 |         val rules = rep_bottom_iff :: @{thms sel_bottom_iff_rules}
 | 
| 54895 | 651 | fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1] | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 652 | fun sel_defin sel = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 653 | let | 
| 40832 | 654 | val (T, U) = dest_cfunT (fastype_of sel) | 
| 655 |             val x = Free ("x", T)
 | |
| 656 | val lhs = mk_eq (sel ` x, mk_bottom U) | |
| 657 | val rhs = mk_eq (x, mk_bottom T) | |
| 658 | val goal = mk_trp (mk_eq (lhs, rhs)) | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 659 | in | 
| 54895 | 660 | prove thy sel_defs goal (tacs o #context) | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 661 | end | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 662 | fun one_arg (false, SOME sel, _) = SOME (sel_defin sel) | 
| 40832 | 663 | | one_arg _ = NONE | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 664 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 665 | case spec2 of | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 666 | [(_, args)] => map_filter one_arg args | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 667 | | _ => [] | 
| 40832 | 668 | end | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 669 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 670 | in | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 671 | (sel_stricts @ sel_defins @ sel_apps, thy) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 672 | end | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 673 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 674 | (******************************************************************************) | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 675 | (************ definitions and theorems for discriminator functions ************) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 676 | (******************************************************************************) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 677 | |
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 678 | fun add_discriminators | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 679 | (bindings : binding list) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 680 | (spec : (term * (bool * typ) list) list) | 
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 681 | (lhsT : typ) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 682 | (exhaust : thm) | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 683 | (case_const : typ -> term) | 
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 684 | (case_rews : thm list) | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 685 | (thy : theory) = | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 686 | let | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 687 | |
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 688 | (* define discriminator functions *) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 689 | local | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 690 | fun dis_fun i (j, (_, args)) = | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 691 | let | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 692 | val (vs, _) = get_vars args | 
| 40832 | 693 |           val tr = if i = j then @{term TT} else @{term FF}
 | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 694 | in | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 695 | big_lambdas vs tr | 
| 40832 | 696 | end | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 697 | fun dis_eqn (i, bind) : binding * term * mixfix = | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 698 | let | 
| 40832 | 699 | val dis_bind = Binding.prefix_name "is_" bind | 
| 700 | val rhs = list_ccomb (case_const trT, map_index (dis_fun i) spec) | |
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 701 | in | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 702 | (dis_bind, rhs, NoSyn) | 
| 40832 | 703 | end | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 704 | in | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 705 | val ((dis_consts, dis_defs), thy) = | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 706 | define_consts (map_index dis_eqn bindings) thy | 
| 40832 | 707 | end | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 708 | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 709 | (* prove discriminator strictness rules *) | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 710 | local | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 711 | fun dis_strict dis = | 
| 40832 | 712 | let val goal = mk_trp (mk_strict dis) | 
| 60754 | 713 | in | 
| 714 | prove thy dis_defs goal | |
| 715 |             (fn {context = ctxt, ...} => [resolve_tac ctxt [hd case_rews] 1])
 | |
| 716 | end | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 717 | in | 
| 40832 | 718 | val dis_stricts = map dis_strict dis_consts | 
| 719 | end | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 720 | |
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 721 | (* prove discriminator/constructor rules *) | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 722 | local | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 723 | fun dis_app (i, dis) (j, (con, args)) = | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 724 | let | 
| 40832 | 725 | val (vs, nonlazy) = get_vars args | 
| 726 | val lhs = dis ` list_ccomb (con, vs) | |
| 727 |           val rhs = if i = j then @{term TT} else @{term FF}
 | |
| 728 | val assms = map (mk_trp o mk_defined) nonlazy | |
| 729 | val concl = mk_trp (mk_eq (lhs, rhs)) | |
| 730 | val goal = Logic.list_implies (assms, concl) | |
| 54895 | 731 | fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1] | 
| 732 | in prove thy dis_defs goal (tacs o #context) end | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 733 | fun one_dis (i, dis) = | 
| 40832 | 734 | map_index (dis_app (i, dis)) spec | 
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 735 | in | 
| 40832 | 736 | val dis_apps = flat (map_index one_dis dis_consts) | 
| 737 | end | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 738 | |
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 739 | (* prove discriminator definedness rules *) | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 740 | local | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 741 | fun dis_defin dis = | 
| 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 742 | let | 
| 40832 | 743 |           val x = Free ("x", lhsT)
 | 
| 744 |           val simps = dis_apps @ @{thms dist_eq_tr}
 | |
| 54895 | 745 | fun tacs ctxt = | 
| 60754 | 746 |             [resolve_tac ctxt @{thms iffI} 1,
 | 
| 54895 | 747 | asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps dis_stricts) 2, | 
| 60754 | 748 | resolve_tac ctxt [exhaust] 1, assume_tac ctxt 1, | 
| 54895 | 749 | ALLGOALS (asm_full_simp_tac (put_simpset simple_ss ctxt addsimps simps))] | 
| 40832 | 750 | val goal = mk_trp (mk_eq (mk_undef (dis ` x), mk_undef x)) | 
| 54895 | 751 | in prove thy [] goal (tacs o #context) end | 
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 752 | in | 
| 40832 | 753 | val dis_defins = map dis_defin dis_consts | 
| 754 | end | |
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 755 | |
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 756 | in | 
| 35461 
34360a1e3537
move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
 huffman parents: 
35460diff
changeset | 757 | (dis_stricts @ dis_defins @ dis_apps, thy) | 
| 40832 | 758 | end | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 759 | |
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 760 | (******************************************************************************) | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 761 | (*************** definitions and theorems for match combinators ***************) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 762 | (******************************************************************************) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 763 | |
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 764 | fun add_match_combinators | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 765 | (bindings : binding list) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 766 | (spec : (term * (bool * typ) list) list) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 767 | (lhsT : typ) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 768 | (case_const : typ -> term) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 769 | (case_rews : thm list) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 770 | (thy : theory) = | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 771 | let | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 772 | |
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 773 | (* get a fresh type variable for the result type *) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 774 | val resultT : typ = | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 775 | let | 
| 40832 | 776 | val ts : string list = map fst (Term.add_tfreesT lhsT []) | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42290diff
changeset | 777 | val t : string = singleton (Name.variant_list ts) "'t" | 
| 40832 | 778 |       in TFree (t, @{sort pcpo}) end
 | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 779 | |
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 780 | (* define match combinators *) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 781 | local | 
| 40832 | 782 |       val x = Free ("x", lhsT)
 | 
| 783 |       fun k args = Free ("k", map snd args -->> mk_matchT resultT)
 | |
| 784 | val fail = mk_fail resultT | |
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 785 | fun mat_fun i (j, (_, args)) = | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 786 | let | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 787 | val (vs, _) = get_vars_avoiding ["x","k"] args | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 788 | in | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 789 | if i = j then k args else big_lambdas vs fail | 
| 40832 | 790 | end | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 791 | fun mat_eqn (i, (bind, (_, args))) : binding * term * mixfix = | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 792 | let | 
| 40832 | 793 | val mat_bind = Binding.prefix_name "match_" bind | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 794 | val funs = map_index (mat_fun i) spec | 
| 40832 | 795 | val body = list_ccomb (case_const (mk_matchT resultT), funs) | 
| 796 | val rhs = big_lambda x (big_lambda (k args) (body ` x)) | |
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 797 | in | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 798 | (mat_bind, rhs, NoSyn) | 
| 40832 | 799 | end | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 800 | in | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 801 | val ((match_consts, match_defs), thy) = | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 802 | define_consts (map_index mat_eqn (bindings ~~ spec)) thy | 
| 40832 | 803 | end | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 804 | |
| 35463 
b20501588930
register match functions from domain_constructors.ML
 huffman parents: 
35462diff
changeset | 805 | (* register match combinators with fixrec package *) | 
| 
b20501588930
register match functions from domain_constructors.ML
 huffman parents: 
35462diff
changeset | 806 | local | 
| 40832 | 807 | val con_names = map (fst o dest_Const o fst) spec | 
| 808 | val mat_names = map (fst o dest_Const) match_consts | |
| 35463 
b20501588930
register match functions from domain_constructors.ML
 huffman parents: 
35462diff
changeset | 809 | in | 
| 40832 | 810 | val thy = Fixrec.add_matchers (con_names ~~ mat_names) thy | 
| 811 | end | |
| 35463 
b20501588930
register match functions from domain_constructors.ML
 huffman parents: 
35462diff
changeset | 812 | |
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 813 | (* prove strictness of match combinators *) | 
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 814 | local | 
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 815 | fun match_strict mat = | 
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 816 | let | 
| 40832 | 817 | val (T, (U, V)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat)) | 
| 818 |           val k = Free ("k", U)
 | |
| 819 | val goal = mk_trp (mk_eq (mat ` mk_bottom T ` k, mk_bottom V)) | |
| 54895 | 820 | fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1] | 
| 821 | in prove thy match_defs goal (tacs o #context) end | |
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 822 | in | 
| 40832 | 823 | val match_stricts = map match_strict match_consts | 
| 824 | end | |
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 825 | |
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 826 | (* prove match/constructor rules *) | 
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 827 | local | 
| 40832 | 828 | val fail = mk_fail resultT | 
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 829 | fun match_app (i, mat) (j, (con, args)) = | 
| 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 830 | let | 
| 40832 | 831 | val (vs, nonlazy) = get_vars_avoiding ["k"] args | 
| 832 | val (_, (kT, _)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat)) | |
| 833 |           val k = Free ("k", kT)
 | |
| 834 | val lhs = mat ` list_ccomb (con, vs) ` k | |
| 835 | val rhs = if i = j then list_ccomb (k, vs) else fail | |
| 836 | val assms = map (mk_trp o mk_defined) nonlazy | |
| 837 | val concl = mk_trp (mk_eq (lhs, rhs)) | |
| 838 | val goal = Logic.list_implies (assms, concl) | |
| 54895 | 839 | fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1] | 
| 840 | in prove thy match_defs goal (tacs o #context) end | |
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 841 | fun one_match (i, mat) = | 
| 40832 | 842 | map_index (match_app (i, mat)) spec | 
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 843 | in | 
| 40832 | 844 | val match_apps = flat (map_index one_match match_consts) | 
| 845 | end | |
| 35466 
9fcfd5763181
move proofs of match_rews to domain_constructors.ML
 huffman parents: 
35463diff
changeset | 846 | |
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 847 | in | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 848 | (match_stricts @ match_apps, thy) | 
| 40832 | 849 | end | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 850 | |
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 851 | (******************************************************************************) | 
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 852 | (******************************* main function ********************************) | 
| 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 853 | (******************************************************************************) | 
| 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 854 | |
| 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 855 | fun add_domain_constructors | 
| 35777 
bcc77916b7b9
pass binding as argument to add_domain_constructors; proper binding for case combinator
 huffman parents: 
35561diff
changeset | 856 | (dbind : binding) | 
| 35481 
7bb9157507a9
add_domain_constructors takes iso_info record as argument
 huffman parents: 
35476diff
changeset | 857 | (spec : (binding * (bool * binding option * typ) list * mixfix) list) | 
| 35514 
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
 huffman parents: 
35496diff
changeset | 858 | (iso_info : Domain_Take_Proofs.iso_info) | 
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 859 | (thy : theory) = | 
| 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 860 | let | 
| 40832 | 861 | val dname = Binding.name_of dbind | 
| 862 |     val _ = writeln ("Proving isomorphism properties of domain "^dname^" ...")
 | |
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 863 | |
| 40832 | 864 | val bindings = map #1 spec | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 865 | |
| 35481 
7bb9157507a9
add_domain_constructors takes iso_info record as argument
 huffman parents: 
35476diff
changeset | 866 | (* retrieve facts about rep/abs *) | 
| 40832 | 867 | val lhsT = #absT iso_info | 
| 868 |     val {rep_const, abs_const, ...} = iso_info
 | |
| 869 | val abs_iso_thm = #abs_inverse iso_info | |
| 870 | val rep_iso_thm = #rep_inverse iso_info | |
| 871 |     val iso_locale = @{thm iso.intro} OF [abs_iso_thm, rep_iso_thm]
 | |
| 872 |     val rep_strict = iso_locale RS @{thm iso.rep_strict}
 | |
| 873 |     val abs_strict = iso_locale RS @{thm iso.abs_strict}
 | |
| 874 |     val rep_bottom_iff = iso_locale RS @{thm iso.rep_bottom_iff}
 | |
| 875 | val iso_rews = [abs_iso_thm, rep_iso_thm, abs_strict, rep_strict] | |
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 876 | |
| 35487 | 877 | (* qualify constants and theorems with domain name *) | 
| 40832 | 878 | val thy = Sign.add_path dname thy | 
| 35487 | 879 | |
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 880 | (* define constructor functions *) | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35456diff
changeset | 881 | val (con_result, thy) = | 
| 35454 
cf6ba350b7ca
remove unnecessary stuff from argument to add_constructors function
 huffman parents: 
35453diff
changeset | 882 | let | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 883 | fun prep_arg (lazy, _, T) = (lazy, T) | 
| 40832 | 884 | fun prep_con (b, args, mx) = (b, map prep_arg args, mx) | 
| 885 | val con_spec = map prep_con spec | |
| 35454 
cf6ba350b7ca
remove unnecessary stuff from argument to add_constructors function
 huffman parents: 
35453diff
changeset | 886 | in | 
| 
cf6ba350b7ca
remove unnecessary stuff from argument to add_constructors function
 huffman parents: 
35453diff
changeset | 887 | add_constructors con_spec abs_const iso_locale thy | 
| 40832 | 888 | end | 
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 889 |     val {con_consts, con_betas, nchotomy, exhaust, compacts, con_rews,
 | 
| 40832 | 890 | inverts, injects, dist_les, dist_eqs} = con_result | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 891 | |
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 892 | (* prepare constructor spec *) | 
| 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 893 | val con_specs : (term * (bool * typ) list) list = | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 894 | let | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 895 | fun prep_arg (lazy, _, T) = (lazy, T) | 
| 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 896 | fun prep_con c (_, args, _) = (c, map prep_arg args) | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 897 | in | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 898 | map2 prep_con con_consts spec | 
| 40832 | 899 | end | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 900 | |
| 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 901 | (* define case combinator *) | 
| 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 902 | val ((case_const : typ -> term, cases : thm list), thy) = | 
| 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 903 | add_case_combinator con_specs lhsT dbind | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 904 | con_betas iso_locale rep_const thy | 
| 35450 
e9ef2b50ac59
move constructor-specific stuff to a separate function
 huffman parents: 
35449diff
changeset | 905 | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 906 | (* define and prove theorems for selector functions *) | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 907 | val (sel_thms : thm list, thy : theory) = | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 908 | let | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 909 | val sel_spec : (term * (bool * binding option * typ) list) list = | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 910 | map2 (fn con => fn (_, args, _) => (con, args)) con_consts spec | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 911 | in | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 912 | add_selectors sel_spec rep_const | 
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40214diff
changeset | 913 | abs_iso_thm rep_strict rep_bottom_iff con_betas thy | 
| 40832 | 914 | end | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 915 | |
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 916 | (* define and prove theorems for discriminator functions *) | 
| 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 917 | val (dis_thms : thm list, thy : theory) = | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 918 | add_discriminators bindings con_specs lhsT | 
| 40832 | 919 | exhaust case_const cases thy | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 920 | |
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 921 | (* define and prove theorems for match combinators *) | 
| 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 922 | val (match_thms : thm list, thy : theory) = | 
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 923 | add_match_combinators bindings con_specs lhsT | 
| 44080 
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
 huffman parents: 
43324diff
changeset | 924 | case_const cases thy | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 925 | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35445diff
changeset | 926 | (* restore original signature path *) | 
| 40832 | 927 | val thy = Sign.parent_path thy | 
| 35444 | 928 | |
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 929 | (* bind theorem names in global theory *) | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 930 | val (_, thy) = | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 931 | let | 
| 63003 | 932 | fun qualified name = Binding.qualify_name true dbind name | 
| 40832 | 933 | val names = "bottom" :: map (fn (b,_,_) => Binding.name_of b) spec | 
| 934 | val dname = fst (dest_Type lhsT) | |
| 935 | val simp = Simplifier.simp_add | |
| 936 | val case_names = Rule_Cases.case_names names | |
| 937 | val cases_type = Induct.cases_type dname | |
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 938 | in | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 939 | Global_Theory.add_thmss [ | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 940 | ((qualified "iso_rews" , iso_rews ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 941 | ((qualified "nchotomy" , [nchotomy] ), []), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 942 | ((qualified "exhaust" , [exhaust] ), [case_names, cases_type]), | 
| 40213 
b63e966564da
rename case combinators generated by domain package to 'foo_case' instead of 'foo_when'
 huffman parents: 
40212diff
changeset | 943 | ((qualified "case_rews" , cases ), [simp]), | 
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 944 | ((qualified "compacts" , compacts ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 945 | ((qualified "con_rews" , con_rews ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 946 | ((qualified "sel_rews" , sel_thms ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 947 | ((qualified "dis_rews" , dis_thms ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 948 | ((qualified "dist_les" , dist_les ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 949 | ((qualified "dist_eqs" , dist_eqs ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 950 | ((qualified "inverts" , inverts ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 951 | ((qualified "injects" , injects ), [simp]), | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 952 | ((qualified "match_rews", match_thms ), [simp])] thy | 
| 40832 | 953 | end | 
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 954 | |
| 35444 | 955 | val result = | 
| 40017 | 956 |       {
 | 
| 957 | iso_info = iso_info, | |
| 40019 
05cda34d36e7
put constructor argument specs in constr_info type
 huffman parents: 
40017diff
changeset | 958 | con_specs = con_specs, | 
| 35451 | 959 | con_betas = con_betas, | 
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 960 | nchotomy = nchotomy, | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35777diff
changeset | 961 | exhaust = exhaust, | 
| 40016 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 962 | compacts = compacts, | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 963 | con_rews = con_rews, | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 964 | inverts = inverts, | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 965 | injects = injects, | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 966 | dist_les = dist_les, | 
| 
2eff1cbc1ccb
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
 huffman parents: 
40014diff
changeset | 967 | dist_eqs = dist_eqs, | 
| 35459 
3d8acfae6fb8
move proofs of when_rews intro domain_constructors.ML
 huffman parents: 
35458diff
changeset | 968 | cases = cases, | 
| 35460 
8cb42aa19358
move definition of discriminators to domain_constructors.ML
 huffman parents: 
35459diff
changeset | 969 | sel_rews = sel_thms, | 
| 35462 
f5461b02d754
move definition of match combinators to domain_constructors.ML
 huffman parents: 
35461diff
changeset | 970 | dis_rews = dis_thms, | 
| 40017 | 971 | match_rews = match_thms | 
| 40832 | 972 | } | 
| 35444 | 973 | in | 
| 974 | (result, thy) | |
| 40832 | 975 | end | 
| 35444 | 976 | |
| 40832 | 977 | end |