| author | blanchet | 
| Fri, 21 Feb 2014 00:09:55 +0100 | |
| changeset 55641 | 5b466efedd2c | 
| parent 54398 | 100c0eaf63d5 | 
| child 56254 | a2dd9200854d | 
| 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_prop.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: characteristic properties of datatypes. | 
| 5177 | 5 | *) | 
| 6 | ||
| 7 | signature DATATYPE_PROP = | |
| 8 | sig | |
| 45896 | 9 | type descr = Datatype_Aux.descr | 
| 8434 | 10 | val indexify_names: string list -> string list | 
| 13465 | 11 | val make_tnames: typ list -> string list | 
| 45822 | 12 | val make_injs : descr list -> term list list | 
| 45889 | 13 | val make_distincts : descr list -> term list list (*no symmetric inequalities*) | 
| 45822 | 14 | val make_ind : descr list -> term | 
| 15 | val make_casedists : descr list -> term list | |
| 16 | val make_primrec_Ts : descr list -> string list -> typ list * typ list | |
| 17 | val make_primrecs : string list -> descr list -> theory -> term list | |
| 18 | val make_cases : string list -> descr list -> theory -> term list list | |
| 19 | val make_splits : string list -> descr list -> theory -> (term * term) list | |
| 20 | val make_case_combs : string list -> descr list -> theory -> string -> term list | |
| 21 | val make_weak_case_congs : string list -> descr list -> theory -> term list | |
| 22 | val make_case_congs : string list -> descr list -> theory -> term list | |
| 23 | val make_nchotomys : descr list -> term list | |
| 5177 | 24 | end; | 
| 25 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 26 | structure Datatype_Prop : DATATYPE_PROP = | 
| 5177 | 27 | struct | 
| 28 | ||
| 45896 | 29 | type descr = Datatype_Aux.descr; | 
| 30 | ||
| 31 | ||
| 54398 
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
 blanchet parents: 
50239diff
changeset | 32 | val indexify_names = Case_Translation.indexify_names; | 
| 
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
 blanchet parents: 
50239diff
changeset | 33 | val make_tnames = Case_Translation.make_tnames; | 
| 8434 | 34 | |
| 5177 | 35 | fun make_tnames Ts = | 
| 36 | let | |
| 40720 | 37 | fun type_name (TFree (name, _)) = unprefix "'" name | 
| 38 | | type_name (Type (name, _)) = | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 39 | let val name' = Long_Name.base_name name | 
| 50239 | 40 | in if Symbol_Pos.is_identifier name' then name' else "x" end; | 
| 8434 | 41 | in indexify_names (map type_name Ts) end; | 
| 5177 | 42 | |
| 43 | ||
| 44 | (************************* injectivity of constructors ************************) | |
| 45 | ||
| 45822 | 46 | fun make_injs descr = | 
| 5177 | 47 | let | 
| 21078 | 48 | val descr' = flat descr; | 
| 49 | fun make_inj T (cname, cargs) = | |
| 45700 | 50 | if null cargs then I | 
| 51 | else | |
| 5177 | 52 | let | 
| 45822 | 53 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 5177 | 54 | val constr_t = Const (cname, Ts ---> T); | 
| 55 | val tnames = make_tnames Ts; | |
| 56 | val frees = map Free (tnames ~~ Ts); | |
| 45743 | 57 | val frees' = map Free (map (suffix "'") tnames ~~ Ts); | 
| 58 | in | |
| 59 | cons (HOLogic.mk_Trueprop (HOLogic.mk_eq | |
| 60 | (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')), | |
| 61 |              foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
 | |
| 62 | (map HOLogic.mk_eq (frees ~~ frees'))))) | |
| 5177 | 63 | end; | 
| 21078 | 64 | in | 
| 65 | map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) []) | |
| 45822 | 66 | (hd descr) (take (length (hd descr)) (Datatype_Aux.get_rec_types descr')) | 
| 5177 | 67 | end; | 
| 68 | ||
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 69 | |
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 70 | (************************* distinctness of constructors ***********************) | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 71 | |
| 45822 | 72 | fun make_distincts descr = | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 73 | let | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 74 | val descr' = flat descr; | 
| 45822 | 75 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 33957 | 76 | val newTs = take (length (hd descr)) recTs; | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 77 | |
| 45822 | 78 | fun prep_constr (cname, cargs) = (cname, map (Datatype_Aux.typ_of_dtyp descr') cargs); | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 79 | |
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 80 | fun make_distincts' _ [] = [] | 
| 45700 | 81 | | make_distincts' T ((cname, cargs) :: constrs) = | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 82 | let | 
| 45743 | 83 | val frees = map Free (make_tnames cargs ~~ cargs); | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 84 | val t = list_comb (Const (cname, cargs ---> T), frees); | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 85 | |
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 86 | fun make_distincts'' (cname', cargs') = | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 87 | let | 
| 45700 | 88 | val frees' = map Free (map (suffix "'") (make_tnames cargs') ~~ cargs'); | 
| 89 | val t' = list_comb (Const (cname', cargs' ---> T), frees'); | |
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 90 | in | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 91 | HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.mk_eq (t, t')) | 
| 45700 | 92 | end; | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 93 | in map make_distincts'' constrs @ make_distincts' T constrs end; | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 94 | in | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 95 | map2 (fn ((_, (_, _, constrs))) => fn T => | 
| 45889 | 96 | make_distincts' T (map prep_constr constrs)) (hd descr) newTs | 
| 27300 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 97 | end; | 
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 98 | |
| 
4cb3101d2bf7
DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
 haftmann parents: 
27002diff
changeset | 99 | |
| 5177 | 100 | (********************************* induction **********************************) | 
| 101 | ||
| 45822 | 102 | fun make_ind descr = | 
| 5177 | 103 | let | 
| 32952 | 104 | val descr' = flat descr; | 
| 45822 | 105 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 41423 | 106 | val pnames = | 
| 107 | if length descr' = 1 then ["P"] | |
| 5177 | 108 | else map (fn i => "P" ^ string_of_int i) (1 upto length descr'); | 
| 109 | ||
| 110 | fun make_pred i T = | |
| 111 | let val T' = T --> HOLogic.boolT | |
| 42364 | 112 | in Free (nth pnames i, T') end; | 
| 5177 | 113 | |
| 114 | fun make_ind_prem k T (cname, cargs) = | |
| 115 | let | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 116 | fun mk_prem ((dt, s), T) = | 
| 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 117 | let val (Us, U) = strip_type T | 
| 41423 | 118 | in | 
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
46215diff
changeset | 119 | Logic.list_all (map (pair "x") Us, | 
| 41423 | 120 | HOLogic.mk_Trueprop | 
| 121 | (make_pred (Datatype_Aux.body_index dt) U $ | |
| 122 | Datatype_Aux.app_bnds (Free (s, T)) (length Us))) | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 123 | end; | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 124 | |
| 41423 | 125 | val recs = filter Datatype_Aux.is_rec_type cargs; | 
| 45822 | 126 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 127 | val recTs' = map (Datatype_Aux.typ_of_dtyp descr') recs; | |
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19233diff
changeset | 128 | val tnames = Name.variant_list pnames (make_tnames Ts); | 
| 41423 | 129 | val rec_tnames = map fst (filter (Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs)); | 
| 5177 | 130 | val frees = tnames ~~ Ts; | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 131 | val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); | 
| 45700 | 132 | in | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 133 | fold_rev (Logic.all o Free) frees | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 134 | (Logic.list_implies (prems, | 
| 45700 | 135 | HOLogic.mk_Trueprop (make_pred k T $ | 
| 136 | list_comb (Const (cname, Ts ---> T), map Free frees)))) | |
| 5177 | 137 | end; | 
| 138 | ||
| 45700 | 139 | val prems = | 
| 140 | maps (fn ((i, (_, _, constrs)), T) => map (make_ind_prem i T) constrs) (descr' ~~ recTs); | |
| 5177 | 141 | val tnames = make_tnames recTs; | 
| 45700 | 142 | val concl = | 
| 143 |       HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
 | |
| 144 | (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T)) | |
| 145 | (descr' ~~ recTs ~~ tnames))); | |
| 5177 | 146 | |
| 147 | in Logic.list_implies (prems, concl) end; | |
| 148 | ||
| 149 | (******************************* case distinction *****************************) | |
| 150 | ||
| 45822 | 151 | fun make_casedists descr = | 
| 5177 | 152 | let | 
| 32952 | 153 | val descr' = flat descr; | 
| 5177 | 154 | |
| 155 | fun make_casedist_prem T (cname, cargs) = | |
| 156 | let | |
| 45822 | 157 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19233diff
changeset | 158 | val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts; | 
| 45700 | 159 | val free_ts = map Free frees; | 
| 160 | in | |
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 161 | fold_rev (Logic.all o Free) frees | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 162 | (Logic.mk_implies (HOLogic.mk_Trueprop | 
| 45700 | 163 |             (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
 | 
| 164 |               HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))))
 | |
| 5177 | 165 | end; | 
| 166 | ||
| 33957 | 167 | fun make_casedist ((_, (_, _, constrs))) T = | 
| 5177 | 168 | let val prems = map (make_casedist_prem T) constrs | 
| 45700 | 169 |       in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))) end;
 | 
| 5177 | 170 | |
| 41423 | 171 | in | 
| 172 | map2 make_casedist (hd descr) | |
| 45822 | 173 | (take (length (hd descr)) (Datatype_Aux.get_rec_types descr')) | 
| 41423 | 174 | end; | 
| 5177 | 175 | |
| 176 | (*************** characteristic equations for primrec combinator **************) | |
| 177 | ||
| 45822 | 178 | fun make_primrec_Ts descr used = | 
| 5177 | 179 | let | 
| 32952 | 180 | val descr' = flat descr; | 
| 5177 | 181 | |
| 45700 | 182 | val rec_result_Ts = | 
| 183 | map TFree | |
| 184 | (Name.variant_list used (replicate (length descr') "'t") ~~ | |
| 185 | replicate (length descr') HOLogic.typeS); | |
| 5177 | 186 | |
| 32952 | 187 | val reccomb_fn_Ts = maps (fn (i, (_, _, constrs)) => | 
| 5177 | 188 | map (fn (_, cargs) => | 
| 189 | let | |
| 45822 | 190 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 41423 | 191 | val recs = filter (Datatype_Aux.is_rec_type o fst) (cargs ~~ Ts); | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 192 | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 193 | fun mk_argT (dt, T) = | 
| 42364 | 194 | binder_types T ---> nth rec_result_Ts (Datatype_Aux.body_index dt); | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 195 | |
| 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 196 | val argTs = Ts @ map mk_argT recs | 
| 42364 | 197 | in argTs ---> nth rec_result_Ts i end) constrs) descr'; | 
| 5177 | 198 | |
| 15459 
16dd63c78049
Introduced function DatatypeProp.make_primrec_Ts to avoid code duplication.
 berghofe parents: 
14981diff
changeset | 199 | in (rec_result_Ts, reccomb_fn_Ts) end; | 
| 
16dd63c78049
Introduced function DatatypeProp.make_primrec_Ts to avoid code duplication.
 berghofe parents: 
14981diff
changeset | 200 | |
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 201 | fun make_primrecs reccomb_names descr thy = | 
| 15459 
16dd63c78049
Introduced function DatatypeProp.make_primrec_Ts to avoid code duplication.
 berghofe parents: 
14981diff
changeset | 202 | let | 
| 32952 | 203 | val descr' = flat descr; | 
| 45822 | 204 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 45738 | 205 | val used = fold Term.add_tfree_namesT recTs []; | 
| 15459 
16dd63c78049
Introduced function DatatypeProp.make_primrec_Ts to avoid code duplication.
 berghofe parents: 
14981diff
changeset | 206 | |
| 45822 | 207 | val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr used; | 
| 15459 
16dd63c78049
Introduced function DatatypeProp.make_primrec_Ts to avoid code duplication.
 berghofe parents: 
14981diff
changeset | 208 | |
| 45700 | 209 | val rec_fns = | 
| 210 | map (uncurry (Datatype_Aux.mk_Free "f")) | |
| 211 | (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts))); | |
| 5177 | 212 | |
| 45700 | 213 | val reccombs = | 
| 214 | map (fn ((name, T), T') => list_comb (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns)) | |
| 5177 | 215 | (reccomb_names ~~ recTs ~~ rec_result_Ts); | 
| 216 | ||
| 45700 | 217 | fun make_primrec T comb_t (cname, cargs) (ts, f :: fs) = | 
| 5177 | 218 | let | 
| 41423 | 219 | val recs = filter Datatype_Aux.is_rec_type cargs; | 
| 45822 | 220 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 221 | val recTs' = map (Datatype_Aux.typ_of_dtyp descr') recs; | |
| 5177 | 222 | val tnames = make_tnames Ts; | 
| 41423 | 223 | val rec_tnames = map fst (filter (Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs)); | 
| 5177 | 224 | val frees = map Free (tnames ~~ Ts); | 
| 225 | val frees' = map Free (rec_tnames ~~ recTs'); | |
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 226 | |
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 227 | fun mk_reccomb ((dt, T), t) = | 
| 42364 | 228 | let val (Us, U) = strip_type T in | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
46218diff
changeset | 229 | fold_rev (Term.abs o pair "x") Us | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
46218diff
changeset | 230 | (nth reccombs (Datatype_Aux.body_index dt) $ Datatype_Aux.app_bnds t (length Us)) | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
13585diff
changeset | 231 | end; | 
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 232 | |
| 45700 | 233 | val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees'); | 
| 5177 | 234 | |
| 45700 | 235 | in | 
| 236 | (ts @ [HOLogic.mk_Trueprop | |
| 237 | (HOLogic.mk_eq (comb_t $ list_comb (Const (cname, Ts ---> T), frees), | |
| 238 | list_comb (f, frees @ reccombs')))], fs) | |
| 33244 | 239 | end; | 
| 240 | in | |
| 241 | fold (fn ((dt, T), comb_t) => fold (make_primrec T comb_t) (#3 (snd dt))) | |
| 242 | (descr' ~~ recTs ~~ reccombs) ([], rec_fns) | |
| 243 | |> fst | |
| 5177 | 244 | end; | 
| 245 | ||
| 246 | (****************** make terms of form t_case f1 ... fn *********************) | |
| 247 | ||
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 248 | fun make_case_combs case_names descr thy fname = | 
| 5177 | 249 | let | 
| 32952 | 250 | val descr' = flat descr; | 
| 45822 | 251 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 45738 | 252 | val used = fold Term.add_tfree_namesT recTs []; | 
| 33957 | 253 | val newTs = take (length (hd descr)) recTs; | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42364diff
changeset | 254 | val T' = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS); | 
| 5177 | 255 | |
| 256 | val case_fn_Ts = map (fn (i, (_, _, constrs)) => | |
| 257 | map (fn (_, cargs) => | |
| 45822 | 258 | let val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs | 
| 5177 | 259 | in Ts ---> T' end) constrs) (hd descr); | 
| 260 | in | |
| 261 | map (fn ((name, Ts), T) => list_comb | |
| 262 | (Const (name, Ts @ [T] ---> T'), | |
| 41423 | 263 | map (uncurry (Datatype_Aux.mk_Free fname)) (Ts ~~ (1 upto length Ts)))) | 
| 5177 | 264 | (case_names ~~ case_fn_Ts ~~ newTs) | 
| 265 | end; | |
| 266 | ||
| 267 | (**************** characteristic equations for case combinator ****************) | |
| 268 | ||
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 269 | fun make_cases case_names descr thy = | 
| 5177 | 270 | let | 
| 32952 | 271 | val descr' = flat descr; | 
| 45822 | 272 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 33957 | 273 | val newTs = take (length (hd descr)) recTs; | 
| 5177 | 274 | |
| 275 | fun make_case T comb_t ((cname, cargs), f) = | |
| 276 | let | |
| 45822 | 277 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 45700 | 278 | val frees = map Free ((make_tnames Ts) ~~ Ts); | 
| 279 | in | |
| 280 | HOLogic.mk_Trueprop | |
| 281 | (HOLogic.mk_eq (comb_t $ list_comb (Const (cname, Ts ---> T), frees), | |
| 282 | list_comb (f, frees))) | |
| 283 | end; | |
| 284 | in | |
| 285 | map (fn (((_, (_, _, constrs)), T), comb_t) => | |
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 286 | map (make_case T comb_t) (constrs ~~ snd (strip_comb comb_t))) | 
| 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 287 | (hd descr ~~ newTs ~~ make_case_combs case_names descr thy "f") | 
| 5177 | 288 | end; | 
| 289 | ||
| 7015 
85be09eb136c
- Datatype package now also supports arbitrarily branching datatypes
 berghofe parents: 
6394diff
changeset | 290 | |
| 5177 | 291 | (*************************** the "split" - equations **************************) | 
| 292 | ||
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 293 | fun make_splits case_names descr thy = | 
| 5177 | 294 | let | 
| 32952 | 295 | val descr' = flat descr; | 
| 45822 | 296 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 45738 | 297 | val used' = fold Term.add_tfree_namesT recTs []; | 
| 33957 | 298 | val newTs = take (length (hd descr)) recTs; | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42364diff
changeset | 299 | val T' = TFree (singleton (Name.variant_list used') "'t", HOLogic.typeS); | 
| 5177 | 300 |     val P = Free ("P", T' --> HOLogic.boolT);
 | 
| 301 | ||
| 302 | fun make_split (((_, (_, _, constrs)), T), comb_t) = | |
| 303 | let | |
| 304 | val (_, fs) = strip_comb comb_t; | |
| 45700 | 305 | val used = ["P", "x"] @ map (fst o dest_Free) fs; | 
| 5177 | 306 | |
| 33338 | 307 | fun process_constr ((cname, cargs), f) (t1s, t2s) = | 
| 5177 | 308 | let | 
| 45822 | 309 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19233diff
changeset | 310 | val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts); | 
| 45700 | 311 |             val eqn = HOLogic.mk_eq (Free ("x", T), list_comb (Const (cname, Ts ---> T), frees));
 | 
| 312 | val P' = P $ list_comb (f, frees); | |
| 313 | in | |
| 314 | (fold_rev (fn Free (s, T) => fn t => HOLogic.mk_all (s, T, t)) frees | |
| 315 | (HOLogic.imp $ eqn $ P') :: t1s, | |
| 316 | fold_rev (fn Free (s, T) => fn t => HOLogic.mk_exists (s, T, t)) frees | |
| 317 | (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) :: t2s) | |
| 5177 | 318 | end; | 
| 319 | ||
| 33338 | 320 | val (t1s, t2s) = fold_rev process_constr (constrs ~~ fs) ([], []); | 
| 45700 | 321 |         val lhs = P $ (comb_t $ Free ("x", T));
 | 
| 5177 | 322 | in | 
| 41423 | 323 | (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, Datatype_Aux.mk_conj t1s)), | 
| 324 | HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ Datatype_Aux.mk_disj t2s))) | |
| 5177 | 325 | end | 
| 326 | ||
| 45700 | 327 | in | 
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 328 | map make_split (hd descr ~~ newTs ~~ make_case_combs case_names descr thy "f") | 
| 5177 | 329 | end; | 
| 330 | ||
| 331 | (************************* additional rules for TFL ***************************) | |
| 332 | ||
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 333 | fun make_weak_case_congs case_names descr thy = | 
| 8601 | 334 | let | 
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 335 | val case_combs = make_case_combs case_names descr thy "f"; | 
| 8601 | 336 | |
| 337 | fun mk_case_cong comb = | |
| 45700 | 338 | let | 
| 8601 | 339 |         val Type ("fun", [T, _]) = fastype_of comb;
 | 
| 340 |         val M = Free ("M", T);
 | |
| 341 |         val M' = Free ("M'", T);
 | |
| 342 | in | |
| 343 | Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')), | |
| 344 | HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb $ M'))) | |
| 45700 | 345 | end; | 
| 8601 | 346 | in | 
| 347 | map mk_case_cong case_combs | |
| 348 | end; | |
| 45700 | 349 | |
| 8601 | 350 | |
| 5177 | 351 | (*--------------------------------------------------------------------------- | 
| 352 | * Structure of case congruence theorem looks like this: | |
| 353 | * | |
| 45700 | 354 | * (M = M') | 
| 355 | * ==> (!!x1,...,xk. (M' = C1 x1..xk) ==> (f1 x1..xk = g1 x1..xk)) | |
| 356 | * ==> ... | |
| 357 | * ==> (!!x1,...,xj. (M' = Cn x1..xj) ==> (fn x1..xj = gn x1..xj)) | |
| 5177 | 358 | * ==> | 
| 359 | * (ty_case f1..fn M = ty_case g1..gn M') | |
| 360 | *---------------------------------------------------------------------------*) | |
| 361 | ||
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 362 | fun make_case_congs case_names descr thy = | 
| 5177 | 363 | let | 
| 45879 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 364 | val case_combs = make_case_combs case_names descr thy "f"; | 
| 
71b8d0d170b1
avoid fragile Sign.intern_const -- pass internal names directly;
 wenzelm parents: 
45822diff
changeset | 365 | val case_combs' = make_case_combs case_names descr thy "g"; | 
| 5177 | 366 | |
| 367 | fun mk_case_cong ((comb, comb'), (_, (_, _, constrs))) = | |
| 368 | let | |
| 369 |         val Type ("fun", [T, _]) = fastype_of comb;
 | |
| 370 | val (_, fs) = strip_comb comb; | |
| 371 | val (_, gs) = strip_comb comb'; | |
| 372 | val used = ["M", "M'"] @ map (fst o dest_Free) (fs @ gs); | |
| 373 |         val M = Free ("M", T);
 | |
| 374 |         val M' = Free ("M'", T);
 | |
| 375 | ||
| 376 | fun mk_clause ((f, g), (cname, _)) = | |
| 377 | let | |
| 40844 | 378 | val Ts = binder_types (fastype_of f); | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19233diff
changeset | 379 | val tnames = Name.variant_list used (make_tnames Ts); | 
| 45743 | 380 | val frees = map Free (tnames ~~ Ts); | 
| 5177 | 381 | in | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 382 | fold_rev Logic.all frees | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 383 | (Logic.mk_implies | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 384 | (HOLogic.mk_Trueprop | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 385 | (HOLogic.mk_eq (M', list_comb (Const (cname, Ts ---> T), frees))), | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 386 | HOLogic.mk_Trueprop | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
45896diff
changeset | 387 | (HOLogic.mk_eq (list_comb (f, frees), list_comb (g, frees))))) | 
| 45700 | 388 | end; | 
| 5177 | 389 | in | 
| 390 | Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')) :: | |
| 391 | map mk_clause (fs ~~ gs ~~ constrs), | |
| 392 | HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb' $ M'))) | |
| 45700 | 393 | end; | 
| 5177 | 394 | in | 
| 395 | map mk_case_cong (case_combs ~~ case_combs' ~~ hd descr) | |
| 396 | end; | |
| 397 | ||
| 398 | (*--------------------------------------------------------------------------- | |
| 399 | * Structure of exhaustion theorem looks like this: | |
| 400 | * | |
| 401 | * !v. (? y1..yi. v = C1 y1..yi) | ... | (? y1..yj. v = Cn y1..yj) | |
| 402 | *---------------------------------------------------------------------------*) | |
| 403 | ||
| 45822 | 404 | fun make_nchotomys descr = | 
| 5177 | 405 | let | 
| 32952 | 406 | val descr' = flat descr; | 
| 45822 | 407 | val recTs = Datatype_Aux.get_rec_types descr'; | 
| 33957 | 408 | val newTs = take (length (hd descr)) recTs; | 
| 5177 | 409 | |
| 410 | fun mk_eqn T (cname, cargs) = | |
| 411 | let | |
| 45822 | 412 | val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19233diff
changeset | 413 | val tnames = Name.variant_list ["v"] (make_tnames Ts); | 
| 45743 | 414 | val frees = tnames ~~ Ts; | 
| 5177 | 415 | in | 
| 33338 | 416 | fold_rev (fn (s, T') => fn t => HOLogic.mk_exists (s, T', t)) frees | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 417 |           (HOLogic.mk_eq (Free ("v", T),
 | 
| 33338 | 418 | list_comb (Const (cname, Ts ---> T), map Free frees))) | 
| 45700 | 419 | end; | 
| 420 | in | |
| 421 | map (fn ((_, (_, _, constrs)), T) => | |
| 422 | HOLogic.mk_Trueprop | |
| 423 |           (HOLogic.mk_all ("v", T, Datatype_Aux.mk_disj (map (mk_eqn T) constrs))))
 | |
| 5177 | 424 | (hd descr ~~ newTs) | 
| 425 | end; | |
| 426 | ||
| 427 | end; |