| author | paulson | 
| Fri, 02 Oct 1998 10:44:20 +0200 | |
| changeset 5607 | 5db9e2343ade | 
| parent 5290 | b755c7240348 | 
| child 5698 | 2b5d9bdec5af | 
| permissions | -rw-r--r-- | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Tools/record_package.ML | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 3 | Author: Wolfgang Naraschewski and Markus Wenzel, TU Muenchen | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 4 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 5 | Extensible records with structural subtyping in HOL. | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 6 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 7 | TODO: | 
| 5235 | 8 | - field types: datatype; | 
| 4894 | 9 | - operations and theorems: split, split_all/ex, ...; | 
| 5197 | 10 | - field constructor: more specific type for snd component (x_more etc. classes); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 11 | *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 12 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 13 | signature RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 14 | sig | 
| 4890 | 15 | val moreS: sort | 
| 16 | val mk_fieldT: (string * typ) * typ -> typ | |
| 17 | val dest_fieldT: typ -> (string * typ) * typ | |
| 18 | val mk_field: (string * term) * term -> term | |
| 19 | val mk_fst: term -> term | |
| 20 | val mk_snd: term -> term | |
| 21 | val mk_recordT: (string * typ) list * typ -> typ | |
| 22 | val dest_recordT: typ -> (string * typ) list * typ | |
| 23 | val mk_record: (string * term) list * term -> term | |
| 24 | val mk_sel: term -> string -> term | |
| 25 | val mk_update: term -> string * term -> term | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 26 | val print_records: theory -> unit | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 27 | val add_record: (string list * bstring) -> string option | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 28 | -> (bstring * string) list -> theory -> theory | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 29 | val add_record_i: (string list * bstring) -> (typ list * string) option | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 30 | -> (bstring * typ) list -> theory -> theory | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 31 | val setup: (theory -> theory) list | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 32 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 33 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 34 | structure RecordPackage: RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 35 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 36 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 37 | |
| 4894 | 38 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 39 | |
| 4894 | 40 | (* definitions and equations *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 41 | |
| 4894 | 42 | infix 0 :== === ; | 
| 43 | ||
| 44 | val (op :==) = Logic.mk_defpair; | |
| 45 | val (op ===) = HOLogic.mk_Trueprop o HOLogic.mk_eq; | |
| 46 | ||
| 47 | fun get_defs thy specs = map (PureThy.get_tthm thy o fst) specs; | |
| 48 | ||
| 49 | ||
| 50 | (* proof by simplification *) | |
| 4890 | 51 | |
| 4895 | 52 | fun prove_simp thy simps = | 
| 4894 | 53 | let | 
| 4895 | 54 | val sign = Theory.sign_of thy; | 
| 55 | val ss = Simplifier.addsimps (HOL_basic_ss, map Attribute.thm_of simps); | |
| 56 | ||
| 57 | fun prove goal = | |
| 4894 | 58 | Attribute.tthm_of | 
| 4895 | 59 | (Goals.prove_goalw_cterm [] (Thm.cterm_of sign goal) | 
| 4894 | 60 | (K [ALLGOALS (Simplifier.simp_tac ss)]) | 
| 61 |         handle ERROR => error ("The error(s) above occurred while trying to prove "
 | |
| 4895 | 62 | ^ quote (Sign.string_of_term sign goal))); | 
| 4894 | 63 | in prove end; | 
| 64 | ||
| 65 | ||
| 66 | ||
| 67 | (*** syntax operations ***) | |
| 68 | ||
| 69 | (** name components **) | |
| 70 | ||
| 71 | val moreN = "more"; | |
| 72 | val schemeN = "_scheme"; | |
| 73 | val fieldN = "_field"; | |
| 74 | val field_typeN = "_field_type"; | |
| 75 | val fstN = "_fst"; | |
| 76 | val sndN = "_snd"; | |
| 77 | val updateN = "_update"; | |
| 78 | val makeN = "make"; | |
| 79 | val make_schemeN = "make_scheme"; | |
| 4890 | 80 | |
| 81 | ||
| 82 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 83 | (** tuple operations **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 84 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 85 | (* more type class *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 86 | |
| 5210 | 87 | val moreS = ["Record.more"]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 88 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 89 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 90 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 91 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 92 | fun mk_fieldT ((c, T), U) = Type (suffix field_typeN c, [T, U]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 93 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 94 | fun dest_fieldT (typ as Type (c_field_type, [T, U])) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 95 | (case try (unsuffix field_typeN) c_field_type of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 96 |         None => raise TYPE ("dest_fieldT", [typ], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 97 | | Some c => ((c, T), U)) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 98 |   | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 99 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 100 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 101 | (* constructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 102 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 103 | fun mk_fieldC U (c, T) = (suffix fieldN c, T --> U --> mk_fieldT ((c, T), U)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 104 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 105 | fun mk_field ((c, t), u) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 106 | let val T = fastype_of t and U = fastype_of u | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 107 | in Const (suffix fieldN c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 108 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 109 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 110 | (* destructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 111 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 112 | fun mk_fstC U (c, T) = (suffix fstN c, mk_fieldT ((c, T), U) --> T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 113 | fun mk_sndC U (c, T) = (suffix sndN c, mk_fieldT ((c, T), U) --> U); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 114 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 115 | fun dest_field fst_or_snd p = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 116 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 117 | val pT = fastype_of p; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 118 | val ((c, T), U) = dest_fieldT pT; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 119 | val (destN, destT) = if fst_or_snd then (fstN, T) else (sndN, U); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 120 | in Const (suffix destN c, pT --> destT) $ p end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 121 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 122 | val mk_fst = dest_field true; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 123 | val mk_snd = dest_field false; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 124 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 125 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 126 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 127 | (** record operations **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 128 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 129 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 130 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 131 | val mk_recordT = foldr mk_fieldT; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 132 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 133 | fun dest_recordT T = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 134 | (case try dest_fieldT T of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 135 | None => ([], T) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 136 | | Some (c_T, U) => apfst (cons c_T) (dest_recordT U)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 137 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 138 | fun find_fieldT c rT = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 139 | (case assoc (fst (dest_recordT rT), c) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 140 |     None => raise TYPE ("find_field: " ^ c, [rT], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 141 | | Some T => T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 142 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 143 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 144 | (* constructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 145 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 146 | val mk_record = foldr mk_field; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 147 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 148 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 149 | (* selectors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 150 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 151 | fun mk_selC rT (c, T) = (c, rT --> T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 152 | |
| 4890 | 153 | fun mk_sel r c = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 154 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 155 | in Const (mk_selC rT (c, find_fieldT c rT)) $ r end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 156 | |
| 4894 | 157 | val mk_moreC = mk_selC; | 
| 158 | ||
| 159 | fun mk_more r c = | |
| 160 | let val rT = fastype_of r | |
| 161 | in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end; | |
| 162 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 163 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 164 | (* updates *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 165 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 166 | fun mk_updateC rT (c, T) = (suffix updateN c, T --> rT --> rT); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 167 | |
| 4890 | 168 | fun mk_update r (c, x) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 169 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 170 | in Const (mk_updateC rT (c, find_fieldT c rT)) $ x $ r end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 171 | |
| 5197 | 172 | val mk_more_updateC = mk_updateC; | 
| 173 | ||
| 174 | fun mk_more_update r (c, x) = | |
| 175 | let val rT = fastype_of r | |
| 176 | in Const (mk_more_updateC rT (c, snd (dest_recordT rT))) $ x $ r end; | |
| 177 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 178 | |
| 4890 | 179 | (* make *) | 
| 180 | ||
| 181 | fun mk_makeC rT (c, Ts) = (c, Ts ---> rT); | |
| 182 | ||
| 183 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 184 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 185 | (** concrete syntax for records **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 186 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 187 | (* parse translations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 188 | |
| 5197 | 189 | fun gen_field_tr mark sfx (t as Const (c, _) $ Free (name, _) $ arg) = | 
| 190 | if c = mark then Syntax.const (suffix sfx name) $ arg | |
| 191 |       else raise TERM ("gen_field_tr: " ^ mark, [t])
 | |
| 192 |   | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 193 | |
| 5197 | 194 | fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = | 
| 195 | if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u | |
| 5201 | 196 | else [gen_field_tr mark sfx tm] | 
| 197 | | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; | |
| 5197 | 198 | |
| 199 | fun gen_record_tr sep mark sfx unit [t] = foldr (op $) (gen_fields_tr sep mark sfx t, unit) | |
| 5201 | 200 |   | gen_record_tr _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
 | 
| 5197 | 201 | |
| 202 | fun gen_record_scheme_tr sep mark sfx [t, more] = foldr (op $) (gen_fields_tr sep mark sfx t, more) | |
| 5201 | 203 |   | gen_record_scheme_tr _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
 | 
| 5197 | 204 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 205 | |
| 5197 | 206 | val record_type_tr = gen_record_tr "_field_types" "_field_type" field_typeN (Syntax.const "unit"); | 
| 207 | val record_type_scheme_tr = gen_record_scheme_tr "_field_types" "_field_type" field_typeN; | |
| 208 | ||
| 209 | val record_tr = gen_record_tr "_fields" "_field" fieldN HOLogic.unit; | |
| 210 | val record_scheme_tr = gen_record_scheme_tr "_fields" "_field" fieldN; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 211 | |
| 5197 | 212 | fun record_update_tr [t, u] = | 
| 213 | foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) | |
| 214 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | |
| 215 | ||
| 216 | ||
| 217 | val parse_translation = | |
| 218 |  [("_record_type", record_type_tr),
 | |
| 219 |   ("_record_type_scheme", record_type_scheme_tr),
 | |
| 220 |   ("_record", record_tr),
 | |
| 221 |   ("_record_scheme", record_scheme_tr),
 | |
| 222 |   ("_record_update", record_update_tr)];
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 223 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 224 | |
| 4890 | 225 | (* print translations *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 226 | |
| 5197 | 227 | fun gen_fields_tr' mark sfx (tm as Const (name_field, _) $ t $ u) = | 
| 228 | (case try (unsuffix sfx) name_field of | |
| 229 | Some name => | |
| 230 | apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_fields_tr' mark sfx u) | |
| 231 | | None => ([], tm)) | |
| 232 | | gen_fields_tr' _ _ tm = ([], tm); | |
| 233 | ||
| 234 | fun gen_record_tr' sep mark sfx is_unit record record_scheme tm = | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 235 | let | 
| 5197 | 236 | val (ts, u) = gen_fields_tr' mark sfx tm; | 
| 237 | val t' = foldr1 (fn (v, w) => Syntax.const sep $ v $ w) ts; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 238 | in | 
| 5197 | 239 | if is_unit u then Syntax.const record $ t' | 
| 240 | else Syntax.const record_scheme $ t' $ u | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 241 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 242 | |
| 5197 | 243 | |
| 244 | val record_type_tr' = | |
| 245 | gen_record_tr' "_field_types" "_field_type" field_typeN | |
| 246 |     (fn Const ("unit", _) => true | _ => false) "_record_type" "_record_type_scheme";
 | |
| 247 | ||
| 248 | val record_tr' = | |
| 249 | gen_record_tr' "_fields" "_field" fieldN HOLogic.is_unit "_record" "_record_scheme"; | |
| 250 | ||
| 251 | fun record_update_tr' tm = | |
| 252 | let val (ts, u) = gen_fields_tr' "_update" updateN tm in | |
| 253 | Syntax.const "_record_update" $ u $ | |
| 254 | foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) | |
| 255 | end; | |
| 256 | ||
| 257 | ||
| 5201 | 258 | fun gen_field_tr' sfx tr' name = | 
| 259 | let val name_sfx = suffix sfx name | |
| 260 | in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; | |
| 261 | ||
| 5197 | 262 | fun print_translation names = | 
| 263 | map (gen_field_tr' field_typeN record_type_tr') names @ | |
| 264 | map (gen_field_tr' fieldN record_tr') names @ | |
| 265 | map (gen_field_tr' updateN record_update_tr') names; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 266 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 267 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 268 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 269 | (*** extend theory by record definition ***) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 270 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 271 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 272 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 273 | (* type record_info and parent_info *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 274 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 275 | type record_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 276 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 277 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 278 | fields: (string * typ) list, | 
| 4895 | 279 | simps: tthm list}; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 280 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 281 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 282 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 283 | fields: (string * typ) list, | 
| 4895 | 284 | simps: tthm list}; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 285 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 286 | |
| 5052 | 287 | (* data kind 'HOL/records' *) | 
| 5001 | 288 | |
| 5006 | 289 | structure RecordsArgs = | 
| 290 | struct | |
| 291 | val name = "HOL/records"; | |
| 292 | type T = record_info Symtab.table; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 293 | |
| 5006 | 294 | val empty = Symtab.empty; | 
| 295 | val prep_ext = I; | |
| 296 | val merge: T * T -> T = Symtab.merge (K true); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 297 | |
| 5006 | 298 | fun print sg tab = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 299 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 300 | val prt_typ = Sign.pretty_typ sg; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 301 | val ext_const = Sign.cond_extern sg Sign.constK; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 302 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 303 | fun pretty_parent None = [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 304 | | pretty_parent (Some (Ts, name)) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 305 | [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 306 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 307 | fun pretty_field (c, T) = Pretty.block | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 308 | [Pretty.str (ext_const c), Pretty.str " ::", Pretty.brk 1, Pretty.quote (prt_typ T)]; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 309 | |
| 4895 | 310 |       fun pretty_record (name, {args, parent, fields, simps = _}) = Pretty.block (Pretty.fbreaks
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 311 | (Pretty.block [prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 312 | pretty_parent parent @ map pretty_field fields)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 313 | in | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 314 | seq (Pretty.writeln o pretty_record) (Symtab.dest tab) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 315 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 316 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 317 | |
| 5006 | 318 | structure RecordsData = TheoryDataFun(RecordsArgs); | 
| 319 | val print_records = RecordsData.print; | |
| 320 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 321 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 322 | (* get and put records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 323 | |
| 5006 | 324 | fun get_record thy name = Symtab.lookup (RecordsData.get thy, name); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 325 | |
| 4890 | 326 | fun put_record name info thy = | 
| 5006 | 327 | RecordsData.put (Symtab.update ((name, info), RecordsData.get thy)) thy; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 328 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 329 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 330 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 331 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 332 | fun inst_record thy (types, name) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 333 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 334 | val sign = Theory.sign_of thy; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 335 | fun err msg = error (msg ^ " parent record " ^ quote name); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 336 | |
| 4895 | 337 |     val {args, parent, fields, simps} =
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 338 | (case get_record thy name of Some info => info | None => err "Unknown"); | 
| 4895 | 339 | val _ = if length types <> length args then err "Bad number of arguments for" else (); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 340 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 341 | fun bad_inst ((x, S), T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 342 | if Sign.of_sort sign (T, S) then None else Some x | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 343 | val bads = mapfilter bad_inst (args ~~ types); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 344 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 345 | val inst = map fst args ~~ types; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 346 | val subst = Term.map_type_tfree (fn (x, _) => the (assoc (inst, x))); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 347 | in | 
| 4895 | 348 | if not (null bads) then | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 349 |       err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
 | 
| 4895 | 350 | else (apsome (apfst (map subst)) parent, map (apsnd subst) fields, simps) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 351 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 352 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 353 | fun add_parents thy (None, parents) = parents | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 354 | | add_parents thy (Some (types, name), parents) = | 
| 4895 | 355 | let val (pparent, pfields, psimps) = inst_record thy (types, name) | 
| 356 |       in add_parents thy (pparent, {name = name, fields = pfields, simps = psimps} :: parents) end;
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 357 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 358 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 359 | |
| 4890 | 360 | (** internal theory extenders **) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 361 | |
| 4894 | 362 | (* field_definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 363 | |
| 4894 | 364 | (*theorems from Prod.thy*) | 
| 4890 | 365 | val prod_convs = map Attribute.tthm_of [fst_conv, snd_conv]; | 
| 366 | ||
| 367 | ||
| 4894 | 368 | fun field_definitions fields names zeta moreT more vars named_vars thy = | 
| 4890 | 369 | let | 
| 370 | val base = Sign.base_name; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 371 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 372 | |
| 4890 | 373 | (* prepare declarations and definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 374 | |
| 4890 | 375 | (*field types*) | 
| 376 | fun mk_fieldT_spec c = | |
| 377 | (suffix field_typeN c, ["'a", zeta], | |
| 378 |         HOLogic.mk_prodT (TFree ("'a", HOLogic.termS), moreT), Syntax.NoSyn);
 | |
| 379 | val fieldT_specs = map (mk_fieldT_spec o base) names; | |
| 380 | ||
| 381 | (*field declarations*) | |
| 382 | val field_decls = map (mk_fieldC moreT) fields; | |
| 383 | val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields; | |
| 384 | ||
| 385 | (*field constructors*) | |
| 386 | fun mk_field_spec (c, v) = | |
| 4894 | 387 | mk_field ((c, v), more) :== HOLogic.mk_prod (v, more); | 
| 388 | val field_specs = map mk_field_spec named_vars; | |
| 4890 | 389 | |
| 390 | (*field destructors*) | |
| 391 | fun mk_dest_spec dest dest' (c, T) = | |
| 392 | let | |
| 4895 | 393 |         val p = Free ("p", mk_fieldT ((c, T), moreT));
 | 
| 394 |         val p' = Free ("p", HOLogic.mk_prodT (T, moreT));
 | |
| 4894 | 395 | (*note: field types are just abbreviations*) | 
| 396 | in dest p :== dest' p' end; | |
| 4890 | 397 | val dest_specs = | 
| 398 | map (mk_dest_spec mk_fst HOLogic.mk_fst) fields @ | |
| 399 | map (mk_dest_spec mk_snd HOLogic.mk_snd) fields; | |
| 400 | ||
| 401 | ||
| 402 | (* prepare theorems *) | |
| 4894 | 403 | |
| 4890 | 404 | fun mk_dest_prop dest dest' (c, v) = | 
| 4894 | 405 | dest (mk_field ((c, v), more)) === dest' (v, more); | 
| 4890 | 406 | val dest_props = | 
| 4895 | 407 | map (mk_dest_prop mk_fst fst) named_vars @ | 
| 408 | map (mk_dest_prop mk_snd snd) named_vars; | |
| 4890 | 409 | |
| 410 | ||
| 411 | (* 1st stage: defs_thy *) | |
| 412 | ||
| 413 | val defs_thy = | |
| 414 | thy | |
| 415 | |> Theory.add_tyabbrs_i fieldT_specs | |
| 416 | |> (Theory.add_consts_i o map (Syntax.no_syn o apfst base)) | |
| 417 | (field_decls @ dest_decls) | |
| 5197 | 418 | |> (PureThy.add_defs_i o map (fn x => (x, [Attribute.tag_internal]))) | 
| 4890 | 419 | (field_specs @ dest_specs); | 
| 420 | ||
| 421 | val field_defs = get_defs defs_thy field_specs; | |
| 422 | val dest_defs = get_defs defs_thy dest_specs; | |
| 423 | ||
| 424 | ||
| 425 | (* 2nd stage: thms_thy *) | |
| 426 | ||
| 4894 | 427 | val dest_convs = | 
| 4895 | 428 | map (prove_simp defs_thy (field_defs @ dest_defs @ prod_convs)) dest_props; | 
| 4894 | 429 | |
| 4890 | 430 | val thms_thy = | 
| 431 | defs_thy | |
| 432 | |> (PureThy.add_tthmss o map Attribute.none) | |
| 433 |         [("field_defs", field_defs),
 | |
| 434 |           ("dest_defs", dest_defs),
 | |
| 435 |           ("dest_convs", dest_convs)];
 | |
| 436 | ||
| 437 | in (thms_thy, dest_convs) end; | |
| 438 | ||
| 439 | ||
| 440 | (* record_definition *) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 441 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 442 | fun record_definition (args, bname) parent (parents: parent_info list) bfields thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 443 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 444 | val sign = Theory.sign_of thy; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 445 | val full = Sign.full_name_path sign bname; | 
| 4890 | 446 | val base = Sign.base_name; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 447 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 448 | |
| 4890 | 449 | (* basic components *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 450 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 451 | val alphas = map fst args; | 
| 5197 | 452 | val name = Sign.full_name sign bname; (*not made part of record name space!*) | 
| 4890 | 453 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 454 | val parent_fields = flat (map #fields parents); | 
| 4890 | 455 | val parent_names = map fst parent_fields; | 
| 456 | val parent_types = map snd parent_fields; | |
| 457 | val parent_len = length parent_fields; | |
| 458 | val parent_xs = variantlist (map (base o fst) parent_fields, [moreN]); | |
| 459 | val parent_vars = ListPair.map Free (parent_xs, parent_types); | |
| 4894 | 460 | val parent_named_vars = parent_names ~~ parent_vars; | 
| 4890 | 461 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 462 | val fields = map (apfst full) bfields; | 
| 4890 | 463 | val names = map fst fields; | 
| 464 | val types = map snd fields; | |
| 465 | val len = length fields; | |
| 466 | val xs = variantlist (map fst bfields, moreN :: parent_xs); | |
| 467 | val vars = ListPair.map Free (xs, types); | |
| 4894 | 468 | val named_vars = names ~~ vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 469 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 470 | val all_fields = parent_fields @ fields; | 
| 4890 | 471 | val all_names = parent_names @ names; | 
| 472 | val all_types = parent_types @ types; | |
| 473 | val all_len = parent_len + len; | |
| 474 | val all_xs = parent_xs @ xs; | |
| 475 | val all_vars = parent_vars @ vars; | |
| 4894 | 476 | val all_named_vars = parent_named_vars @ named_vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 477 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 478 | val zeta = variant alphas "'z"; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 479 | val moreT = TFree (zeta, moreS); | 
| 4895 | 480 | val more = Free (moreN, moreT); | 
| 5197 | 481 | val full_moreN = full moreN; | 
| 482 | fun more_part t = mk_more t full_moreN; | |
| 483 | fun more_part_update t x = mk_more_update t (full_moreN, x); | |
| 4894 | 484 | |
| 485 | val parent_more = funpow parent_len mk_snd; | |
| 486 | val idxs = 0 upto (len - 1); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 487 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 488 | val rec_schemeT = mk_recordT (all_fields, moreT); | 
| 4894 | 489 | val rec_scheme = mk_record (all_named_vars, more); | 
| 4890 | 490 |     val r = Free ("r", rec_schemeT);
 | 
| 4894 | 491 | val recT = mk_recordT (all_fields, HOLogic.unitT); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 492 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 493 | |
| 4890 | 494 | (* prepare print translation functions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 495 | |
| 5197 | 496 | val accesses = distinct (flat (map NameSpace.accesses (full_moreN :: names))); | 
| 497 | val (_, _, tr'_names, _) = Syntax.trfun_names (Theory.syn_of thy); | |
| 498 | val field_tr's = filter_out (fn (c, _) => c mem tr'_names) (print_translation accesses); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 499 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 500 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 501 | (* prepare declarations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 502 | |
| 5197 | 503 | val sel_decls = map (mk_selC rec_schemeT) bfields @ | 
| 504 | [mk_moreC rec_schemeT (moreN, moreT)]; | |
| 505 | val update_decls = map (mk_updateC rec_schemeT) bfields @ | |
| 506 | [mk_more_updateC rec_schemeT (moreN, moreT)]; | |
| 4890 | 507 | val make_decls = | 
| 508 | [(mk_makeC rec_schemeT (make_schemeN, all_types @ [moreT])), | |
| 509 | (mk_makeC recT (makeN, all_types))]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 510 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 511 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 512 | (* prepare definitions *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 513 | |
| 4895 | 514 | (*record (scheme) type abbreviation*) | 
| 4890 | 515 | val recordT_specs = | 
| 516 | [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn), | |
| 517 | (bname, alphas, recT, Syntax.NoSyn)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 518 | |
| 4894 | 519 | (*selectors*) | 
| 4890 | 520 | fun mk_sel_spec (i, c) = | 
| 4894 | 521 | mk_sel r c :== mk_fst (funpow i mk_snd (parent_more r)); | 
| 522 | val sel_specs = | |
| 523 | ListPair.map mk_sel_spec (idxs, names) @ | |
| 524 | [more_part r :== funpow len mk_snd (parent_more r)]; | |
| 4890 | 525 | |
| 526 | (*updates*) | |
| 4894 | 527 | val all_sels = all_names ~~ map (mk_sel r) all_names; | 
| 4890 | 528 | fun mk_upd_spec (i, (c, x)) = | 
| 4894 | 529 | mk_update r (c, x) :== | 
| 530 | mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part r) | |
| 5197 | 531 | val update_specs = | 
| 532 | ListPair.map mk_upd_spec (idxs, named_vars) @ | |
| 533 | [more_part_update r more :== mk_record (all_sels, more)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 534 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 535 | (*makes*) | 
| 4890 | 536 | val make_scheme = Const (mk_makeC rec_schemeT (full make_schemeN, all_types @ [moreT])); | 
| 537 | val make = Const (mk_makeC recT (full makeN, all_types)); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 538 | val make_specs = | 
| 4894 | 539 | [list_comb (make_scheme, all_vars) $ more :== rec_scheme, | 
| 540 | list_comb (make, all_vars) :== mk_record (all_named_vars, HOLogic.unit)]; | |
| 541 | ||
| 542 | ||
| 543 | (* prepare propositions *) | |
| 544 | ||
| 545 | (*selectors*) | |
| 546 | val sel_props = | |
| 547 | map (fn (c, x) => mk_sel rec_scheme c === x) named_vars @ | |
| 548 | [more_part rec_scheme === more]; | |
| 549 | ||
| 550 | (*updates*) | |
| 551 | fun mk_upd_prop (i, (c, T)) = | |
| 552 | let val x' = Free (variant all_xs (base c ^ "'"), T) in | |
| 553 | mk_update rec_scheme (c, x') === | |
| 554 | mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more) | |
| 555 | end; | |
| 5197 | 556 | val update_props = | 
| 557 | ListPair.map mk_upd_prop (idxs, fields) @ | |
| 558 | let val more' = Free (variant all_xs (moreN ^ "'"), moreT) | |
| 559 | in [more_part_update rec_scheme more' === mk_record (all_named_vars, more')] end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 560 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 561 | |
| 4890 | 562 | (* 1st stage: fields_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 563 | |
| 4890 | 564 | val (fields_thy, field_simps) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 565 | thy | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 566 | |> Theory.add_path bname | 
| 4894 | 567 | |> field_definitions fields names zeta moreT more vars named_vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 568 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 569 | |
| 4890 | 570 | (* 2nd stage: defs_thy *) | 
| 571 | ||
| 572 | val defs_thy = | |
| 573 | fields_thy | |
| 574 | |> Theory.parent_path | |
| 5197 | 575 | |> Theory.add_tyabbrs_i recordT_specs (*not made part of record name space!*) | 
| 4890 | 576 | |> Theory.add_path bname | 
| 5197 | 577 | |> Theory.add_trfuns ([], [], field_tr's, []) | 
| 4894 | 578 | |> (Theory.add_consts_i o map Syntax.no_syn) | 
| 579 | (sel_decls @ update_decls @ make_decls) | |
| 5197 | 580 | |> (PureThy.add_defs_i o map (fn x => (x, [Attribute.tag_internal]))) | 
| 5212 | 581 | (sel_specs @ update_specs) | 
| 582 | |> (PureThy.add_defs_i o map Attribute.none) make_specs; | |
| 4890 | 583 | |
| 584 | val sel_defs = get_defs defs_thy sel_specs; | |
| 585 | val update_defs = get_defs defs_thy update_specs; | |
| 586 | val make_defs = get_defs defs_thy make_specs; | |
| 587 | ||
| 588 | ||
| 589 | (* 3rd stage: thms_thy *) | |
| 590 | ||
| 4895 | 591 | val parent_simps = flat (map #simps parents); | 
| 592 | val prove = prove_simp defs_thy; | |
| 4890 | 593 | |
| 4895 | 594 | val sel_convs = map (prove (parent_simps @ sel_defs @ field_simps)) sel_props; | 
| 595 | val update_convs = map (prove (parent_simps @ update_defs @ sel_convs)) update_props; | |
| 4894 | 596 | |
| 597 | val simps = field_simps @ sel_convs @ update_convs @ make_defs; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 598 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 599 | val thms_thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 600 | defs_thy | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 601 | |> (PureThy.add_tthmss o map Attribute.none) | 
| 4895 | 602 |         [("select_defs", sel_defs),
 | 
| 4890 | 603 |           ("update_defs", update_defs),
 | 
| 4894 | 604 |           ("make_defs", make_defs),
 | 
| 4895 | 605 |           ("select_convs", sel_convs),
 | 
| 4894 | 606 |           ("update_convs", update_convs)]
 | 
| 607 |       |> PureThy.add_tthmss [(("simps", simps), [Simplifier.simp_add_global])];
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 608 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 609 | |
| 4890 | 610 | (* 4th stage: final_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 611 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 612 | val final_thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 613 | thms_thy | 
| 4895 | 614 |       |> put_record name {args = args, parent = parent, fields = fields, simps = simps}
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 615 | |> Theory.parent_path; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 616 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 617 | in final_thy end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 618 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 619 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 620 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 621 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 622 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 623 | (* prepare arguments *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 624 | |
| 4894 | 625 | (*note: read_raw_typ avoids expanding type abbreviations*) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 626 | fun read_raw_parent sign s = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 627 | (case Sign.read_raw_typ (sign, K None) s handle TYPE (msg, _, _) => error msg of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 628 | Type (name, Ts) => (Ts, name) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 629 |   | _ => error ("Bad parent record specification: " ^ quote s));
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 630 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 631 | fun read_typ sign (env, s) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 632 | let | 
| 5060 | 633 | fun def_sort (x, ~1) = assoc (env, x) | 
| 634 | | def_sort _ = None; | |
| 635 | val T = Type.no_tvars (Sign.read_typ (sign, def_sort) s) handle TYPE (msg, _, _) => error msg; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 636 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 637 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 638 | fun cert_typ sign (env, raw_T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 639 | let val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle TYPE (msg, _, _) => error msg | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 640 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 641 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 642 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 643 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 644 | |
| 4895 | 645 | (*we do all preparations and error checks here, deferring the real | 
| 646 | work to record_definition*) | |
| 4890 | 647 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 648 | fun gen_add_record prep_typ prep_raw_parent (params, bname) raw_parent raw_fields thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 649 | let | 
| 4970 | 650 | val _ = Theory.requires thy "Record" "record definitions"; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 651 | val sign = Theory.sign_of thy; | 
| 4895 | 652 |     val _ = writeln ("Defining record " ^ quote bname ^ " ...");
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 653 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 654 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 655 | (* parents *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 656 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 657 | fun prep_inst T = snd (cert_typ sign ([], T)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 658 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 659 | val parent = apsome (apfst (map prep_inst) o prep_raw_parent sign) raw_parent | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 660 |       handle ERROR => error ("The error(s) above in parent record specification");
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 661 | val parents = add_parents thy (parent, []); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 662 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 663 | val init_env = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 664 | (case parent of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 665 | None => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 666 | | Some (types, _) => foldr Term.add_typ_tfrees (types, [])); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 667 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 668 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 669 | (* fields *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 670 | |
| 4967 | 671 | fun prep_field (env, (c, raw_T)) = | 
| 672 | let val (env', T) = prep_typ sign (env, raw_T) handle ERROR => | |
| 673 |         error ("The error(s) above occured in field " ^ quote c)
 | |
| 674 | in (env', (c, T)) end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 675 | |
| 4967 | 676 | val (envir, bfields) = foldl_map prep_field (init_env, raw_fields); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 677 | val envir_names = map fst envir; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 678 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 679 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 680 | (* args *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 681 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 682 | val defaultS = Sign.defaultS sign; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 683 | val args = map (fn x => (x, if_none (assoc (envir, x)) defaultS)) params; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 684 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 685 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 686 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 687 | |
| 4890 | 688 | val name = Sign.full_name sign bname; | 
| 689 | val err_dup_record = | |
| 690 | if is_none (get_record thy name) then [] | |
| 691 | else ["Duplicate definition of record " ^ quote name]; | |
| 692 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 693 | val err_dup_parms = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 694 | (case duplicates params of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 695 | [] => [] | 
| 4890 | 696 | | dups => ["Duplicate parameter(s) " ^ commas dups]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 697 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 698 | val err_extra_frees = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 699 | (case gen_rems (op =) (envir_names, params) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 700 | [] => [] | 
| 4890 | 701 | | extras => ["Extra free type variable(s) " ^ commas extras]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 702 | |
| 4890 | 703 | val err_no_fields = if null bfields then ["No fields present"] else []; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 704 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 705 | val err_dup_fields = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 706 | (case duplicates (map fst bfields) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 707 | [] => [] | 
| 4890 | 708 | | dups => ["Duplicate field(s) " ^ commas_quote dups]); | 
| 709 | ||
| 710 | val err_bad_fields = | |
| 711 | if forall (not_equal moreN o fst) bfields then [] | |
| 712 | else ["Illegal field name " ^ quote moreN]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 713 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 714 | val err_dup_sorts = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 715 | (case duplicates envir_names of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 716 | [] => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 717 | | dups => ["Inconsistent sort constraints for " ^ commas dups]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 718 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 719 | val errs = | 
| 4890 | 720 | err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @ | 
| 721 | err_dup_fields @ err_bad_fields @ err_dup_sorts; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 722 | in | 
| 4890 | 723 | if null errs then () else error (cat_lines errs); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 724 | thy |> record_definition (args, bname) parent parents bfields | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 725 | end | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 726 |   handle ERROR => error ("Failed to define record " ^ quote bname);
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 727 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 728 | val add_record = gen_add_record read_typ read_raw_parent; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 729 | val add_record_i = gen_add_record cert_typ (K I); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 730 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 731 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 732 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 733 | (** setup theory **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 734 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 735 | val setup = | 
| 5006 | 736 | [RecordsData.init, | 
| 5197 | 737 | Theory.add_trfuns ([], parse_translation, [], [])]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 738 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 739 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 740 | end; |