| author | wenzelm | 
| Mon, 12 Jul 1999 22:23:31 +0200 | |
| changeset 6977 | 4781c0673e83 | 
| parent 6851 | 526c0b90bcef | 
| child 7178 | 50b9849cf6ad | 
| 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. | 
| 5698 | 6 | *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 7 | |
| 5698 | 8 | signature BASIC_RECORD_PACKAGE = | 
| 9 | sig | |
| 10 | val record_split_tac: int -> tactic | |
| 5713 | 11 | val record_split_name: string | 
| 5698 | 12 | val record_split_wrapper: string * wrapper | 
| 13 | end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 14 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 15 | signature RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 16 | sig | 
| 5698 | 17 | include BASIC_RECORD_PACKAGE | 
| 18 | val quiet_mode: bool ref | |
| 4890 | 19 | val moreS: sort | 
| 20 | val mk_fieldT: (string * typ) * typ -> typ | |
| 21 | val dest_fieldT: typ -> (string * typ) * typ | |
| 22 | val mk_field: (string * term) * term -> term | |
| 23 | val mk_fst: term -> term | |
| 24 | val mk_snd: term -> term | |
| 25 | val mk_recordT: (string * typ) list * typ -> typ | |
| 26 | val dest_recordT: typ -> (string * typ) list * typ | |
| 27 | val mk_record: (string * term) list * term -> term | |
| 28 | val mk_sel: term -> string -> term | |
| 29 | val mk_update: term -> string * term -> term | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 30 | val print_records: theory -> unit | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 31 | val add_record: (string list * bstring) -> string option | 
| 6519 | 32 |     -> (bstring * string) list -> theory -> theory * {simps: thm list, iffs: thm list}
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 33 | val add_record_i: (string list * bstring) -> (typ list * string) option | 
| 6519 | 34 |     -> (bstring * typ) list -> theory -> theory * {simps: thm list, iffs: thm list}
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 35 | val setup: (theory -> theory) list | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 36 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 37 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 38 | structure RecordPackage: RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 39 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 40 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 41 | |
| 4894 | 42 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 43 | |
| 5698 | 44 | (* messages *) | 
| 45 | ||
| 46 | val quiet_mode = ref false; | |
| 47 | fun message s = if ! quiet_mode then () else writeln s; | |
| 48 | ||
| 49 | ||
| 6519 | 50 | (* wrappers *) | 
| 5707 | 51 | |
| 52 | fun add_wrapper wrapper thy = | |
| 5713 | 53 | let val r = Classical.claset_ref_of thy | 
| 5707 | 54 | in r := ! r addSWrapper wrapper; thy end; | 
| 55 | ||
| 56 | ||
| 4894 | 57 | (* definitions and equations *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 58 | |
| 5698 | 59 | infix 0 :== ===; | 
| 4894 | 60 | |
| 61 | val (op :==) = Logic.mk_defpair; | |
| 62 | val (op ===) = HOLogic.mk_Trueprop o HOLogic.mk_eq; | |
| 63 | ||
| 6092 | 64 | fun get_defs thy specs = map (PureThy.get_thm thy o fst) specs; | 
| 4894 | 65 | |
| 66 | ||
| 67 | (* proof by simplification *) | |
| 4890 | 68 | |
| 5698 | 69 | fun prove_simp thy tacs simps = | 
| 4894 | 70 | let | 
| 4895 | 71 | val sign = Theory.sign_of thy; | 
| 6092 | 72 | val ss = Simplifier.addsimps (HOL_basic_ss, simps); | 
| 4895 | 73 | |
| 74 | fun prove goal = | |
| 6092 | 75 | Goals.prove_goalw_cterm [] (Thm.cterm_of sign goal) | 
| 76 | (K (tacs @ [ALLGOALS (Simplifier.simp_tac ss)])) | |
| 77 |       handle ERROR => error ("The error(s) above occurred while trying to prove "
 | |
| 78 | ^ quote (Sign.string_of_term sign goal)); | |
| 4894 | 79 | in prove end; | 
| 80 | ||
| 81 | ||
| 82 | ||
| 83 | (*** syntax operations ***) | |
| 84 | ||
| 85 | (** name components **) | |
| 86 | ||
| 87 | val moreN = "more"; | |
| 88 | val schemeN = "_scheme"; | |
| 89 | val fieldN = "_field"; | |
| 5698 | 90 | val raw_fieldN = "_raw_field"; | 
| 4894 | 91 | val field_typeN = "_field_type"; | 
| 5698 | 92 | val fstN = "_val"; | 
| 93 | val sndN = "_more"; | |
| 4894 | 94 | val updateN = "_update"; | 
| 95 | val makeN = "make"; | |
| 96 | val make_schemeN = "make_scheme"; | |
| 4890 | 97 | |
| 5698 | 98 | (*see datatype package*) | 
| 99 | val caseN = "_case"; | |
| 100 | ||
| 101 | ||
| 102 | ||
| 103 | (** generic operations **) | |
| 104 | ||
| 5713 | 105 | (* adhoc priming of vars *) | 
| 106 | ||
| 107 | fun prime (Free (x, T)) = Free (x ^ "'", T) | |
| 108 |   | prime t = raise TERM ("prime: no free variable", [t]);
 | |
| 109 | ||
| 110 | ||
| 111 | (* product case *) | |
| 112 | ||
| 5698 | 113 | fun fst_fn T U = Abs ("x", T, Abs ("y", U, Bound 1));
 | 
| 114 | fun snd_fn T U = Abs ("x", T, Abs ("y", U, Bound 0));
 | |
| 115 | ||
| 116 | fun mk_prod_case name f p = | |
| 117 | let | |
| 118 |     val fT as Type ("fun", [A, Type ("fun", [B, C])]) = fastype_of f;
 | |
| 119 | val pT = fastype_of p; | |
| 120 | in Const (suffix caseN name, fT --> pT --> C) $ f $ p end; | |
| 121 | ||
| 4890 | 122 | |
| 123 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 124 | (** tuple operations **) | 
| 
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 | (* more type class *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 127 | |
| 5210 | 128 | val moreS = ["Record.more"]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 129 | |
| 
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 | (* types *) | 
| 
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 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 | 134 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 135 | 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 | 136 | (case try (unsuffix field_typeN) c_field_type of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 137 |         None => raise TYPE ("dest_fieldT", [typ], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 138 | | Some c => ((c, T), U)) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 139 |   | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 140 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 141 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 142 | (* constructors *) | 
| 
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 | 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 | 145 | |
| 5698 | 146 | fun gen_mk_field sfx ((c, t), u) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 147 | let val T = fastype_of t and U = fastype_of u | 
| 5698 | 148 | in Const (suffix sfx c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end; | 
| 149 | ||
| 150 | val mk_field = gen_mk_field fieldN; | |
| 151 | val mk_raw_field = gen_mk_field raw_fieldN; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 152 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 153 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 154 | (* destructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 155 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 156 | 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 | 157 | 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 | 158 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 159 | fun dest_field fst_or_snd p = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 160 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 161 | val pT = fastype_of p; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 162 | val ((c, T), U) = dest_fieldT pT; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 163 | 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 | 164 | in Const (suffix destN c, pT --> destT) $ p end; | 
| 
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 | val mk_fst = dest_field true; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 167 | val mk_snd = dest_field false; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 168 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 169 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 170 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 171 | (** record operations **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 172 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 173 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 174 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 175 | val mk_recordT = foldr mk_fieldT; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 176 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 177 | fun dest_recordT T = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 178 | (case try dest_fieldT T of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 179 | None => ([], T) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 180 | | 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 | 181 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 182 | fun find_fieldT c rT = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 183 | (case assoc (fst (dest_recordT rT), c) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 184 |     None => raise TYPE ("find_field: " ^ c, [rT], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 185 | | Some T => T); | 
| 
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 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 188 | (* constructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 189 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 190 | val mk_record = foldr mk_field; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 191 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 192 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 193 | (* selectors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 194 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 195 | fun mk_selC rT (c, T) = (c, rT --> T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 196 | |
| 4890 | 197 | fun mk_sel r c = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 198 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 199 | 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 | 200 | |
| 4894 | 201 | val mk_moreC = mk_selC; | 
| 202 | ||
| 203 | fun mk_more r c = | |
| 204 | let val rT = fastype_of r | |
| 205 | in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end; | |
| 206 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 207 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 208 | (* updates *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 209 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 210 | 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 | 211 | |
| 4890 | 212 | fun mk_update r (c, x) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 213 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 214 | 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 | 215 | |
| 5197 | 216 | val mk_more_updateC = mk_updateC; | 
| 217 | ||
| 218 | fun mk_more_update r (c, x) = | |
| 219 | let val rT = fastype_of r | |
| 220 | in Const (mk_more_updateC rT (c, snd (dest_recordT rT))) $ x $ r end; | |
| 221 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 222 | |
| 4890 | 223 | (* make *) | 
| 224 | ||
| 225 | fun mk_makeC rT (c, Ts) = (c, Ts ---> rT); | |
| 226 | ||
| 227 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 228 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 229 | (** concrete syntax for records **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 230 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 231 | (* parse translations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 232 | |
| 5197 | 233 | fun gen_field_tr mark sfx (t as Const (c, _) $ Free (name, _) $ arg) = | 
| 234 | if c = mark then Syntax.const (suffix sfx name) $ arg | |
| 235 |       else raise TERM ("gen_field_tr: " ^ mark, [t])
 | |
| 236 |   | 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 | 237 | |
| 5197 | 238 | fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = | 
| 239 | if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u | |
| 5201 | 240 | else [gen_field_tr mark sfx tm] | 
| 241 | | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; | |
| 5197 | 242 | |
| 243 | fun gen_record_tr sep mark sfx unit [t] = foldr (op $) (gen_fields_tr sep mark sfx t, unit) | |
| 5201 | 244 |   | gen_record_tr _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
 | 
| 5197 | 245 | |
| 246 | fun gen_record_scheme_tr sep mark sfx [t, more] = foldr (op $) (gen_fields_tr sep mark sfx t, more) | |
| 5201 | 247 |   | gen_record_scheme_tr _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
 | 
| 5197 | 248 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 249 | |
| 5197 | 250 | val record_type_tr = gen_record_tr "_field_types" "_field_type" field_typeN (Syntax.const "unit"); | 
| 251 | val record_type_scheme_tr = gen_record_scheme_tr "_field_types" "_field_type" field_typeN; | |
| 252 | ||
| 253 | val record_tr = gen_record_tr "_fields" "_field" fieldN HOLogic.unit; | |
| 254 | 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 | 255 | |
| 5197 | 256 | fun record_update_tr [t, u] = | 
| 257 | foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) | |
| 258 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | |
| 259 | ||
| 260 | ||
| 261 | val parse_translation = | |
| 262 |  [("_record_type", record_type_tr),
 | |
| 263 |   ("_record_type_scheme", record_type_scheme_tr),
 | |
| 264 |   ("_record", record_tr),
 | |
| 265 |   ("_record_scheme", record_scheme_tr),
 | |
| 266 |   ("_record_update", record_update_tr)];
 | |
| 4867 
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 | |
| 4890 | 269 | (* print translations *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 270 | |
| 5197 | 271 | fun gen_fields_tr' mark sfx (tm as Const (name_field, _) $ t $ u) = | 
| 272 | (case try (unsuffix sfx) name_field of | |
| 273 | Some name => | |
| 274 | apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_fields_tr' mark sfx u) | |
| 275 | | None => ([], tm)) | |
| 276 | | gen_fields_tr' _ _ tm = ([], tm); | |
| 277 | ||
| 278 | 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 | 279 | let | 
| 5197 | 280 | val (ts, u) = gen_fields_tr' mark sfx tm; | 
| 281 | 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 | 282 | in | 
| 5197 | 283 | if is_unit u then Syntax.const record $ t' | 
| 284 | else Syntax.const record_scheme $ t' $ u | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 285 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 286 | |
| 5197 | 287 | |
| 288 | val record_type_tr' = | |
| 289 | gen_record_tr' "_field_types" "_field_type" field_typeN | |
| 290 |     (fn Const ("unit", _) => true | _ => false) "_record_type" "_record_type_scheme";
 | |
| 291 | ||
| 292 | val record_tr' = | |
| 293 | gen_record_tr' "_fields" "_field" fieldN HOLogic.is_unit "_record" "_record_scheme"; | |
| 294 | ||
| 295 | fun record_update_tr' tm = | |
| 296 | let val (ts, u) = gen_fields_tr' "_update" updateN tm in | |
| 297 | Syntax.const "_record_update" $ u $ | |
| 298 | foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) | |
| 299 | end; | |
| 300 | ||
| 301 | ||
| 5201 | 302 | fun gen_field_tr' sfx tr' name = | 
| 303 | let val name_sfx = suffix sfx name | |
| 304 | in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; | |
| 305 | ||
| 5197 | 306 | fun print_translation names = | 
| 307 | map (gen_field_tr' field_typeN record_type_tr') names @ | |
| 308 | map (gen_field_tr' fieldN record_tr') names @ | |
| 309 | map (gen_field_tr' updateN record_update_tr') names; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 310 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 311 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 312 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 313 | (*** extend theory by record definition ***) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 314 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 315 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 316 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 317 | (* type record_info and parent_info *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 318 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 319 | type record_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 320 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 321 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 322 | fields: (string * typ) list, | 
| 6092 | 323 | simps: thm list}; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 324 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 325 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 326 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 327 | fields: (string * typ) list, | 
| 6092 | 328 | simps: thm list}; | 
| 4867 
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 | |
| 5052 | 331 | (* data kind 'HOL/records' *) | 
| 5001 | 332 | |
| 5006 | 333 | structure RecordsArgs = | 
| 334 | struct | |
| 335 | val name = "HOL/records"; | |
| 5698 | 336 | type T = | 
| 5707 | 337 | record_info Symtab.table * (*records*) | 
| 338 | (thm Symtab.table * Simplifier.simpset); (*field split rules*) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 339 | |
| 5698 | 340 | val empty = (Symtab.empty, (Symtab.empty, HOL_basic_ss)); | 
| 6556 | 341 | val copy = I; | 
| 5006 | 342 | val prep_ext = I; | 
| 5698 | 343 | fun merge ((recs1, (sps1, ss1)), (recs2, (sps2, ss2))) = | 
| 344 | (Symtab.merge (K true) (recs1, recs2), | |
| 345 | (Symtab.merge (K true) (sps1, sps2), Simplifier.merge_ss (ss1, ss2))); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 346 | |
| 5698 | 347 | fun print sg (recs, _) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 348 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 349 | val prt_typ = Sign.pretty_typ sg; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 350 | val ext_const = Sign.cond_extern sg Sign.constK; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 351 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 352 | fun pretty_parent None = [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 353 | | pretty_parent (Some (Ts, name)) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 354 | [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 355 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 356 | fun pretty_field (c, T) = Pretty.block | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 357 | [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 | 358 | |
| 4895 | 359 |       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 | 360 | (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 | 361 | pretty_parent parent @ map pretty_field fields)); | 
| 6851 | 362 | in seq (Pretty.writeln o pretty_record) (Sign.cond_extern_table sg Sign.typeK recs) end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 363 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 364 | |
| 5006 | 365 | structure RecordsData = TheoryDataFun(RecordsArgs); | 
| 366 | val print_records = RecordsData.print; | |
| 367 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 368 | |
| 5698 | 369 | (* get and put data *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 370 | |
| 5698 | 371 | fun get_record thy name = Symtab.lookup (#1 (RecordsData.get thy), name); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 372 | |
| 4890 | 373 | fun put_record name info thy = | 
| 5698 | 374 | let val (tab, sp) = RecordsData.get thy | 
| 375 | in RecordsData.put (Symtab.update ((name, info), tab), sp) thy end; | |
| 376 | ||
| 377 | fun add_record_splits splits thy = | |
| 378 | let | |
| 379 | val (tab, (sps, ss)) = RecordsData.get thy; | |
| 380 | val simps = map #2 splits; | |
| 381 | in RecordsData.put (tab, (Symtab.extend (sps, splits), Simplifier.addsimps (ss, simps))) thy end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 382 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 383 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 384 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 385 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 386 | fun inst_record thy (types, name) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 387 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 388 | val sign = Theory.sign_of thy; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 389 | fun err msg = error (msg ^ " parent record " ^ quote name); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 390 | |
| 4895 | 391 |     val {args, parent, fields, simps} =
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 392 | (case get_record thy name of Some info => info | None => err "Unknown"); | 
| 4895 | 393 | 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 | 394 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 395 | fun bad_inst ((x, S), T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 396 | 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 | 397 | val bads = mapfilter bad_inst (args ~~ types); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 398 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 399 | val inst = map fst args ~~ types; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 400 | 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 | 401 | in | 
| 4895 | 402 | if not (null bads) then | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 403 |       err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
 | 
| 4895 | 404 | 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 | 405 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 406 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 407 | fun add_parents thy (None, parents) = parents | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 408 | | add_parents thy (Some (types, name), parents) = | 
| 4895 | 409 | let val (pparent, pfields, psimps) = inst_record thy (types, name) | 
| 410 |       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 | 411 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 412 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 413 | |
| 5698 | 414 | (** record field splitting **) | 
| 415 | ||
| 6358 | 416 | (* tactic *) | 
| 417 | ||
| 5698 | 418 | fun record_split_tac i st = | 
| 419 | let | |
| 420 | val (_, (sps, ss)) = RecordsData.get_sg (Thm.sign_of_thm st); | |
| 421 | ||
| 422 | fun is_fieldT (_, Type (a, [_, _])) = is_some (Symtab.lookup (sps, a)) | |
| 423 | | is_fieldT _ = false; | |
| 424 | val params = Logic.strip_params (Library.nth_elem (i - 1, Thm.prems_of st)); | |
| 425 | in | |
| 426 | if exists is_fieldT params then Simplifier.full_simp_tac ss i st | |
| 427 | else Seq.empty | |
| 428 | end handle Library.LIST _ => Seq.empty; | |
| 429 | ||
| 6358 | 430 | |
| 431 | (* wrapper *) | |
| 432 | ||
| 5707 | 433 | val record_split_name = "record_split_tac"; | 
| 434 | val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac); | |
| 5698 | 435 | |
| 436 | ||
| 6358 | 437 | (* method *) | 
| 438 | ||
| 439 | val record_split_method = | |
| 440 |   ("record_split", Method.no_args (Method.METHOD0 (FIRSTGOAL record_split_tac)),
 | |
| 441 | "split record fields"); | |
| 442 | ||
| 443 | ||
| 5698 | 444 | |
| 4890 | 445 | (** internal theory extenders **) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 446 | |
| 5698 | 447 | (* field_type_defs *) | 
| 448 | ||
| 5713 | 449 | fun field_type_def ((thy, simps), (name, tname, vs, T, U)) = | 
| 5698 | 450 | let | 
| 6394 | 451 | val full = Sign.full_name (Theory.sign_of thy); | 
| 5713 | 452 |     val (thy', {simps = simps', ...}) =
 | 
| 5698 | 453 | thy | 
| 454 | |> setmp DatatypePackage.quiet_mode true | |
| 455 | (DatatypePackage.add_datatype_i true [tname] | |
| 456 | [(vs, tname, Syntax.NoSyn, [(name, [T, U], Syntax.NoSyn)])]); | |
| 457 | val thy'' = | |
| 458 | thy' | |
| 459 | |> setmp AxClass.quiet_mode true | |
| 460 | (AxClass.add_inst_arity_i (full tname, [HOLogic.termS, moreS], moreS) [] [] None); | |
| 5713 | 461 | in (thy'', simps' @ simps) end; | 
| 5698 | 462 | |
| 5713 | 463 | fun field_type_defs args thy = foldl field_type_def ((thy, []), args); | 
| 5698 | 464 | |
| 465 | ||
| 4894 | 466 | (* field_definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 467 | |
| 4894 | 468 | fun field_definitions fields names zeta moreT more vars named_vars thy = | 
| 4890 | 469 | let | 
| 5698 | 470 | val sign = Theory.sign_of thy; | 
| 4890 | 471 | val base = Sign.base_name; | 
| 5698 | 472 | val full_path = Sign.full_name_path sign; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 473 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 474 | |
| 4890 | 475 | (* prepare declarations and definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 476 | |
| 4890 | 477 | (*field types*) | 
| 478 | fun mk_fieldT_spec c = | |
| 5698 | 479 | (suffix raw_fieldN c, suffix field_typeN c, | 
| 480 |         ["'a", zeta], TFree ("'a", HOLogic.termS), moreT);
 | |
| 4890 | 481 | val fieldT_specs = map (mk_fieldT_spec o base) names; | 
| 482 | ||
| 5698 | 483 | (*field constructors*) | 
| 4890 | 484 | val field_decls = map (mk_fieldC moreT) fields; | 
| 485 | ||
| 486 | fun mk_field_spec (c, v) = | |
| 5698 | 487 | mk_field ((c, v), more) :== mk_raw_field ((c, v), more); | 
| 4894 | 488 | val field_specs = map mk_field_spec named_vars; | 
| 4890 | 489 | |
| 490 | (*field destructors*) | |
| 5698 | 491 | val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields; | 
| 492 | ||
| 493 | fun mk_dest_spec dest f (c, T) = | |
| 494 |       let val p = Free ("p", mk_fieldT ((c, T), moreT));
 | |
| 495 | in dest p :== mk_prod_case (suffix field_typeN c) (f T moreT) p end; | |
| 4890 | 496 | val dest_specs = | 
| 5698 | 497 | map (mk_dest_spec mk_fst fst_fn) fields @ | 
| 498 | map (mk_dest_spec mk_snd snd_fn) fields; | |
| 4890 | 499 | |
| 500 | ||
| 501 | (* prepare theorems *) | |
| 4894 | 502 | |
| 5713 | 503 | (*constructor injects*) | 
| 504 | fun mk_inject_prop (c, v) = | |
| 505 | HOLogic.mk_eq (mk_field ((c, v), more), mk_field ((c, prime v), prime more)) === | |
| 506 | (HOLogic.conj $ HOLogic.mk_eq (v, prime v) $ HOLogic.mk_eq (more, prime more)); | |
| 507 | val inject_props = map mk_inject_prop named_vars; | |
| 508 | ||
| 5698 | 509 | (*destructor conversions*) | 
| 4890 | 510 | fun mk_dest_prop dest dest' (c, v) = | 
| 4894 | 511 | dest (mk_field ((c, v), more)) === dest' (v, more); | 
| 4890 | 512 | val dest_props = | 
| 4895 | 513 | map (mk_dest_prop mk_fst fst) named_vars @ | 
| 514 | map (mk_dest_prop mk_snd snd) named_vars; | |
| 4890 | 515 | |
| 5698 | 516 | (*surjective pairing*) | 
| 517 | fun mk_surj_prop (c, T) = | |
| 518 |       let val p = Free ("p", mk_fieldT ((c, T), moreT));
 | |
| 519 | in p === mk_field ((c, mk_fst p), mk_snd p) end; | |
| 520 | val surj_props = map mk_surj_prop fields; | |
| 4890 | 521 | |
| 5698 | 522 | |
| 523 | (* 1st stage: types_thy *) | |
| 524 | ||
| 6092 | 525 | val (types_thy, datatype_simps) = | 
| 5698 | 526 | thy | 
| 527 | |> field_type_defs fieldT_specs; | |
| 528 | ||
| 529 | ||
| 530 | (* 2nd stage: defs_thy *) | |
| 4890 | 531 | |
| 532 | val defs_thy = | |
| 5698 | 533 | types_thy | 
| 534 | |> (Theory.add_consts_i o map (Syntax.no_syn o apfst base)) | |
| 535 | (field_decls @ dest_decls) | |
| 6092 | 536 | |> (PureThy.add_defs_i o map (fn x => (x, [Drule.tag_internal]))) | 
| 5698 | 537 | (field_specs @ dest_specs); | 
| 4890 | 538 | |
| 539 | val field_defs = get_defs defs_thy field_specs; | |
| 540 | val dest_defs = get_defs defs_thy dest_specs; | |
| 541 | ||
| 542 | ||
| 5698 | 543 | (* 3rd stage: thms_thy *) | 
| 544 | ||
| 545 | val prove = prove_simp defs_thy; | |
| 5713 | 546 | val prove_std = prove [] (field_defs @ dest_defs @ datatype_simps); | 
| 4890 | 547 | |
| 5713 | 548 | val field_injects = map prove_std inject_props; | 
| 549 | val dest_convs = map prove_std dest_props; | |
| 5698 | 550 | val surj_pairs = map (prove [DatatypePackage.induct_tac "p" 1] | 
| 6092 | 551 | (map Thm.symmetric field_defs @ dest_convs)) surj_props; | 
| 5698 | 552 | |
| 5707 | 553 | fun mk_split th = SplitPairedAll.rule (th RS eq_reflection); | 
| 6092 | 554 | val field_splits = map mk_split surj_pairs; | 
| 4894 | 555 | |
| 4890 | 556 | val thms_thy = | 
| 557 | defs_thy | |
| 6092 | 558 | |> (PureThy.add_thmss o map Thm.no_attributes) | 
| 4890 | 559 |         [("field_defs", field_defs),
 | 
| 560 |           ("dest_defs", dest_defs),
 | |
| 5698 | 561 |           ("dest_convs", dest_convs),
 | 
| 562 |           ("surj_pairs", surj_pairs),
 | |
| 5713 | 563 |           ("field_splits", field_splits)];
 | 
| 4890 | 564 | |
| 5713 | 565 | in (thms_thy, dest_convs, field_injects, field_splits) end; | 
| 4890 | 566 | |
| 567 | ||
| 568 | (* record_definition *) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 569 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 570 | 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 | 571 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 572 | val sign = Theory.sign_of thy; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 573 | val full = Sign.full_name_path sign bname; | 
| 4890 | 574 | val base = Sign.base_name; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 575 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 576 | |
| 4890 | 577 | (* basic components *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 578 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 579 | val alphas = map fst args; | 
| 5197 | 580 | val name = Sign.full_name sign bname; (*not made part of record name space!*) | 
| 4890 | 581 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 582 | val parent_fields = flat (map #fields parents); | 
| 4890 | 583 | val parent_names = map fst parent_fields; | 
| 584 | val parent_types = map snd parent_fields; | |
| 585 | val parent_len = length parent_fields; | |
| 586 | val parent_xs = variantlist (map (base o fst) parent_fields, [moreN]); | |
| 587 | val parent_vars = ListPair.map Free (parent_xs, parent_types); | |
| 4894 | 588 | val parent_named_vars = parent_names ~~ parent_vars; | 
| 4890 | 589 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 590 | val fields = map (apfst full) bfields; | 
| 4890 | 591 | val names = map fst fields; | 
| 592 | val types = map snd fields; | |
| 593 | val len = length fields; | |
| 594 | val xs = variantlist (map fst bfields, moreN :: parent_xs); | |
| 595 | val vars = ListPair.map Free (xs, types); | |
| 4894 | 596 | val named_vars = names ~~ vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 597 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 598 | val all_fields = parent_fields @ fields; | 
| 4890 | 599 | val all_names = parent_names @ names; | 
| 600 | val all_types = parent_types @ types; | |
| 601 | val all_len = parent_len + len; | |
| 602 | val all_xs = parent_xs @ xs; | |
| 603 | val all_vars = parent_vars @ vars; | |
| 4894 | 604 | val all_named_vars = parent_named_vars @ named_vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 605 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 606 | val zeta = variant alphas "'z"; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 607 | val moreT = TFree (zeta, moreS); | 
| 4895 | 608 | val more = Free (moreN, moreT); | 
| 5197 | 609 | val full_moreN = full moreN; | 
| 610 | fun more_part t = mk_more t full_moreN; | |
| 611 | fun more_part_update t x = mk_more_update t (full_moreN, x); | |
| 4894 | 612 | |
| 613 | val parent_more = funpow parent_len mk_snd; | |
| 614 | val idxs = 0 upto (len - 1); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 615 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 616 | val rec_schemeT = mk_recordT (all_fields, moreT); | 
| 4894 | 617 | val rec_scheme = mk_record (all_named_vars, more); | 
| 4890 | 618 |     val r = Free ("r", rec_schemeT);
 | 
| 4894 | 619 | val recT = mk_recordT (all_fields, HOLogic.unitT); | 
| 4867 
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 | |
| 4890 | 622 | (* prepare print translation functions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 623 | |
| 5698 | 624 | val field_tr's = | 
| 625 | print_translation (distinct (flat (map NameSpace.accesses (full_moreN :: names)))); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 626 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 627 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 628 | (* prepare declarations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 629 | |
| 5197 | 630 | val sel_decls = map (mk_selC rec_schemeT) bfields @ | 
| 631 | [mk_moreC rec_schemeT (moreN, moreT)]; | |
| 632 | val update_decls = map (mk_updateC rec_schemeT) bfields @ | |
| 633 | [mk_more_updateC rec_schemeT (moreN, moreT)]; | |
| 4890 | 634 | val make_decls = | 
| 635 | [(mk_makeC rec_schemeT (make_schemeN, all_types @ [moreT])), | |
| 636 | (mk_makeC recT (makeN, all_types))]; | |
| 4867 
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 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 639 | (* prepare definitions *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 640 | |
| 4895 | 641 | (*record (scheme) type abbreviation*) | 
| 4890 | 642 | val recordT_specs = | 
| 643 | [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn), | |
| 644 | (bname, alphas, recT, Syntax.NoSyn)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 645 | |
| 4894 | 646 | (*selectors*) | 
| 4890 | 647 | fun mk_sel_spec (i, c) = | 
| 4894 | 648 | mk_sel r c :== mk_fst (funpow i mk_snd (parent_more r)); | 
| 649 | val sel_specs = | |
| 650 | ListPair.map mk_sel_spec (idxs, names) @ | |
| 651 | [more_part r :== funpow len mk_snd (parent_more r)]; | |
| 4890 | 652 | |
| 653 | (*updates*) | |
| 4894 | 654 | val all_sels = all_names ~~ map (mk_sel r) all_names; | 
| 4890 | 655 | fun mk_upd_spec (i, (c, x)) = | 
| 4894 | 656 | mk_update r (c, x) :== | 
| 657 | mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part r) | |
| 5197 | 658 | val update_specs = | 
| 659 | ListPair.map mk_upd_spec (idxs, named_vars) @ | |
| 660 | [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 | 661 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 662 | (*makes*) | 
| 4890 | 663 | val make_scheme = Const (mk_makeC rec_schemeT (full make_schemeN, all_types @ [moreT])); | 
| 664 | 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 | 665 | val make_specs = | 
| 4894 | 666 | [list_comb (make_scheme, all_vars) $ more :== rec_scheme, | 
| 667 | list_comb (make, all_vars) :== mk_record (all_named_vars, HOLogic.unit)]; | |
| 668 | ||
| 669 | ||
| 670 | (* prepare propositions *) | |
| 671 | ||
| 672 | (*selectors*) | |
| 673 | val sel_props = | |
| 674 | map (fn (c, x) => mk_sel rec_scheme c === x) named_vars @ | |
| 675 | [more_part rec_scheme === more]; | |
| 676 | ||
| 677 | (*updates*) | |
| 678 | fun mk_upd_prop (i, (c, T)) = | |
| 679 | let val x' = Free (variant all_xs (base c ^ "'"), T) in | |
| 680 | mk_update rec_scheme (c, x') === | |
| 681 | mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more) | |
| 682 | end; | |
| 5197 | 683 | val update_props = | 
| 684 | ListPair.map mk_upd_prop (idxs, fields) @ | |
| 685 | let val more' = Free (variant all_xs (moreN ^ "'"), moreT) | |
| 686 | 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 | 687 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 688 | |
| 4890 | 689 | (* 1st stage: fields_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 690 | |
| 5713 | 691 | val (fields_thy, field_simps, field_injects, field_splits) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 692 | thy | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 693 | |> Theory.add_path bname | 
| 4894 | 694 | |> 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 | 695 | |
| 6092 | 696 | val named_splits = map2 (fn (c, th) => (suffix field_typeN c, th)) (names, field_splits); | 
| 5698 | 697 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 698 | |
| 4890 | 699 | (* 2nd stage: defs_thy *) | 
| 700 | ||
| 701 | val defs_thy = | |
| 702 | fields_thy | |
| 703 | |> Theory.parent_path | |
| 5197 | 704 | |> Theory.add_tyabbrs_i recordT_specs (*not made part of record name space!*) | 
| 4890 | 705 | |> Theory.add_path bname | 
| 5197 | 706 | |> Theory.add_trfuns ([], [], field_tr's, []) | 
| 4894 | 707 | |> (Theory.add_consts_i o map Syntax.no_syn) | 
| 708 | (sel_decls @ update_decls @ make_decls) | |
| 6092 | 709 | |> (PureThy.add_defs_i o map (fn x => (x, [Drule.tag_internal]))) | 
| 5212 | 710 | (sel_specs @ update_specs) | 
| 6092 | 711 | |> (PureThy.add_defs_i o map Thm.no_attributes) make_specs; | 
| 4890 | 712 | |
| 713 | val sel_defs = get_defs defs_thy sel_specs; | |
| 714 | val update_defs = get_defs defs_thy update_specs; | |
| 715 | val make_defs = get_defs defs_thy make_specs; | |
| 716 | ||
| 717 | ||
| 718 | (* 3rd stage: thms_thy *) | |
| 719 | ||
| 4895 | 720 | val parent_simps = flat (map #simps parents); | 
| 5698 | 721 | val prove = prove_simp defs_thy []; | 
| 4890 | 722 | |
| 4895 | 723 | val sel_convs = map (prove (parent_simps @ sel_defs @ field_simps)) sel_props; | 
| 724 | val update_convs = map (prove (parent_simps @ update_defs @ sel_convs)) update_props; | |
| 4894 | 725 | |
| 726 | val simps = field_simps @ sel_convs @ update_convs @ make_defs; | |
| 6519 | 727 | val iffs = field_injects; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 728 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 729 | val thms_thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 730 | defs_thy | 
| 6092 | 731 | |> (PureThy.add_thmss o map Thm.no_attributes) | 
| 4895 | 732 |         [("select_defs", sel_defs),
 | 
| 4890 | 733 |           ("update_defs", update_defs),
 | 
| 4894 | 734 |           ("make_defs", make_defs),
 | 
| 4895 | 735 |           ("select_convs", sel_convs),
 | 
| 4894 | 736 |           ("update_convs", update_convs)]
 | 
| 6092 | 737 | |> PureThy.add_thmss | 
| 5707 | 738 |         [(("simps", simps), [Simplifier.simp_add_global]),
 | 
| 6519 | 739 |          (("iffs", iffs), [iff_add_global])];
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 740 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 741 | |
| 4890 | 742 | (* 4th stage: final_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 743 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 744 | val final_thy = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 745 | thms_thy | 
| 4895 | 746 |       |> put_record name {args = args, parent = parent, fields = fields, simps = simps}
 | 
| 5713 | 747 | |> add_record_splits named_splits | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 748 | |> Theory.parent_path; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 749 | |
| 6519 | 750 |   in (final_thy, {simps = simps, iffs = iffs}) end;
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 751 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 752 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 753 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 754 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 755 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 756 | (* prepare arguments *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 757 | |
| 4894 | 758 | (*note: read_raw_typ avoids expanding type abbreviations*) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 759 | fun read_raw_parent sign s = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 760 | (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 | 761 | Type (name, Ts) => (Ts, name) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 762 |   | _ => error ("Bad parent record specification: " ^ quote s));
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 763 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 764 | fun read_typ sign (env, s) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 765 | let | 
| 5060 | 766 | fun def_sort (x, ~1) = assoc (env, x) | 
| 767 | | def_sort _ = None; | |
| 768 | 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 | 769 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 770 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 771 | fun cert_typ sign (env, raw_T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 772 | 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 | 773 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 774 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 775 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 776 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 777 | |
| 4895 | 778 | (*we do all preparations and error checks here, deferring the real | 
| 779 | work to record_definition*) | |
| 4890 | 780 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 781 | 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 | 782 | let | 
| 4970 | 783 | val _ = Theory.requires thy "Record" "record definitions"; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 784 | val sign = Theory.sign_of thy; | 
| 5698 | 785 |     val _ = message ("Defining record " ^ quote bname ^ " ...");
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 786 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 787 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 788 | (* parents *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 789 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 790 | fun prep_inst T = snd (cert_typ sign ([], T)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 791 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 792 | 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 | 793 |       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 | 794 | val parents = add_parents thy (parent, []); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 795 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 796 | val init_env = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 797 | (case parent of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 798 | None => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 799 | | Some (types, _) => foldr Term.add_typ_tfrees (types, [])); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 800 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 801 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 802 | (* fields *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 803 | |
| 4967 | 804 | fun prep_field (env, (c, raw_T)) = | 
| 805 | let val (env', T) = prep_typ sign (env, raw_T) handle ERROR => | |
| 806 |         error ("The error(s) above occured in field " ^ quote c)
 | |
| 807 | in (env', (c, T)) end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 808 | |
| 4967 | 809 | 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 | 810 | val envir_names = map fst envir; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 811 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 812 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 813 | (* args *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 814 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 815 | val defaultS = Sign.defaultS sign; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 816 | 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 | 817 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 818 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 819 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 820 | |
| 4890 | 821 | val name = Sign.full_name sign bname; | 
| 822 | val err_dup_record = | |
| 823 | if is_none (get_record thy name) then [] | |
| 824 | else ["Duplicate definition of record " ^ quote name]; | |
| 825 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 826 | val err_dup_parms = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 827 | (case duplicates params of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 828 | [] => [] | 
| 4890 | 829 | | dups => ["Duplicate parameter(s) " ^ commas dups]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 830 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 831 | val err_extra_frees = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 832 | (case gen_rems (op =) (envir_names, params) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 833 | [] => [] | 
| 4890 | 834 | | extras => ["Extra free type variable(s) " ^ commas extras]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 835 | |
| 4890 | 836 | 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 | 837 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 838 | val err_dup_fields = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 839 | (case duplicates (map fst bfields) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 840 | [] => [] | 
| 4890 | 841 | | dups => ["Duplicate field(s) " ^ commas_quote dups]); | 
| 842 | ||
| 843 | val err_bad_fields = | |
| 844 | if forall (not_equal moreN o fst) bfields then [] | |
| 845 | else ["Illegal field name " ^ quote moreN]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 846 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 847 | val err_dup_sorts = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 848 | (case duplicates envir_names of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 849 | [] => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 850 | | dups => ["Inconsistent sort constraints for " ^ commas dups]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 851 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 852 | val errs = | 
| 4890 | 853 | err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @ | 
| 854 | err_dup_fields @ err_bad_fields @ err_dup_sorts; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 855 | in | 
| 4890 | 856 | if null errs then () else error (cat_lines errs); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 857 | thy |> record_definition (args, bname) parent parents bfields | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 858 | end | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 859 |   handle ERROR => error ("Failed to define record " ^ quote bname);
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 860 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 861 | 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 | 862 | 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 | 863 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 864 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 865 | |
| 6358 | 866 | (** package setup **) | 
| 867 | ||
| 868 | (* setup theory *) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 869 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 870 | val setup = | 
| 5006 | 871 | [RecordsData.init, | 
| 5698 | 872 | Theory.add_trfuns ([], parse_translation, [], []), | 
| 6358 | 873 | Method.add_methods [record_split_method], | 
| 5698 | 874 | add_wrapper record_split_wrapper]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 875 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 876 | |
| 6358 | 877 | (* outer syntax *) | 
| 878 | ||
| 6723 | 879 | local structure P = OuterParse and K = OuterSyntax.Keyword in | 
| 6358 | 880 | |
| 881 | val record_decl = | |
| 6723 | 882 | P.type_args -- P.name -- (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") | 
| 6729 | 883 | -- Scan.repeat1 (P.name -- (P.$$$ "::" |-- P.typ) --| P.marg_comment)); | 
| 6358 | 884 | |
| 885 | val recordP = | |
| 6723 | 886 | OuterSyntax.command "record" "define extensible record" K.thy_decl | 
| 6519 | 887 | (record_decl >> (fn (x, (y, z)) => Toplevel.theory (#1 o add_record x y z))); | 
| 6358 | 888 | |
| 889 | val _ = OuterSyntax.add_parsers [recordP]; | |
| 890 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 891 | end; | 
| 5698 | 892 | |
| 893 | ||
| 6384 | 894 | end; | 
| 895 | ||
| 5698 | 896 | structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage; | 
| 897 | open BasicRecordPackage; |