| author | blanchet | 
| Thu, 21 Apr 2011 22:32:00 +0200 | |
| changeset 42450 | 2765d4fb2b9c | 
| parent 42375 | 774df7c59508 | 
| child 43324 | 2b47822868e4 | 
| permissions | -rw-r--r-- | 
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 1 | (* Title: HOL/Tools/Datatype/datatype_abs_proofs.ML | 
| 11539 | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 5177 | 3 | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 4 | Datatype package: proofs and defintions independent of concrete | 
| 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 5 | representation of datatypes (i.e. requiring only abstract | 
| 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 6 | properties: injectivity / distinctness of constructors and induction). | 
| 5177 | 7 | *) | 
| 8 | ||
| 9 | signature DATATYPE_ABS_PROOFS = | |
| 10 | sig | |
| 31737 | 11 | include DATATYPE_COMMON | 
| 12 | val prove_casedist_thms : config -> string list -> | |
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 13 | descr list -> (string * sort) list -> thm -> | 
| 18728 | 14 | attribute list -> theory -> thm list * theory | 
| 31737 | 15 | val prove_primrec_thms : config -> string list -> | 
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 16 | descr list -> (string * sort) list -> | 
| 32915 
a7a97960054b
more appropriate abstraction over distinctness rules
 haftmann parents: 
32906diff
changeset | 17 | (string -> thm list) -> thm list list -> thm list list * thm list list -> | 
| 
a7a97960054b
more appropriate abstraction over distinctness rules
 haftmann parents: 
32906diff
changeset | 18 | thm -> theory -> (string list * thm list) * theory | 
| 31737 | 19 | val prove_case_thms : config -> string list -> | 
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 20 | descr list -> (string * sort) list -> | 
| 18314 | 21 | string list -> thm list -> theory -> (thm list list * string list) * theory | 
| 31737 | 22 | val prove_split_thms : config -> string list -> | 
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 23 | descr list -> (string * sort) list -> | 
| 5177 | 24 | thm list list -> thm list list -> thm list -> thm list list -> theory -> | 
| 18314 | 25 | (thm * thm) list * theory | 
| 31737 | 26 | val prove_nchotomys : config -> string list -> descr list -> | 
| 18314 | 27 | (string * sort) list -> thm list -> theory -> thm list * theory | 
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 28 | val prove_weak_case_congs : string list -> descr list -> | 
| 18314 | 29 | (string * sort) list -> theory -> thm list * theory | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 30 | val prove_case_congs : string list -> | 
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 31 | descr list -> (string * sort) list -> | 
| 18314 | 32 | thm list -> thm list list -> theory -> thm list * theory | 
| 5177 | 33 | end; | 
| 34 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 35 | structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS = | 
| 5177 | 36 | struct | 
| 37 | ||
| 38 | (************************ case distinction theorems ***************************) | |
| 39 | ||
| 41423 | 40 | fun prove_casedist_thms (config : Datatype_Aux.config) | 
| 41 | new_type_names descr sorts induct case_names_exhausts thy = | |
| 5177 | 42 | let | 
| 41423 | 43 | val _ = Datatype_Aux.message config "Proving case distinction theorems ..."; | 
| 5177 | 44 | |
| 32952 | 45 | val descr' = flat descr; | 
| 41423 | 46 | val recTs = Datatype_Aux.get_rec_types descr' sorts; | 
| 33957 | 47 | val newTs = take (length (hd descr)) recTs; | 
| 5177 | 48 | |
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 49 |     val {maxidx, ...} = rep_thm induct;
 | 
| 8305 | 50 | val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); | 
| 5177 | 51 | |
| 33063 | 52 | fun prove_casedist_thm (i, (T, t)) = | 
| 5177 | 53 | let | 
| 54 | val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) => | |
| 35364 | 55 |           Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
 | 
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 56 |         val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
 | 
| 5177 | 57 |           Var (("P", 0), HOLogic.boolT))
 | 
| 33957 | 58 | val insts = take i dummyPs @ (P::(drop (i + 1) dummyPs)); | 
| 17985 | 59 | val cert = cterm_of thy; | 
| 5177 | 60 | val insts' = (map cert induct_Ps) ~~ (map cert insts); | 
| 32905 | 61 | val induct' = refl RS ((nth | 
| 41423 | 62 | (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i) RSN (2, rev_mp)) | 
| 5177 | 63 | |
| 17985 | 64 | in | 
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 65 | Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) | 
| 26711 | 66 |           (fn {prems, ...} => EVERY
 | 
| 17985 | 67 | [rtac induct' 1, | 
| 68 | REPEAT (rtac TrueI 1), | |
| 69 | REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)), | |
| 20046 | 70 | REPEAT (rtac TrueI 1)]) | 
| 5177 | 71 | end; | 
| 72 | ||
| 33063 | 73 | val casedist_thms = map_index prove_casedist_thm | 
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 74 | (newTs ~~ Datatype_Prop.make_casedists descr sorts) | 
| 18314 | 75 | in | 
| 76 | thy | |
| 41423 | 77 | |> Datatype_Aux.store_thms_atts "exhaust" new_type_names | 
| 78 | (map single case_names_exhausts) casedist_thms | |
| 18314 | 79 | end; | 
| 5177 | 80 | |
| 81 | ||
| 82 | (*************************** primrec combinators ******************************) | |
| 83 | ||
| 41423 | 84 | fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr sorts | 
| 32915 
a7a97960054b
more appropriate abstraction over distinctness rules
 haftmann parents: 
32906diff
changeset | 85 | injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy = | 
| 5177 | 86 | let | 
| 41423 | 87 | val _ = Datatype_Aux.message config "Constructing primrec combinators ..."; | 
| 5661 | 88 | |
| 89 | val big_name = space_implode "_" new_type_names; | |
| 32124 | 90 | val thy0 = Sign.add_path big_name thy; | 
| 5177 | 91 | |
| 32952 | 92 | val descr' = flat descr; | 
| 41423 | 93 | val recTs = Datatype_Aux.get_rec_types descr' sorts; | 
| 30190 | 94 | val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; | 
| 33957 | 95 | val newTs = take (length (hd descr)) recTs; | 
| 5177 | 96 | |
| 8305 | 97 | val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); | 
| 5177 | 98 | |
| 5661 | 99 | val big_rec_name' = big_name ^ "_rec_set"; | 
| 21021 | 100 | val rec_set_names' = | 
| 101 | if length descr' = 1 then [big_rec_name'] else | |
| 102 | map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int) | |
| 103 | (1 upto (length descr')); | |
| 28965 | 104 | val rec_set_names = map (Sign.full_bname thy0) rec_set_names'; | 
| 5177 | 105 | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 106 | val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr sorts used; | 
| 5177 | 107 | |
| 21021 | 108 | val rec_set_Ts = map (fn (T1, T2) => | 
| 109 | reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); | |
| 5177 | 110 | |
| 41423 | 111 | val rec_fns = map (uncurry (Datatype_Aux.mk_Free "f")) | 
| 5177 | 112 | (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts))); | 
| 21021 | 113 | val rec_sets' = map (fn c => list_comb (Free c, rec_fns)) | 
| 114 | (rec_set_names' ~~ rec_set_Ts); | |
| 5177 | 115 | val rec_sets = map (fn c => list_comb (Const c, rec_fns)) | 
| 116 | (rec_set_names ~~ rec_set_Ts); | |
| 117 | ||
| 118 | (* introduction rules for graph of primrec function *) | |
| 119 | ||
| 32906 | 120 | fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) = | 
| 5177 | 121 | let | 
| 33338 | 122 | fun mk_prem (dt, U) (j, k, prems, t1s, t2s) = | 
| 41423 | 123 | let val free1 = Datatype_Aux.mk_Free "x" U j | 
| 124 | in (case (Datatype_Aux.strip_dtyp dt, strip_type U) of | |
| 125 | ((_, Datatype_Aux.DtRec m), (Us, _)) => | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 126 | let | 
| 41423 | 127 | val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k; | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 128 | val i = length Us | 
| 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 129 | in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all | 
| 32905 | 130 | (map (pair "x") Us, nth rec_sets' m $ | 
| 41423 | 131 | Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems, | 
| 5177 | 132 | free1::t1s, free2::t2s) | 
| 133 | end | |
| 134 | | _ => (j + 1, k, prems, free1::t1s, t2s)) | |
| 135 | end; | |
| 136 | ||
| 41423 | 137 | val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs; | 
| 33338 | 138 | val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []) | 
| 5177 | 139 | |
| 21021 | 140 | in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop | 
| 141 | (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $ | |
| 32905 | 142 | list_comb (nth rec_fns l, t1s @ t2s)))], l + 1) | 
| 5177 | 143 | end; | 
| 144 | ||
| 32906 | 145 | val (rec_intr_ts, _) = fold (fn ((d, T), set_name) => | 
| 146 | fold (make_rec_intr T set_name) (#3 (snd d))) | |
| 147 | (descr' ~~ recTs ~~ rec_sets') ([], 0); | |
| 5177 | 148 | |
| 21365 
4ee8e2702241
InductivePackage.add_inductive_i: canonical argument order;
 wenzelm parents: 
21291diff
changeset | 149 |     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
 | 
| 33278 | 150 | thy0 | 
| 151 | |> Sign.map_naming Name_Space.conceal | |
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33669diff
changeset | 152 | |> Inductive.add_inductive_global | 
| 33669 | 153 |           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
 | 
| 154 | coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false} | |
| 28965 | 155 | (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) | 
| 26128 | 156 | (map dest_Free rec_fns) | 
| 33278 | 157 | (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] | 
| 158 | ||> Sign.restore_naming thy0 | |
| 159 | ||> Theory.checkpoint; | |
| 5177 | 160 | |
| 161 | (* prove uniqueness and termination of primrec combinators *) | |
| 162 | ||
| 41423 | 163 | val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ..."; | 
| 5177 | 164 | |
| 32906 | 165 | fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) = | 
| 5177 | 166 | let | 
| 21021 | 167 | val distinct_tac = | 
| 5177 | 168 | (if i < length newTs then | 
| 32905 | 169 | full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1 | 
| 32915 
a7a97960054b
more appropriate abstraction over distinctness rules
 haftmann parents: 
32906diff
changeset | 170 | else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1); | 
| 5177 | 171 | |
| 172 | val inject = map (fn r => r RS iffD1) | |
| 32905 | 173 | (if i < length newTs then nth constr_inject i | 
| 32729 | 174 | else injects_of tname); | 
| 5177 | 175 | |
| 32906 | 176 | fun mk_unique_constr_tac n (cname, cargs) (tac, intr::intrs, j) = | 
| 5177 | 177 | let | 
| 41423 | 178 | val k = length (filter Datatype_Aux.is_rec_type cargs) | 
| 5177 | 179 | |
| 180 | in (EVERY [DETERM tac, | |
| 181 | REPEAT (etac ex1E 1), rtac ex1I 1, | |
| 182 | DEPTH_SOLVE_1 (ares_tac [intr] 1), | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 183 | REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1), | 
| 5177 | 184 | etac elim 1, | 
| 185 | REPEAT_DETERM_N j distinct_tac, | |
| 21021 | 186 | TRY (dresolve_tac inject 1), | 
| 5177 | 187 | REPEAT (etac conjE 1), hyp_subst_tac 1, | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12910diff
changeset | 188 | REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]), | 
| 5177 | 189 | TRY (hyp_subst_tac 1), | 
| 190 | rtac refl 1, | |
| 191 | REPEAT_DETERM_N (n - j - 1) distinct_tac], | |
| 192 | intrs, j + 1) | |
| 193 | end; | |
| 194 | ||
| 32906 | 195 | val (tac', intrs', _) = fold (mk_unique_constr_tac (length constrs)) | 
| 196 | constrs (tac, intrs, 0); | |
| 5177 | 197 | |
| 198 | in (tac', intrs') end; | |
| 199 | ||
| 200 | val rec_unique_thms = | |
| 201 | let | |
| 202 | val rec_unique_ts = map (fn (((set_t, T1), T2), i) => | |
| 35364 | 203 |           Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
 | 
| 41423 | 204 |             absfree ("y", T2, set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
 | 
| 21021 | 205 | (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs)); | 
| 17985 | 206 | val cert = cterm_of thy1 | 
| 5177 | 207 |         val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
 | 
| 208 | ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts); | |
| 209 | val induct' = cterm_instantiate ((map cert induct_Ps) ~~ | |
| 210 | (map cert insts)) induct; | |
| 32906 | 211 | val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts) | 
| 35625 | 212 | (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1 | 
| 35410 | 213 |               THEN rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs));
 | 
| 5177 | 214 | |
| 41423 | 215 | in | 
| 216 | Datatype_Aux.split_conj_thm (Skip_Proof.prove_global thy1 [] [] | |
| 217 | (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) (K tac)) | |
| 5177 | 218 | end; | 
| 219 | ||
| 35410 | 220 |     val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
 | 
| 5177 | 221 | |
| 222 | (* define primrec combinators *) | |
| 223 | ||
| 224 | val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec"; | |
| 28965 | 225 | val reccomb_names = map (Sign.full_bname thy1) | 
| 5177 | 226 | (if length descr' = 1 then [big_reccomb_name] else | 
| 227 | (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) | |
| 228 | (1 upto (length descr')))); | |
| 229 | val reccombs = map (fn ((name, T), T') => list_comb | |
| 230 | (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns)) | |
| 231 | (reccomb_names ~~ recTs ~~ rec_result_Ts); | |
| 232 | ||
| 18358 | 233 | val (reccomb_defs, thy2) = | 
| 234 | thy1 | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24699diff
changeset | 235 | |> Sign.add_consts_i (map (fn ((name, T), T') => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 236 | (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn)) | 
| 18358 | 237 | (reccomb_names ~~ recTs ~~ rec_result_Ts)) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 238 | |> (Global_Theory.add_defs false o map Thm.no_attributes) | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 239 | (map (fn ((((name, comb), set), T), T') => | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 240 | (Binding.name (Long_Name.base_name name ^ "_def"), | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 241 |               Logic.mk_equals (comb, absfree ("x", T,
 | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 242 |                Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
 | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 243 |                  set $ Free ("x", T) $ Free ("y", T'))))))
 | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 244 | (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)) | 
| 32124 | 245 | ||> Sign.parent_path | 
| 28361 
232fcbba2e4e
explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
 wenzelm parents: 
28110diff
changeset | 246 | ||> Theory.checkpoint; | 
| 5177 | 247 | |
| 248 | ||
| 249 | (* prove characteristic equations for primrec combinators *) | |
| 250 | ||
| 41423 | 251 | val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ..."; | 
| 5177 | 252 | |
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 253 | val rec_thms = map (fn t => Skip_Proof.prove_global thy2 [] [] t | 
| 17985 | 254 | (fn _ => EVERY | 
| 255 | [rewrite_goals_tac reccomb_defs, | |
| 35410 | 256 |          rtac @{thm the1_equality} 1,
 | 
| 5177 | 257 | resolve_tac rec_unique_thms 1, | 
| 258 | resolve_tac rec_intrs 1, | |
| 20046 | 259 | REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)])) | 
| 41423 | 260 | (Datatype_Prop.make_primrecs new_type_names descr sorts thy2); | 
| 5177 | 261 | in | 
| 18314 | 262 | thy2 | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24699diff
changeset | 263 | |> Sign.add_path (space_implode "_" new_type_names) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 264 | |> Global_Theory.add_thmss [((Binding.name "recs", rec_thms), [Nitpick_Simps.add])] | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24699diff
changeset | 265 | ||> Sign.parent_path | 
| 28361 
232fcbba2e4e
explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
 wenzelm parents: 
28110diff
changeset | 266 | ||> Theory.checkpoint | 
| 32906 | 267 | |-> (fn thms => pair (reccomb_names, flat thms)) | 
| 5177 | 268 | end; | 
| 269 | ||
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 270 | |
| 5177 | 271 | (***************************** case combinators *******************************) | 
| 272 | ||
| 41423 | 273 | fun prove_case_thms (config : Datatype_Aux.config) | 
| 274 | new_type_names descr sorts reccomb_names primrec_thms thy = | |
| 5177 | 275 | let | 
| 41423 | 276 | val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ..."; | 
| 5661 | 277 | |
| 32124 | 278 | val thy1 = Sign.add_path (space_implode "_" new_type_names) thy; | 
| 5177 | 279 | |
| 32952 | 280 | val descr' = flat descr; | 
| 41423 | 281 | val recTs = Datatype_Aux.get_rec_types descr' sorts; | 
| 30190 | 282 | val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; | 
| 33957 | 283 | val newTs = take (length (hd descr)) recTs; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20046diff
changeset | 284 | val T' = TFree (Name.variant used "'t", HOLogic.typeS); | 
| 5177 | 285 | |
| 41423 | 286 | fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' sorts dt) ---> T'; | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6522diff
changeset | 287 | |
| 5177 | 288 | val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) => | 
| 289 | let | |
| 41423 | 290 | val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs; | 
| 291 | val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs) | |
| 28524 | 292 |       in Const (@{const_name undefined}, Ts @ Ts' ---> T')
 | 
| 5177 | 293 | end) constrs) descr'; | 
| 294 | ||
| 28965 | 295 | val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names; | 
| 5177 | 296 | |
| 297 | (* define case combinators via primrec combinators *) | |
| 298 | ||
| 32906 | 299 | val (case_defs, thy2) = fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) => | 
| 5177 | 300 | let | 
| 32906 | 301 | val (fns1, fns2) = split_list (map (fn ((_, cargs), j) => | 
| 5177 | 302 | let | 
| 41423 | 303 | val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs; | 
| 304 | val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs); | |
| 305 | val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts'); | |
| 33957 | 306 | val frees = take (length cargs) frees'; | 
| 41423 | 307 | val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j | 
| 5177 | 308 | in | 
| 309 | (free, list_abs_free (map dest_Free frees', | |
| 310 | list_comb (free, frees))) | |
| 311 | end) (constrs ~~ (1 upto length constrs))); | |
| 312 | ||
| 313 | val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T'; | |
| 33957 | 314 | val fns = flat (take i case_dummy_fns) @ | 
| 315 | fns2 @ flat (drop (i + 1) case_dummy_fns); | |
| 5177 | 316 | val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T'); | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 317 | val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn); | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 318 | val def = (Binding.name (Long_Name.base_name name ^ "_def"), | 
| 5177 | 319 | Logic.mk_equals (list_comb (Const (name, caseT), fns1), | 
| 33957 | 320 | list_comb (reccomb, (flat (take i case_dummy_fns)) @ | 
| 321 | fns2 @ (flat (drop (i + 1) case_dummy_fns))))); | |
| 18358 | 322 | val ([def_thm], thy') = | 
| 323 | thy | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41423diff
changeset | 324 | |> Sign.declare_const_global decl |> snd | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38864diff
changeset | 325 | |> (Global_Theory.add_defs false o map Thm.no_attributes) [def]; | 
| 5177 | 326 | |
| 8436 | 327 | in (defs @ [def_thm], thy') | 
| 32906 | 328 | end) (hd descr ~~ newTs ~~ case_names ~~ | 
| 33957 | 329 | take (length newTs) reccomb_names) ([], thy1) | 
| 28361 
232fcbba2e4e
explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
 wenzelm parents: 
28110diff
changeset | 330 | ||> Theory.checkpoint; | 
| 5177 | 331 | |
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 332 | val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t | 
| 20046 | 333 | (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))) | 
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 334 | (Datatype_Prop.make_cases new_type_names descr sorts thy2) | 
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 335 | in | 
| 18314 | 336 | thy2 | 
| 33459 | 337 | |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms) | 
| 32124 | 338 | |> Sign.parent_path | 
| 41423 | 339 | |> Datatype_Aux.store_thmss "cases" new_type_names case_thms | 
| 18314 | 340 | |-> (fn thmss => pair (thmss, case_names)) | 
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 341 | end; | 
| 5177 | 342 | |
| 343 | ||
| 344 | (******************************* case splitting *******************************) | |
| 345 | ||
| 41423 | 346 | fun prove_split_thms (config : Datatype_Aux.config) | 
| 347 | new_type_names descr sorts constr_inject dist_rewrites casedist_thms case_thms thy = | |
| 5177 | 348 | let | 
| 41423 | 349 | val _ = Datatype_Aux.message config "Proving equations for case splitting ..."; | 
| 5177 | 350 | |
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
31604diff
changeset | 351 | val descr' = flat descr; | 
| 41423 | 352 | val recTs = Datatype_Aux.get_rec_types descr' sorts; | 
| 33957 | 353 | val newTs = take (length (hd descr)) recTs; | 
| 5177 | 354 | |
| 355 | fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), | |
| 356 | exhaustion), case_thms'), T) = | |
| 357 | let | |
| 17985 | 358 | val cert = cterm_of thy; | 
| 5177 | 359 | val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion))); | 
| 360 | val exhaustion' = cterm_instantiate | |
| 361 |           [(cert lhs, cert (Free ("x", T)))] exhaustion;
 | |
| 17985 | 362 | val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac | 
| 363 | (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))]) | |
| 5177 | 364 | in | 
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 365 | (Skip_Proof.prove_global thy [] [] t1 tacf, | 
| 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 366 | Skip_Proof.prove_global thy [] [] t2 tacf) | 
| 5177 | 367 | end; | 
| 368 | ||
| 369 | val split_thm_pairs = map prove_split_thms | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 370 | ((Datatype_Prop.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ | 
| 5177 | 371 | dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs); | 
| 372 | ||
| 373 | val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs | |
| 374 | ||
| 375 | in | |
| 18314 | 376 | thy | 
| 41423 | 377 | |> Datatype_Aux.store_thms "split" new_type_names split_thms | 
| 378 | ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms | |
| 18314 | 379 | |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2)) | 
| 5177 | 380 | end; | 
| 381 | ||
| 8601 | 382 | fun prove_weak_case_congs new_type_names descr sorts thy = | 
| 383 | let | |
| 384 | fun prove_weak_case_cong t = | |
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 385 | Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) | 
| 26711 | 386 |          (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
 | 
| 8601 | 387 | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 388 | val weak_case_congs = map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs | 
| 8601 | 389 | new_type_names descr sorts thy) | 
| 390 | ||
| 41423 | 391 | in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end; | 
| 8477 
17231d71171a
- Fixed bug in prove_casedist_thms (proof failed because of
 berghofe parents: 
8436diff
changeset | 392 | |
| 5177 | 393 | (************************* additional theorems for TFL ************************) | 
| 394 | ||
| 41423 | 395 | fun prove_nchotomys (config : Datatype_Aux.config) | 
| 396 | new_type_names descr sorts casedist_thms thy = | |
| 5177 | 397 | let | 
| 41423 | 398 | val _ = Datatype_Aux.message config "Proving additional theorems for TFL ..."; | 
| 5177 | 399 | |
| 400 | fun prove_nchotomy (t, exhaustion) = | |
| 401 | let | |
| 402 | (* For goal i, select the correct disjunct to attack, then prove it *) | |
| 403 | fun tac i 0 = EVERY [TRY (rtac disjI1 i), | |
| 404 | hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i] | |
| 405 | | tac i n = rtac disjI2 i THEN tac i (n - 1) | |
| 406 | in | |
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 407 | Skip_Proof.prove_global thy [] [] t (fn _ => | 
| 17985 | 408 | EVERY [rtac allI 1, | 
| 41423 | 409 | Datatype_Aux.exh_tac (K exhaustion) 1, | 
| 20046 | 410 | ALLGOALS (fn i => tac i (i-1))]) | 
| 5177 | 411 | end; | 
| 412 | ||
| 413 | val nchotomys = | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 414 | map prove_nchotomy (Datatype_Prop.make_nchotomys descr sorts ~~ casedist_thms) | 
| 5177 | 415 | |
| 41423 | 416 | in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end; | 
| 5177 | 417 | |
| 418 | fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy = | |
| 419 | let | |
| 420 | fun prove_case_cong ((t, nchotomy), case_rewrites) = | |
| 421 | let | |
| 422 |         val (Const ("==>", _) $ tm $ _) = t;
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
35625diff
changeset | 423 |         val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma)) = tm;
 | 
| 22578 | 424 | val cert = cterm_of thy; | 
| 5177 | 425 | val nchotomy' = nchotomy RS spec; | 
| 29264 | 426 | val [v] = Term.add_vars (concl_of nchotomy') []; | 
| 427 | val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy' | |
| 5177 | 428 | in | 
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32952diff
changeset | 429 | Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) | 
| 26711 | 430 |           (fn {prems, ...} => 
 | 
| 17985 | 431 | let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites)) | 
| 432 | in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1, | |
| 433 | cut_facts_tac [nchotomy''] 1, | |
| 434 | REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1), | |
| 435 | REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)] | |
| 20046 | 436 | end) | 
| 5177 | 437 | end; | 
| 438 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 439 | val case_congs = map prove_case_cong (Datatype_Prop.make_case_congs | 
| 5177 | 440 | new_type_names descr sorts thy ~~ nchotomys ~~ case_thms) | 
| 441 | ||
| 41423 | 442 | in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end; | 
| 443 | ||
| 444 | ||
| 445 | open Datatype_Aux; | |
| 5177 | 446 | |
| 447 | end; |