| author | berghofe | 
| Mon, 21 Oct 2002 17:07:58 +0200 | |
| changeset 13660 | e36798726ca4 | 
| parent 13462 | 56610e2ba220 | 
| child 13904 | c13e6e218a69 | 
| 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 | 
| 9230 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 5 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 6 | Extensible records with structural subtyping in HOL. | 
| 5698 | 7 | *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 8 | |
| 5698 | 9 | signature BASIC_RECORD_PACKAGE = | 
| 10 | sig | |
| 7178 | 11 | val record_simproc: simproc | 
| 5698 | 12 | val record_split_tac: int -> tactic | 
| 5713 | 13 | val record_split_name: string | 
| 5698 | 14 | val record_split_wrapper: string * wrapper | 
| 15 | end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 16 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 17 | signature RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 18 | sig | 
| 5698 | 19 | include BASIC_RECORD_PACKAGE | 
| 20 | val quiet_mode: bool ref | |
| 8574 | 21 | val updateN: string | 
| 4890 | 22 | val mk_fieldT: (string * typ) * typ -> typ | 
| 23 | val dest_fieldT: typ -> (string * typ) * typ | |
| 24 | val mk_field: (string * term) * term -> term | |
| 25 | val mk_fst: term -> term | |
| 26 | val mk_snd: term -> term | |
| 27 | val mk_recordT: (string * typ) list * typ -> typ | |
| 28 | val dest_recordT: typ -> (string * typ) list * typ | |
| 29 | val mk_record: (string * term) list * term -> term | |
| 30 | val mk_sel: term -> string -> term | |
| 31 | val mk_update: term -> string * term -> term | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 32 | val print_records: theory -> unit | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 33 | val add_record: (string list * bstring) -> string option | 
| 12506 | 34 |     -> (bstring * string * mixfix) 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 add_record_i: (string list * bstring) -> (typ list * string) option | 
| 12506 | 36 |     -> (bstring * typ * mixfix) list -> theory -> theory * {simps: thm list, iffs: thm list}
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 37 | val setup: (theory -> theory) list | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 38 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 39 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 40 | structure RecordPackage: RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 41 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 42 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 43 | |
| 11832 | 44 | (*** theory context references ***) | 
| 45 | ||
| 46 | val product_typeN = "Record.product_type"; | |
| 47 | ||
| 13419 | 48 | val product_type_intro = thm "product_type.intro"; | 
| 13413 | 49 | val product_type_inject = thm "product_type.inject"; | 
| 50 | val product_type_conv1 = thm "product_type.conv1"; | |
| 51 | val product_type_conv2 = thm "product_type.conv2"; | |
| 52 | val product_type_induct = thm "product_type.induct"; | |
| 53 | val product_type_cases = thm "product_type.cases"; | |
| 54 | val product_type_split_paired_all = thm "product_type.split_paired_all"; | |
| 11832 | 55 | |
| 56 | ||
| 57 | ||
| 4894 | 58 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 59 | |
| 5698 | 60 | (* messages *) | 
| 61 | ||
| 62 | val quiet_mode = ref false; | |
| 63 | fun message s = if ! quiet_mode then () else writeln s; | |
| 64 | ||
| 65 | ||
| 12255 | 66 | (* syntax *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 67 | |
| 12247 | 68 | fun prune n xs = Library.drop (n, xs); | 
| 11832 | 69 | fun prefix_base s = NameSpace.map_base (fn bname => s ^ bname); | 
| 70 | ||
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 71 | val Trueprop = HOLogic.mk_Trueprop; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 72 | fun All xs t = Term.list_all_free (xs, t); | 
| 4894 | 73 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 74 | infix 9 $$; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 75 | infix 0 :== ===; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 76 | infixr 0 ==>; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 77 | |
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 78 | val (op $$) = Term.list_comb; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 79 | val (op :==) = Logic.mk_defpair; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 80 | val (op ===) = Trueprop o HOLogic.mk_eq; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 81 | val (op ==>) = Logic.mk_implies; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 82 | |
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 83 | |
| 12255 | 84 | (* attributes *) | 
| 85 | ||
| 12302 | 86 | fun case_names_fields x = RuleCases.case_names ["fields"] x; | 
| 12255 | 87 | fun induct_type_global name = [case_names_fields, InductAttrib.induct_type_global name]; | 
| 88 | fun cases_type_global name = [case_names_fields, InductAttrib.cases_type_global name]; | |
| 89 | ||
| 90 | ||
| 91 | (* tactics *) | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 92 | |
| 11967 | 93 | fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps)); | 
| 4895 | 94 | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 95 | fun try_param_tac x s rule i st = | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 96 | res_inst_tac [(x, (case Tactic.innermost_params i st of [] => s | (p, _) :: _ => p))] rule i st; | 
| 4894 | 97 | |
| 98 | ||
| 99 | ||
| 12449 | 100 | (*** code generator data ***) | 
| 101 | ||
| 102 | val [prod_code, fst_code, snd_code] = | |
| 103 | map (Codegen.parse_mixfix (K (Bound 0))) ["(_,/ _)", "fst", "snd"]; | |
| 104 | val prodT_code = Codegen.parse_mixfix (K dummyT) "(_ */ _)"; | |
| 105 | ||
| 106 | ||
| 107 | ||
| 4894 | 108 | (*** syntax operations ***) | 
| 109 | ||
| 110 | (** name components **) | |
| 111 | ||
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 112 | val rN = "r"; | 
| 4894 | 113 | val moreN = "more"; | 
| 114 | val schemeN = "_scheme"; | |
| 11832 | 115 | val field_typeN = "_field_type"; | 
| 4894 | 116 | val fieldN = "_field"; | 
| 5698 | 117 | val fstN = "_val"; | 
| 118 | val sndN = "_more"; | |
| 4894 | 119 | val updateN = "_update"; | 
| 120 | val makeN = "make"; | |
| 12265 | 121 | val fieldsN = "fields"; | 
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 122 | val extendN = "extend"; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 123 | val truncateN = "truncate"; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 124 | |
| 4890 | 125 | |
| 11832 | 126 | (*see typedef_package.ML*) | 
| 127 | val RepN = "Rep_"; | |
| 128 | val AbsN = "Abs_"; | |
| 5698 | 129 | |
| 130 | ||
| 5713 | 131 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 132 | (** tuple operations **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 133 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 134 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 135 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 136 | 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 | 137 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 138 | 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 | 139 | (case try (unsuffix field_typeN) c_field_type of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 140 |         None => raise TYPE ("dest_fieldT", [typ], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 141 | | Some c => ((c, T), U)) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 142 |   | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
 | 
| 
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 | |
| 11832 | 145 | (* morphisms *) | 
| 146 | ||
| 147 | fun mk_Rep U (c, T) = | |
| 148 | Const (suffix field_typeN (prefix_base RepN c), | |
| 149 | mk_fieldT ((c, T), U) --> HOLogic.mk_prodT (T, U)); | |
| 150 | ||
| 151 | fun mk_Abs U (c, T) = | |
| 152 | Const (suffix field_typeN (prefix_base AbsN c), | |
| 153 | HOLogic.mk_prodT (T, U) --> mk_fieldT ((c, T), U)); | |
| 11833 | 154 | |
| 11832 | 155 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 156 | (* constructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 157 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 158 | 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 | 159 | |
| 11832 | 160 | fun mk_field ((c, t), u) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 161 | let val T = fastype_of t and U = fastype_of u | 
| 11832 | 162 | in Const (suffix fieldN c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end; | 
| 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 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 165 | (* destructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 166 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 167 | 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 | 168 | 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 | 169 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 170 | fun dest_field fst_or_snd p = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 171 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 172 | val pT = fastype_of p; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 173 | val ((c, T), U) = dest_fieldT pT; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 174 | 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 | 175 | in Const (suffix destN c, pT --> destT) $ p end; | 
| 
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 | val mk_fst = dest_field true; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 178 | val mk_snd = dest_field false; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 179 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 180 | |
| 
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 | (** record operations **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 183 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 184 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 185 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 186 | val mk_recordT = foldr mk_fieldT; | 
| 
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 | fun dest_recordT T = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 189 | (case try dest_fieldT T of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 190 | None => ([], T) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 191 | | 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 | 192 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 193 | fun find_fieldT c rT = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 194 | (case assoc (fst (dest_recordT rT), c) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 195 |     None => raise TYPE ("find_field: " ^ c, [rT], [])
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 196 | | Some T => T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 197 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 198 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 199 | (* constructors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 200 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 201 | val mk_record = foldr mk_field; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 202 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 203 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 204 | (* selectors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 205 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 206 | fun mk_selC rT (c, T) = (c, rT --> T); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 207 | |
| 4890 | 208 | fun mk_sel r c = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 209 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 210 | 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 | 211 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 212 | fun mk_named_sels names r = names ~~ map (mk_sel r) names; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 213 | |
| 4894 | 214 | val mk_moreC = mk_selC; | 
| 215 | ||
| 216 | fun mk_more r c = | |
| 217 | let val rT = fastype_of r | |
| 218 | in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end; | |
| 219 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 220 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 221 | (* updates *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 222 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 223 | 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 | 224 | |
| 4890 | 225 | fun mk_update r (c, x) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 226 | let val rT = fastype_of r | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 227 | 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 | 228 | |
| 5197 | 229 | val mk_more_updateC = mk_updateC; | 
| 230 | ||
| 231 | fun mk_more_update r (c, x) = | |
| 232 | let val rT = fastype_of r | |
| 233 | in Const (mk_more_updateC rT (c, snd (dest_recordT rT))) $ x $ r end; | |
| 234 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 235 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 236 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 237 | (** concrete syntax for records **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 238 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 239 | (* parse translations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 240 | |
| 11473 
4546d8d39221
fix problem with user translations by making field names appear as consts;
 wenzelm parents: 
10008diff
changeset | 241 | fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) = | 
| 5197 | 242 | if c = mark then Syntax.const (suffix sfx name) $ arg | 
| 243 |       else raise TERM ("gen_field_tr: " ^ mark, [t])
 | |
| 244 |   | 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 | 245 | |
| 5197 | 246 | fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = | 
| 247 | if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u | |
| 5201 | 248 | else [gen_field_tr mark sfx tm] | 
| 249 | | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; | |
| 5197 | 250 | |
| 251 | fun gen_record_tr sep mark sfx unit [t] = foldr (op $) (gen_fields_tr sep mark sfx t, unit) | |
| 5201 | 252 |   | gen_record_tr _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
 | 
| 5197 | 253 | |
| 254 | fun gen_record_scheme_tr sep mark sfx [t, more] = foldr (op $) (gen_fields_tr sep mark sfx t, more) | |
| 5201 | 255 |   | gen_record_scheme_tr _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
 | 
| 5197 | 256 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 257 | |
| 5197 | 258 | val record_type_tr = gen_record_tr "_field_types" "_field_type" field_typeN (Syntax.const "unit"); | 
| 259 | val record_type_scheme_tr = gen_record_scheme_tr "_field_types" "_field_type" field_typeN; | |
| 260 | ||
| 261 | val record_tr = gen_record_tr "_fields" "_field" fieldN HOLogic.unit; | |
| 262 | 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 | 263 | |
| 5197 | 264 | fun record_update_tr [t, u] = | 
| 265 | foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) | |
| 266 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | |
| 267 | ||
| 268 | ||
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 269 | fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 270 | | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts | 
| 11833 | 271 |   | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
 | 
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 272 | (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts | 
| 11833 | 273 |   | update_name_tr ts = raise TERM ("update_name_tr", ts);
 | 
| 274 | ||
| 275 | ||
| 5197 | 276 | val parse_translation = | 
| 277 |  [("_record_type", record_type_tr),
 | |
| 278 |   ("_record_type_scheme", record_type_scheme_tr),
 | |
| 279 |   ("_record", record_tr),
 | |
| 280 |   ("_record_scheme", record_scheme_tr),
 | |
| 11833 | 281 |   ("_record_update", record_update_tr),
 | 
| 282 |   ("_update_name", update_name_tr)];
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 283 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 284 | |
| 4890 | 285 | (* print translations *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 286 | |
| 5197 | 287 | fun gen_fields_tr' mark sfx (tm as Const (name_field, _) $ t $ u) = | 
| 288 | (case try (unsuffix sfx) name_field of | |
| 289 | Some name => | |
| 290 | apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_fields_tr' mark sfx u) | |
| 291 | | None => ([], tm)) | |
| 292 | | gen_fields_tr' _ _ tm = ([], tm); | |
| 293 | ||
| 294 | 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 | 295 | let | 
| 5197 | 296 | val (ts, u) = gen_fields_tr' mark sfx tm; | 
| 297 | 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 | 298 | in | 
| 5197 | 299 | if is_unit u then Syntax.const record $ t' | 
| 300 | else Syntax.const record_scheme $ t' $ u | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 301 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 302 | |
| 5197 | 303 | |
| 304 | val record_type_tr' = | |
| 305 | gen_record_tr' "_field_types" "_field_type" field_typeN | |
| 306 |     (fn Const ("unit", _) => true | _ => false) "_record_type" "_record_type_scheme";
 | |
| 307 | ||
| 308 | val record_tr' = | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 309 | gen_record_tr' "_fields" "_field" fieldN | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 310 |     (fn Const ("Unity", _) => true | _ => false) "_record" "_record_scheme";
 | 
| 5197 | 311 | |
| 312 | fun record_update_tr' tm = | |
| 313 | let val (ts, u) = gen_fields_tr' "_update" updateN tm in | |
| 314 | Syntax.const "_record_update" $ u $ | |
| 315 | foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) | |
| 316 | end; | |
| 317 | ||
| 318 | ||
| 5201 | 319 | fun gen_field_tr' sfx tr' name = | 
| 320 | let val name_sfx = suffix sfx name | |
| 321 | in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; | |
| 322 | ||
| 5197 | 323 | fun print_translation names = | 
| 324 | map (gen_field_tr' field_typeN record_type_tr') names @ | |
| 325 | map (gen_field_tr' fieldN record_tr') names @ | |
| 326 | map (gen_field_tr' updateN record_update_tr') names; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 327 | |
| 
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 | (*** extend theory by record definition ***) | 
| 
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 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 333 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 334 | (* type record_info and parent_info *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 335 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 336 | type record_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 337 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 338 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 339 | fields: (string * typ) list, | 
| 12247 | 340 | field_inducts: thm list, | 
| 341 | field_cases: thm list, | |
| 342 | simps: thm list}; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 343 | |
| 12247 | 344 | fun make_record_info args parent fields field_inducts field_cases simps = | 
| 345 |  {args = args, parent = parent, fields = fields, field_inducts = field_inducts,
 | |
| 346 | field_cases = field_cases, simps = simps}: record_info; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 347 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 348 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 349 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 350 | fields: (string * typ) list, | 
| 12247 | 351 | field_inducts: thm list, | 
| 352 | field_cases: thm list, | |
| 353 | simps: thm list}; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 354 | |
| 12247 | 355 | fun make_parent_info name fields field_inducts field_cases simps = | 
| 356 |  {name = name, fields = fields, field_inducts = field_inducts,
 | |
| 357 | field_cases = field_cases, simps = simps}: parent_info; | |
| 4867 
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 | |
| 5052 | 360 | (* data kind 'HOL/records' *) | 
| 5001 | 361 | |
| 7178 | 362 | type record_data = | 
| 363 |  {records: record_info Symtab.table,
 | |
| 364 | sel_upd: | |
| 365 |    {selectors: unit Symtab.table,
 | |
| 366 | updates: string Symtab.table, | |
| 367 | simpset: Simplifier.simpset}, | |
| 368 | field_splits: | |
| 369 |    {fields: unit Symtab.table,
 | |
| 370 | simpset: Simplifier.simpset}}; | |
| 371 | ||
| 372 | fun make_record_data records sel_upd field_splits = | |
| 373 |  {records = records, sel_upd = sel_upd, field_splits = field_splits}: record_data;
 | |
| 374 | ||
| 5006 | 375 | structure RecordsArgs = | 
| 376 | struct | |
| 377 | val name = "HOL/records"; | |
| 7178 | 378 | type T = record_data; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 379 | |
| 7178 | 380 | val empty = | 
| 381 | make_record_data Symtab.empty | |
| 382 |       {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss}
 | |
| 383 |       {fields = Symtab.empty, simpset = HOL_basic_ss};
 | |
| 384 | ||
| 6556 | 385 | val copy = I; | 
| 5006 | 386 | val prep_ext = I; | 
| 7178 | 387 | fun merge | 
| 388 |    ({records = recs1,
 | |
| 389 |      sel_upd = {selectors = sels1, updates = upds1, simpset = ss1},
 | |
| 390 |      field_splits = {fields = flds1, simpset = fld_ss1}},
 | |
| 391 |     {records = recs2,
 | |
| 392 |      sel_upd = {selectors = sels2, updates = upds2, simpset = ss2},
 | |
| 393 |      field_splits = {fields = flds2, simpset = fld_ss2}}) =
 | |
| 394 | make_record_data | |
| 395 | (Symtab.merge (K true) (recs1, recs2)) | |
| 396 |       {selectors = Symtab.merge (K true) (sels1, sels2),
 | |
| 397 | updates = Symtab.merge (K true) (upds1, upds2), | |
| 398 | simpset = Simplifier.merge_ss (ss1, ss2)} | |
| 399 |       {fields = Symtab.merge (K true) (flds1, flds2),
 | |
| 400 | simpset = Simplifier.merge_ss (fld_ss1, fld_ss2)}; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 401 | |
| 7178 | 402 |   fun print sg ({records = recs, ...}: record_data) =
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 403 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 404 | val prt_typ = Sign.pretty_typ sg; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 405 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 406 | fun pretty_parent None = [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 407 | | pretty_parent (Some (Ts, name)) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 408 | [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 409 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 410 | fun pretty_field (c, T) = Pretty.block | 
| 12129 | 411 | [Pretty.str (Sign.cond_extern sg Sign.constK c), Pretty.str " ::", | 
| 412 | Pretty.brk 1, Pretty.quote (prt_typ T)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 413 | |
| 12247 | 414 |       fun pretty_record (name, {args, parent, fields, ...}: record_info) =
 | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 415 | Pretty.block (Pretty.fbreaks (Pretty.block | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 416 | [prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 417 | pretty_parent parent @ map pretty_field fields)); | 
| 12129 | 418 | in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 419 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 420 | |
| 5006 | 421 | structure RecordsData = TheoryDataFun(RecordsArgs); | 
| 422 | val print_records = RecordsData.print; | |
| 423 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 424 | |
| 7178 | 425 | (* access 'records' *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 426 | |
| 7178 | 427 | fun get_record thy name = Symtab.lookup (#records (RecordsData.get thy), name); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 428 | |
| 4890 | 429 | fun put_record name info thy = | 
| 7178 | 430 | let | 
| 431 |     val {records, sel_upd, field_splits} = RecordsData.get thy;
 | |
| 432 | val data = make_record_data (Symtab.update ((name, info), records)) sel_upd field_splits; | |
| 433 | in RecordsData.put data thy end; | |
| 434 | ||
| 435 | ||
| 436 | (* access 'sel_upd' *) | |
| 437 | ||
| 438 | fun get_sel_upd sg = #sel_upd (RecordsData.get_sg sg); | |
| 439 | ||
| 440 | fun get_selectors sg name = Symtab.lookup (#selectors (get_sel_upd sg), name); | |
| 441 | fun get_updates sg name = Symtab.lookup (#updates (get_sel_upd sg), name); | |
| 442 | fun get_simpset sg = #simpset (get_sel_upd sg); | |
| 443 | ||
| 444 | ||
| 445 | fun put_sel_upd names simps thy = | |
| 446 | let | |
| 447 | val sels = map (rpair ()) names; | |
| 448 | val upds = map (suffix updateN) names ~~ names; | |
| 449 | ||
| 450 |     val {records, sel_upd = {selectors, updates, simpset}, field_splits} = RecordsData.get thy;
 | |
| 451 | val data = make_record_data records | |
| 452 |       {selectors = Symtab.extend (selectors, sels),
 | |
| 453 | updates = Symtab.extend (updates, upds), | |
| 454 | simpset = Simplifier.addsimps (simpset, simps)} | |
| 455 | field_splits; | |
| 456 | in RecordsData.put data thy end; | |
| 457 | ||
| 458 | ||
| 459 | (* access 'field_splits' *) | |
| 5698 | 460 | |
| 461 | fun add_record_splits splits thy = | |
| 462 | let | |
| 7178 | 463 |     val {records, sel_upd, field_splits = {fields, simpset}} = RecordsData.get thy;
 | 
| 464 | val flds = map (rpair () o fst) splits; | |
| 465 | val simps = map snd splits; | |
| 466 | val data = make_record_data records sel_upd | |
| 467 |       {fields = Symtab.extend (fields, flds), simpset = Simplifier.addsimps (simpset, simps)};
 | |
| 468 | in RecordsData.put data thy end; | |
| 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 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 471 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 472 | |
| 12247 | 473 | fun add_parents thy None parents = parents | 
| 474 | | add_parents thy (Some (types, name)) parents = | |
| 475 | let | |
| 476 | val sign = Theory.sign_of thy; | |
| 477 | fun err msg = error (msg ^ " parent record " ^ quote name); | |
| 12255 | 478 | |
| 12247 | 479 |         val {args, parent, fields, field_inducts, field_cases, simps} =
 | 
| 480 | (case get_record thy name of Some info => info | None => err "Unknown"); | |
| 481 | val _ = if length types <> length args then err "Bad number of arguments for" else (); | |
| 12255 | 482 | |
| 12247 | 483 | fun bad_inst ((x, S), T) = | 
| 484 | if Sign.of_sort sign (T, S) then None else Some x | |
| 485 | val bads = mapfilter bad_inst (args ~~ types); | |
| 12255 | 486 | |
| 12247 | 487 | val inst = map fst args ~~ types; | 
| 488 | val subst = Term.map_type_tfree (fn (x, _) => the (assoc (inst, x))); | |
| 489 | val parent' = apsome (apfst (map subst)) parent; | |
| 490 | val fields' = map (apsnd subst) fields; | |
| 491 | in | |
| 12255 | 492 | conditional (not (null bads)) (fn () => | 
| 493 |           err ("Ill-sorted instantiation of " ^ commas bads ^ " in"));
 | |
| 494 | add_parents thy parent' | |
| 12247 | 495 | (make_parent_info name fields' field_inducts field_cases simps :: parents) | 
| 496 | end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 497 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 498 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 499 | |
| 7178 | 500 | (** record simproc **) | 
| 501 | ||
| 13462 | 502 | val record_simproc = | 
| 503 | Simplifier.simproc (Theory.sign_of HOL.thy) "record_simp" ["s (u k r)"] | |
| 504 | (fn sg => fn _ => fn t => | |
| 505 | (case t of (sel as Const (s, _)) $ ((upd as Const (u, _)) $ k $ r) => | |
| 506 | (case get_selectors sg s of Some () => | |
| 507 | (case get_updates sg u of Some u_name => | |
| 508 | let | |
| 509 | fun mk_free x t = Free (x, fastype_of t); | |
| 510 | val k' = mk_free "k" k; | |
| 511 | val r' = mk_free "r" r; | |
| 512 | val t' = sel $ (upd $ k' $ r'); | |
| 513 | fun prove prop = | |
| 514 | Tactic.prove sg ["k", "r"] [] prop (K (simp_all_tac (get_simpset sg) [])); | |
| 515 | in | |
| 516 | if u_name = s then Some (prove (Logic.mk_equals (t', k'))) | |
| 517 | else Some (prove (Logic.mk_equals (t', sel $ r'))) | |
| 518 | end | |
| 519 | | None => None) | |
| 520 | | None => None) | |
| 521 | | _ => None)); | |
| 7178 | 522 | |
| 523 | ||
| 524 | ||
| 5698 | 525 | (** record field splitting **) | 
| 526 | ||
| 6358 | 527 | (* tactic *) | 
| 528 | ||
| 5698 | 529 | fun record_split_tac i st = | 
| 530 | let | |
| 7178 | 531 |     val {field_splits = {fields, simpset}, ...} = RecordsData.get_sg (Thm.sign_of_thm st);
 | 
| 5698 | 532 | |
| 7178 | 533 | fun is_fieldT (_, Type (a, [_, _])) = is_some (Symtab.lookup (fields, a)) | 
| 5698 | 534 | | is_fieldT _ = false; | 
| 535 | val params = Logic.strip_params (Library.nth_elem (i - 1, Thm.prems_of st)); | |
| 536 | in | |
| 7178 | 537 | if exists is_fieldT params then Simplifier.full_simp_tac simpset i st | 
| 5698 | 538 | else Seq.empty | 
| 539 | end handle Library.LIST _ => Seq.empty; | |
| 540 | ||
| 6358 | 541 | |
| 542 | (* wrapper *) | |
| 543 | ||
| 5707 | 544 | val record_split_name = "record_split_tac"; | 
| 545 | val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac); | |
| 5698 | 546 | |
| 547 | ||
| 6358 | 548 | (* method *) | 
| 549 | ||
| 550 | val record_split_method = | |
| 9705 | 551 |   ("record_split", Method.no_args (Method.SIMPLE_METHOD' HEADGOAL record_split_tac),
 | 
| 6358 | 552 | "split record fields"); | 
| 553 | ||
| 554 | ||
| 12255 | 555 | |
| 4890 | 556 | (** internal theory extenders **) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 557 | |
| 11832 | 558 | (* field_typedefs *) | 
| 5698 | 559 | |
| 11832 | 560 | fun field_typedefs zeta moreT names theory = | 
| 5698 | 561 | let | 
| 11832 | 562 | val alpha = "'a"; | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12311diff
changeset | 563 | val aT = TFree (alpha, HOLogic.typeS); | 
| 11832 | 564 | val UNIV = HOLogic.mk_UNIV (HOLogic.mk_prodT (aT, moreT)); | 
| 5698 | 565 | |
| 11832 | 566 | fun type_def (thy, name) = | 
| 11940 | 567 |       let val (thy', {type_definition, set_def = Some def, ...}) =
 | 
| 568 | thy |> setmp TypedefPackage.quiet_mode true | |
| 11832 | 569 | (TypedefPackage.add_typedef_i true None | 
| 570 | (suffix field_typeN (Sign.base_name name), [alpha, zeta], Syntax.NoSyn) UNIV None | |
| 571 | (Tactic.rtac UNIV_witness 1)) | |
| 572 | in (thy', Tactic.rewrite_rule [def] type_definition) end | |
| 573 | in foldl_map type_def (theory, names) end; | |
| 5698 | 574 | |
| 575 | ||
| 4894 | 576 | (* field_definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 577 | |
| 11832 | 578 | fun field_definitions fields names xs alphas zeta moreT more vars named_vars thy = | 
| 4890 | 579 | let | 
| 5698 | 580 | val sign = Theory.sign_of thy; | 
| 4890 | 581 | val base = Sign.base_name; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 582 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12311diff
changeset | 583 | val xT = TFree (variant alphas "'x", HOLogic.typeS); | 
| 11832 | 584 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 585 | |
| 4890 | 586 | (* prepare declarations and definitions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 587 | |
| 5698 | 588 | (*field constructors*) | 
| 4890 | 589 | val field_decls = map (mk_fieldC moreT) fields; | 
| 590 | ||
| 11832 | 591 | fun mk_field_spec ((c, T), v) = | 
| 592 | Term.head_of (mk_field ((c, v), more)) :== | |
| 593 | lambda v (lambda more (mk_Abs moreT (c, T) $ (HOLogic.mk_prod (v, more)))); | |
| 594 | val field_specs = map mk_field_spec (fields ~~ vars); | |
| 4890 | 595 | |
| 596 | (*field destructors*) | |
| 5698 | 597 | val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields; | 
| 598 | ||
| 11832 | 599 | fun mk_dest_spec dest sel (c, T) = | 
| 5698 | 600 |       let val p = Free ("p", mk_fieldT ((c, T), moreT));
 | 
| 11832 | 601 | in Term.head_of (dest p) :== lambda p (sel (mk_Rep moreT (c, T) $ p)) end; | 
| 602 | val dest_specs1 = map (mk_dest_spec mk_fst HOLogic.mk_fst) fields; | |
| 603 | val dest_specs2 = map (mk_dest_spec mk_snd HOLogic.mk_snd) fields; | |
| 4890 | 604 | |
| 605 | ||
| 11832 | 606 | (* 1st stage: defs_thy *) | 
| 5713 | 607 | |
| 11832 | 608 | val (defs_thy, (((typedefs, field_defs), dest_defs1), dest_defs2)) = | 
| 609 | thy | |
| 610 | |> field_typedefs zeta moreT names | |
| 611 | |>> (Theory.add_consts_i o map (Syntax.no_syn o apfst base)) (field_decls @ dest_decls) | |
| 612 | |>>> (PureThy.add_defs_i false o map Thm.no_attributes) field_specs | |
| 613 | |>>> (PureThy.add_defs_i false o map Thm.no_attributes) dest_specs1 | |
| 614 | |>>> (PureThy.add_defs_i false o map Thm.no_attributes) dest_specs2; | |
| 4890 | 615 | |
| 13419 | 616 | val prod_types = map (fn (((a, b), c), d) => product_type_intro OF [a, b, c, d]) | 
| 11832 | 617 | (typedefs ~~ field_defs ~~ dest_defs1 ~~ dest_defs2); | 
| 4890 | 618 | |
| 5698 | 619 | |
| 11832 | 620 | (* 2nd stage: thms_thy *) | 
| 5698 | 621 | |
| 11832 | 622 | fun make th = map (fn prod_type => Drule.standard (th OF [prod_type])) prod_types; | 
| 4890 | 623 | |
| 11940 | 624 | val dest_convs = make product_type_conv1 @ make product_type_conv2; | 
| 11832 | 625 | val field_injects = make product_type_inject; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 626 | val field_inducts = make product_type_induct; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 627 | val field_cases = make product_type_cases; | 
| 11832 | 628 | val field_splits = make product_type_split_paired_all; | 
| 4894 | 629 | |
| 11940 | 630 | val (thms_thy, [field_defs', dest_defs', dest_convs', field_injects', | 
| 631 | field_splits', field_inducts', field_cases']) = defs_thy | |
| 12449 | 632 | |> Codegen.assoc_consts_i (flat (map (fn (s, _) => | 
| 633 | [(suffix fieldN s, None, prod_code), | |
| 634 | (suffix fstN s, None, fst_code), | |
| 635 | (suffix sndN s, None, snd_code)]) fields)) | |
| 636 | |> Codegen.assoc_types (map (fn (s, _) => | |
| 637 | (suffix field_typeN s, prodT_code)) fields) | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 638 | |> (PureThy.add_thmss o map Thm.no_attributes) | 
| 11940 | 639 |        [("field_defs", field_defs),
 | 
| 640 |         ("dest_defs", dest_defs1 @ dest_defs2),
 | |
| 641 |         ("dest_convs", dest_convs),
 | |
| 642 |         ("field_injects", field_injects),
 | |
| 643 |         ("field_splits", field_splits),
 | |
| 644 |         ("field_inducts", field_inducts),
 | |
| 645 |         ("field_cases", field_cases)];
 | |
| 4890 | 646 | |
| 11940 | 647 | in (thms_thy, dest_convs', field_injects', field_splits', field_inducts', field_cases') end; | 
| 4890 | 648 | |
| 649 | ||
| 650 | (* record_definition *) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 651 | |
| 12506 | 652 | fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 653 | let | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 654 | val sign = Theory.sign_of thy; | 
| 12247 | 655 | |
| 656 | val alphas = map fst args; | |
| 657 | val name = Sign.full_name sign bname; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 658 | val full = Sign.full_name_path sign bname; | 
| 4890 | 659 | val base = Sign.base_name; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 660 | |
| 12506 | 661 | val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields); | 
| 662 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 663 | |
| 4890 | 664 | (* basic components *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 665 | |
| 12247 | 666 | val ancestry = map (length o flat o map #fields) (Library.prefixes1 parents); | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 667 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 668 | val parent_fields = flat (map #fields parents); | 
| 4890 | 669 | val parent_names = map fst parent_fields; | 
| 670 | val parent_types = map snd parent_fields; | |
| 671 | val parent_len = length parent_fields; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 672 | val parent_xs = variantlist (map (base o fst) parent_fields, [moreN, rN]); | 
| 4890 | 673 | val parent_vars = ListPair.map Free (parent_xs, parent_types); | 
| 4894 | 674 | val parent_named_vars = parent_names ~~ parent_vars; | 
| 4890 | 675 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 676 | val fields = map (apfst full) bfields; | 
| 4890 | 677 | val names = map fst fields; | 
| 678 | val types = map snd fields; | |
| 679 | val len = length fields; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 680 | val xs = variantlist (map fst bfields, moreN :: rN :: parent_xs); | 
| 4890 | 681 | val vars = ListPair.map Free (xs, types); | 
| 4894 | 682 | val named_vars = names ~~ vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 683 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 684 | val all_fields = parent_fields @ fields; | 
| 4890 | 685 | val all_names = parent_names @ names; | 
| 686 | val all_types = parent_types @ types; | |
| 687 | val all_len = parent_len + len; | |
| 688 | val all_xs = parent_xs @ xs; | |
| 689 | val all_vars = parent_vars @ vars; | |
| 4894 | 690 | val all_named_vars = parent_named_vars @ named_vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 691 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 692 | val zeta = variant alphas "'z"; | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12311diff
changeset | 693 | val moreT = TFree (zeta, HOLogic.typeS); | 
| 4895 | 694 | val more = Free (moreN, moreT); | 
| 5197 | 695 | val full_moreN = full moreN; | 
| 696 | fun more_part t = mk_more t full_moreN; | |
| 697 | fun more_part_update t x = mk_more_update t (full_moreN, x); | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 698 | val all_types_more = all_types @ [moreT]; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 699 | val all_xs_more = all_xs @ [moreN]; | 
| 4894 | 700 | |
| 701 | val parent_more = funpow parent_len mk_snd; | |
| 702 | val idxs = 0 upto (len - 1); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 703 | |
| 12265 | 704 | val fieldsT = mk_recordT (fields, HOLogic.unitT); | 
| 12247 | 705 | fun rec_schemeT n = mk_recordT (prune n all_fields, moreT); | 
| 706 | fun rec_scheme n = mk_record (prune n all_named_vars, more); | |
| 707 | fun recT n = mk_recordT (prune n all_fields, HOLogic.unitT); | |
| 12255 | 708 | fun rec_ n = mk_record (prune n all_named_vars, HOLogic.unit); | 
| 12247 | 709 | fun r_scheme n = Free (rN, rec_schemeT n); | 
| 710 | fun r n = Free (rN, recT n); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 711 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 712 | |
| 4890 | 713 | (* prepare print translation functions *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 714 | |
| 5698 | 715 | val field_tr's = | 
| 13333 | 716 | 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 | 717 | |
| 
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 | (* prepare declarations *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 720 | |
| 12247 | 721 | val sel_decls = map (mk_selC (rec_schemeT 0)) bfields @ | 
| 722 | [mk_moreC (rec_schemeT 0) (moreN, moreT)]; | |
| 723 | val update_decls = map (mk_updateC (rec_schemeT 0)) bfields @ | |
| 724 | [mk_more_updateC (rec_schemeT 0) (moreN, moreT)]; | |
| 12265 | 725 | val make_decl = (makeN, all_types ---> recT 0); | 
| 726 | val fields_decl = (fieldsN, types ---> fieldsT); | |
| 12247 | 727 | val extend_decl = (extendN, recT 0 --> moreT --> rec_schemeT 0); | 
| 728 | val truncate_decl = (truncateN, rec_schemeT 0 --> recT 0); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 729 | |
| 
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 | (* prepare definitions *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 732 | |
| 4895 | 733 | (*record (scheme) type abbreviation*) | 
| 4890 | 734 | val recordT_specs = | 
| 12247 | 735 | [(suffix schemeN bname, alphas @ [zeta], rec_schemeT 0, Syntax.NoSyn), | 
| 736 | (bname, alphas, recT 0, Syntax.NoSyn)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 737 | |
| 4894 | 738 | (*selectors*) | 
| 4890 | 739 | fun mk_sel_spec (i, c) = | 
| 12247 | 740 | mk_sel (r_scheme 0) c :== mk_fst (funpow i mk_snd (parent_more (r_scheme 0))); | 
| 4894 | 741 | val sel_specs = | 
| 742 | ListPair.map mk_sel_spec (idxs, names) @ | |
| 12247 | 743 | [more_part (r_scheme 0) :== funpow len mk_snd (parent_more (r_scheme 0))]; | 
| 4890 | 744 | |
| 745 | (*updates*) | |
| 12247 | 746 | val all_sels = mk_named_sels all_names (r_scheme 0); | 
| 4890 | 747 | fun mk_upd_spec (i, (c, x)) = | 
| 12247 | 748 | mk_update (r_scheme 0) (c, x) :== | 
| 749 | mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part (r_scheme 0)) | |
| 5197 | 750 | val update_specs = | 
| 751 | ListPair.map mk_upd_spec (idxs, named_vars) @ | |
| 12247 | 752 | [more_part_update (r_scheme 0) more :== mk_record (all_sels, more)]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 753 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 754 | (*derived operations*) | 
| 12265 | 755 | val make_spec = Const (full makeN, all_types ---> recT 0) $$ all_vars :== | 
| 756 | mk_record (all_named_vars, HOLogic.unit); | |
| 757 | val fields_spec = Const (full fieldsN, types ---> fieldsT) $$ vars :== | |
| 758 | mk_record (named_vars, HOLogic.unit); | |
| 12247 | 759 | val extend_spec = Const (full extendN, recT 0 --> moreT --> rec_schemeT 0) $ r 0 $ more :== | 
| 760 | mk_record (mk_named_sels all_names (r 0), more); | |
| 761 | val truncate_spec = Const (full truncateN, rec_schemeT 0 --> recT 0) $ r_scheme 0 :== | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 762 | mk_record (all_sels, HOLogic.unit); | 
| 4894 | 763 | |
| 764 | ||
| 765 | (* prepare propositions *) | |
| 766 | ||
| 767 | (*selectors*) | |
| 768 | val sel_props = | |
| 12247 | 769 | map (fn (c, x) => mk_sel (rec_scheme 0) c === x) named_vars @ | 
| 770 | [more_part (rec_scheme 0) === more]; | |
| 4894 | 771 | |
| 772 | (*updates*) | |
| 773 | fun mk_upd_prop (i, (c, T)) = | |
| 774 | let val x' = Free (variant all_xs (base c ^ "'"), T) in | |
| 12247 | 775 | mk_update (rec_scheme 0) (c, x') === | 
| 4894 | 776 | mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more) | 
| 777 | end; | |
| 5197 | 778 | val update_props = | 
| 779 | ListPair.map mk_upd_prop (idxs, fields) @ | |
| 780 | let val more' = Free (variant all_xs (moreN ^ "'"), moreT) | |
| 12247 | 781 | in [more_part_update (rec_scheme 0) more' === mk_record (all_named_vars, more')] end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 782 | |
| 9626 | 783 | (*equality*) | 
| 784 | fun mk_sel_eq (t, T) = | |
| 12247 | 785 | let val t' = Term.abstract_over (r_scheme 0, t) | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 786 | in Trueprop (HOLogic.eq_const T $ Term.incr_boundvars 1 t' $ t') end; | 
| 12247 | 787 | val sel_eqs = map2 mk_sel_eq | 
| 788 | (map (mk_sel (r_scheme 0)) all_names @ [more_part (r_scheme 0)], all_types @ [moreT]); | |
| 9626 | 789 | val equality_prop = | 
| 12247 | 790 |       Term.all (rec_schemeT 0) $ (Abs ("r", rec_schemeT 0,
 | 
| 791 |         Term.all (rec_schemeT 0) $ (Abs ("r'", rec_schemeT 0,
 | |
| 9626 | 792 | Logic.list_implies (sel_eqs, | 
| 12247 | 793 | Trueprop (HOLogic.eq_const (rec_schemeT 0) $ Bound 1 $ Bound 0)))))); | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 794 | |
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 795 | (*induct*) | 
| 12247 | 796 | fun induct_scheme_prop n = | 
| 797 |       let val P = Free ("P", rec_schemeT n --> HOLogic.boolT) in
 | |
| 798 | (All (prune n all_xs_more ~~ prune n all_types_more) | |
| 799 | (Trueprop (P $ rec_scheme n)), Trueprop (P $ r_scheme n)) | |
| 800 | end; | |
| 801 | fun induct_prop n = | |
| 802 |       let val P = Free ("P", recT n --> HOLogic.boolT) in
 | |
| 803 | (All (prune n all_xs ~~ prune n all_types) (Trueprop (P $ rec_ n)), Trueprop (P $ r n)) | |
| 804 | end; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 805 | |
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 806 | (*cases*) | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 807 | val C = Trueprop (Free (variant all_xs_more "C", HOLogic.boolT)); | 
| 12247 | 808 | fun cases_scheme_prop n = | 
| 809 | All (prune n all_xs_more ~~ prune n all_types_more) | |
| 810 | ((r_scheme n === rec_scheme n) ==> C) ==> C; | |
| 12255 | 811 | fun cases_prop n = All (prune n all_xs ~~ prune n all_types) ((r n === rec_ n) ==> C) ==> C; | 
| 9626 | 812 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 813 | |
| 4890 | 814 | (* 1st stage: fields_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 815 | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 816 | val (fields_thy, field_simps, field_injects, field_splits, field_inducts, field_cases) = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 817 | thy | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 818 | |> Theory.add_path bname | 
| 11832 | 819 | |> field_definitions fields names xs alphas zeta moreT more vars named_vars; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 820 | |
| 12247 | 821 | val all_field_inducts = flat (map #field_inducts parents) @ field_inducts; | 
| 822 | val all_field_cases = flat (map #field_cases parents) @ field_cases; | |
| 823 | ||
| 6092 | 824 | val named_splits = map2 (fn (c, th) => (suffix field_typeN c, th)) (names, field_splits); | 
| 5698 | 825 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 826 | |
| 4890 | 827 | (* 2nd stage: defs_thy *) | 
| 828 | ||
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 829 | val (defs_thy, (((sel_defs, update_defs), derived_defs))) = | 
| 4890 | 830 | fields_thy | 
| 9626 | 831 | |> add_record_splits named_splits | 
| 4890 | 832 | |> Theory.parent_path | 
| 12255 | 833 | |> Theory.add_tyabbrs_i recordT_specs | 
| 4890 | 834 | |> Theory.add_path bname | 
| 5197 | 835 | |> Theory.add_trfuns ([], [], field_tr's, []) | 
| 12506 | 836 | |> Theory.add_consts_i | 
| 837 | (map2 (fn ((x, T), mx) => (x, T, mx)) (sel_decls, field_syntax @ [Syntax.NoSyn])) | |
| 4894 | 838 | |> (Theory.add_consts_i o map Syntax.no_syn) | 
| 12506 | 839 | (update_decls @ [make_decl, fields_decl, extend_decl, truncate_decl]) | 
| 11832 | 840 | |> (PureThy.add_defs_i false o map Thm.no_attributes) sel_specs | 
| 841 | |>>> (PureThy.add_defs_i false o map Thm.no_attributes) update_specs | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 842 | |>>> (PureThy.add_defs_i false o map Thm.no_attributes) | 
| 12590 
3573830e9b91
hide base name of "make", "fields", "extend", "truncate", "more",
 wenzelm parents: 
12506diff
changeset | 843 | [make_spec, fields_spec, extend_spec, truncate_spec] | 
| 
3573830e9b91
hide base name of "make", "fields", "extend", "truncate", "more",
 wenzelm parents: 
12506diff
changeset | 844 | |>> Theory.hide_consts false [full makeN, full fieldsN, full extendN, full truncateN, | 
| 
3573830e9b91
hide base name of "make", "fields", "extend", "truncate", "more",
 wenzelm parents: 
12506diff
changeset | 845 | full moreN, full (suffix updateN moreN)]; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 846 | |
| 4890 | 847 | |
| 848 | (* 3rd stage: thms_thy *) | |
| 849 | ||
| 12247 | 850 | val prove_standard = Tactic.prove_standard (Theory.sign_of defs_thy); | 
| 11967 | 851 | fun prove_simp simps = | 
| 852 | let val tac = simp_all_tac HOL_basic_ss simps | |
| 853 | in fn prop => prove_standard [] [] prop (K tac) end; | |
| 4890 | 854 | |
| 11967 | 855 | val parent_simps = flat (map #simps parents); | 
| 856 | val sel_convs = map (prove_simp (parent_simps @ sel_defs @ field_simps)) sel_props; | |
| 857 | val update_convs = map (prove_simp (parent_simps @ update_defs @ sel_convs)) update_props; | |
| 4894 | 858 | |
| 12247 | 859 | fun induct_scheme n = | 
| 860 | let val (assm, concl) = induct_scheme_prop n in | |
| 861 | prove_standard [] [assm] concl (fn prems => | |
| 862 | EVERY (map (fn rule => try_param_tac "p" rN rule 1) (prune n all_field_inducts)) | |
| 863 | THEN resolve_tac prems 1) | |
| 864 | end; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 865 | |
| 12247 | 866 | fun cases_scheme n = | 
| 867 | prove_standard [] [] (cases_scheme_prop n) (fn _ => | |
| 868 | EVERY (map (fn rule => try_param_tac "p" rN rule 1) (prune n all_field_cases)) | |
| 11967 | 869 | THEN simp_all_tac HOL_basic_ss []); | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 870 | |
| 12247 | 871 | val induct_scheme0 = induct_scheme 0; | 
| 872 | val cases_scheme0 = cases_scheme 0; | |
| 873 | val more_induct_scheme = map induct_scheme ancestry; | |
| 874 | val more_cases_scheme = map cases_scheme ancestry; | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 875 | |
| 12247 | 876 | val (thms_thy, (([sel_convs', update_convs', sel_defs', update_defs', _], | 
| 877 | [induct_scheme', cases_scheme']), [more_induct_scheme', more_cases_scheme'])) = | |
| 11940 | 878 | defs_thy | 
| 879 | |> (PureThy.add_thmss o map Thm.no_attributes) | |
| 880 |        [("select_convs", sel_convs),
 | |
| 881 |         ("update_convs", update_convs),
 | |
| 882 |         ("select_defs", sel_defs),
 | |
| 883 |         ("update_defs", update_defs),
 | |
| 12265 | 884 |         ("defs", derived_defs)]
 | 
| 11940 | 885 | |>>> PureThy.add_thms | 
| 12255 | 886 |        [(("induct_scheme", induct_scheme0), induct_type_global (suffix schemeN name)),
 | 
| 887 |         (("cases_scheme", cases_scheme0), cases_type_global (suffix schemeN name))]
 | |
| 888 | |>>> PureThy.add_thmss | |
| 889 |         [(("more_induct_scheme", more_induct_scheme), induct_type_global ""),
 | |
| 890 |          (("more_cases_scheme", more_cases_scheme), cases_type_global "")];
 | |
| 12247 | 891 | |
| 892 | ||
| 893 | (* 4th stage: more_thms_thy *) | |
| 894 | ||
| 895 | val prove_standard = Tactic.prove_standard (Theory.sign_of thms_thy); | |
| 11940 | 896 | |
| 12247 | 897 | fun induct (n, scheme) = | 
| 898 | let val (assm, concl) = induct_prop n in | |
| 899 | prove_standard [] [assm] concl (fn prems => | |
| 900 | res_inst_tac [(rN, rN)] scheme 1 | |
| 901 | THEN try_param_tac "x" "more" unit_induct 1 | |
| 902 | THEN resolve_tac prems 1) | |
| 903 | end; | |
| 904 | ||
| 905 | fun cases (n, scheme) = | |
| 906 | prove_standard [] [] (cases_prop n) (fn _ => | |
| 907 | res_inst_tac [(rN, rN)] scheme 1 | |
| 908 | THEN simp_all_tac HOL_basic_ss [unit_all_eq1]); | |
| 909 | ||
| 910 | val induct0 = induct (0, induct_scheme'); | |
| 911 | val cases0 = cases (0, cases_scheme'); | |
| 912 | val more_induct = map induct (ancestry ~~ more_induct_scheme'); | |
| 913 | val more_cases = map cases (ancestry ~~ more_cases_scheme'); | |
| 914 | ||
| 915 | val equality = prove_standard [] [] equality_prop (fn _ => | |
| 11967 | 916 | fn st => let val [r, r'] = map #1 (rev (Tactic.innermost_params 1 st)) in | 
| 917 | st |> (res_inst_tac [(rN, r)] cases_scheme' 1 | |
| 918 | THEN res_inst_tac [(rN, r')] cases_scheme' 1 | |
| 919 | THEN simp_all_tac HOL_basic_ss (parent_simps @ sel_convs)) | |
| 920 | end); | |
| 921 | ||
| 12247 | 922 | val (more_thms_thy, [_, _, equality']) = | 
| 923 | thms_thy |> PureThy.add_thms | |
| 12255 | 924 |        [(("induct", induct0), induct_type_global name),
 | 
| 925 |         (("cases", cases0), cases_type_global name),
 | |
| 12374 | 926 |         (("equality", equality), [ContextRules.intro_bang_global None])]
 | 
| 12255 | 927 | |>> (#1 oo PureThy.add_thmss) | 
| 928 |         [(("more_induct", more_induct), induct_type_global ""),
 | |
| 929 |          (("more_cases", more_cases), cases_type_global "")];
 | |
| 11967 | 930 | |
| 11959 | 931 | val simps = sel_convs' @ update_convs' @ [equality']; | 
| 6519 | 932 | val iffs = field_injects; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 933 | |
| 12247 | 934 | val more_thms_thy' = | 
| 935 | more_thms_thy |> (#1 oo PureThy.add_thmss) | |
| 5707 | 936 |         [(("simps", simps), [Simplifier.simp_add_global]),
 | 
| 6519 | 937 |          (("iffs", iffs), [iff_add_global])];
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 938 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 939 | |
| 12247 | 940 | (* 5th stage: final_thy *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 941 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 942 | val final_thy = | 
| 12247 | 943 | more_thms_thy' | 
| 944 | |> put_record name (make_record_info args parent fields field_inducts field_cases | |
| 945 | (field_simps @ simps)) | |
| 11940 | 946 | |> put_sel_upd (names @ [full_moreN]) (field_simps @ sel_defs' @ update_defs') | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 947 | |> Theory.parent_path; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 948 | |
| 6519 | 949 |   in (final_thy, {simps = simps, iffs = iffs}) end;
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 950 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 951 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 952 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 953 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 954 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 955 | (* prepare arguments *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 956 | |
| 4894 | 957 | (*note: read_raw_typ avoids expanding type abbreviations*) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 958 | fun read_raw_parent sign s = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 959 | (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 | 960 | Type (name, Ts) => (Ts, name) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 961 |   | _ => error ("Bad parent record specification: " ^ quote s));
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 962 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 963 | fun read_typ sign (env, s) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 964 | let | 
| 5060 | 965 | fun def_sort (x, ~1) = assoc (env, x) | 
| 966 | | def_sort _ = None; | |
| 967 | 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 | 968 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 969 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 970 | fun cert_typ sign (env, raw_T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 971 | 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 | 972 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 973 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 974 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 975 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 976 | |
| 4895 | 977 | (*we do all preparations and error checks here, deferring the real | 
| 978 | work to record_definition*) | |
| 4890 | 979 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 980 | 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 | 981 | let | 
| 4970 | 982 | val _ = Theory.requires thy "Record" "record definitions"; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 983 | val sign = Theory.sign_of thy; | 
| 5698 | 984 |     val _ = message ("Defining record " ^ quote bname ^ " ...");
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 985 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 986 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 987 | (* parents *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 988 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 989 | fun prep_inst T = snd (cert_typ sign ([], T)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 990 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 991 | 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 | 992 |       handle ERROR => error ("The error(s) above in parent record specification");
 | 
| 12247 | 993 | val parents = add_parents thy parent []; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 994 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 995 | val init_env = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 996 | (case parent of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 997 | None => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 998 | | Some (types, _) => foldr Term.add_typ_tfrees (types, [])); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 999 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1000 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1001 | (* fields *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1002 | |
| 12506 | 1003 | fun prep_field (env, (c, raw_T, mx)) = | 
| 4967 | 1004 | let val (env', T) = prep_typ sign (env, raw_T) handle ERROR => | 
| 1005 |         error ("The error(s) above occured in field " ^ quote c)
 | |
| 12506 | 1006 | in (env', (c, T, mx)) end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1007 | |
| 4967 | 1008 | 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 | 1009 | val envir_names = map fst envir; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1010 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1011 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1012 | (* args *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1013 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1014 | val defaultS = Sign.defaultS sign; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1015 | 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 | 1016 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1017 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1018 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1019 | |
| 4890 | 1020 | val name = Sign.full_name sign bname; | 
| 1021 | val err_dup_record = | |
| 1022 | if is_none (get_record thy name) then [] | |
| 1023 | else ["Duplicate definition of record " ^ quote name]; | |
| 1024 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1025 | val err_dup_parms = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1026 | (case duplicates params of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1027 | [] => [] | 
| 4890 | 1028 | | dups => ["Duplicate parameter(s) " ^ commas dups]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1029 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1030 | val err_extra_frees = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1031 | (case gen_rems (op =) (envir_names, params) of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1032 | [] => [] | 
| 4890 | 1033 | | extras => ["Extra free type variable(s) " ^ commas extras]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1034 | |
| 4890 | 1035 | 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 | 1036 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1037 | val err_dup_fields = | 
| 12506 | 1038 | (case duplicates (map #1 bfields) of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1039 | [] => [] | 
| 4890 | 1040 | | dups => ["Duplicate field(s) " ^ commas_quote dups]); | 
| 1041 | ||
| 1042 | val err_bad_fields = | |
| 12506 | 1043 | if forall (not_equal moreN o #1) bfields then [] | 
| 4890 | 1044 | else ["Illegal field name " ^ quote moreN]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1045 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1046 | val err_dup_sorts = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1047 | (case duplicates envir_names of | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1048 | [] => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1049 | | dups => ["Inconsistent sort constraints for " ^ commas dups]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1050 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1051 | val errs = | 
| 4890 | 1052 | err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @ | 
| 1053 | err_dup_fields @ err_bad_fields @ err_dup_sorts; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1054 | in | 
| 4890 | 1055 | if null errs then () else error (cat_lines errs); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1056 | thy |> record_definition (args, bname) parent parents bfields | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1057 | end | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1058 |   handle ERROR => error ("Failed to define record " ^ quote bname);
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1059 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1060 | 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 | 1061 | 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 | 1062 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1063 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1064 | |
| 6358 | 1065 | (** package setup **) | 
| 1066 | ||
| 1067 | (* setup theory *) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1068 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1069 | val setup = | 
| 5006 | 1070 | [RecordsData.init, | 
| 11490 
f9ae28f55178
field_name_ast_tr superceded by constify_ast_tr in Pure;
 wenzelm parents: 
11473diff
changeset | 1071 | Theory.add_trfuns ([], parse_translation, [], []), | 
| 6358 | 1072 | Method.add_methods [record_split_method], | 
| 7178 | 1073 | Simplifier.change_simpset_of Simplifier.addsimprocs [record_simproc]]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1074 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1075 | |
| 6358 | 1076 | (* outer syntax *) | 
| 1077 | ||
| 6723 | 1078 | local structure P = OuterParse and K = OuterSyntax.Keyword in | 
| 6358 | 1079 | |
| 1080 | val record_decl = | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12590diff
changeset | 1081 | P.type_args -- P.name -- | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12590diff
changeset | 1082 | (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const); | 
| 6358 | 1083 | |
| 1084 | val recordP = | |
| 6723 | 1085 | OuterSyntax.command "record" "define extensible record" K.thy_decl | 
| 6519 | 1086 | (record_decl >> (fn (x, (y, z)) => Toplevel.theory (#1 o add_record x y z))); | 
| 6358 | 1087 | |
| 1088 | val _ = OuterSyntax.add_parsers [recordP]; | |
| 1089 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1090 | end; | 
| 5698 | 1091 | |
| 6384 | 1092 | end; | 
| 1093 | ||
| 5698 | 1094 | structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage; | 
| 1095 | open BasicRecordPackage; |