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