| author | krauss | 
| Thu, 22 Mar 2007 13:36:57 +0100 | |
| changeset 22498 | 62cdd4b3e96b | 
| parent 22219 | 61b5bab471ce | 
| child 22578 | b0eb5652f210 | 
| 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$ | 
| 14579 | 3 | Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 4 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 5 | Extensible records with structural subtyping in HOL. | 
| 5698 | 6 | *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 7 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 8 | |
| 5698 | 9 | signature BASIC_RECORD_PACKAGE = | 
| 10 | sig | |
| 7178 | 11 | val record_simproc: simproc | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 12 | val record_eq_simproc: simproc | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 13 | val record_upd_simproc: simproc | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 14 | val record_split_simproc: (term -> int) -> simproc | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 15 | val record_ex_sel_eq_simproc: simproc | 
| 5698 | 16 | val record_split_tac: int -> tactic | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 17 | val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic | 
| 5713 | 18 | val record_split_name: string | 
| 5698 | 19 | val record_split_wrapper: string * wrapper | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 20 | val print_record_type_abbr: bool ref | 
| 17261 | 21 | val print_record_type_as_fields: bool ref | 
| 5698 | 22 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 23 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 24 | signature RECORD_PACKAGE = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 25 | sig | 
| 5698 | 26 | include BASIC_RECORD_PACKAGE | 
| 27 | val quiet_mode: bool ref | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 28 | val record_quick_and_dirty_sensitive: bool ref | 
| 8574 | 29 | val updateN: string | 
| 21363 | 30 | val updN: string | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 31 | val ext_typeN: string | 
| 21363 | 32 | val extN: string | 
| 33 | val makeN: string | |
| 34 | val moreN: string | |
| 35 | val ext_dest: string | |
| 36 | val KN:string | |
| 37 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 38 | val last_extT: typ -> (string * typ list) option | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 39 | val dest_recTs : typ -> (string * typ list) list | 
| 16458 | 40 | val get_extT_fields: theory -> typ -> ((string * typ) list * (string * typ)) | 
| 41 | val get_recT_fields: theory -> typ -> ((string * typ) list * (string * typ)) | |
| 42 | val get_extension: theory -> Symtab.key -> (string * typ list) option | |
| 43 | val get_extinjects: theory -> thm list | |
| 44 | val get_simpset: theory -> simpset | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 45 | val print_records: theory -> unit | 
| 16458 | 46 | val add_record: string list * string -> string option -> (string * string * mixfix) list | 
| 47 | -> theory -> theory | |
| 48 | val add_record_i: string list * string -> (typ list * string) option | |
| 49 | -> (string * typ * mixfix) list -> theory -> theory | |
| 18708 | 50 | val setup: theory -> theory | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 51 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 52 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 53 | |
| 17960 | 54 | structure RecordPackage: RECORD_PACKAGE = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 55 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 56 | |
| 21546 | 57 | val eq_reflection = thm "eq_reflection"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 58 | val rec_UNIV_I = thm "rec_UNIV_I"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 59 | val rec_True_simp = thm "rec_True_simp"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 60 | val Pair_eq = thm "Product_Type.Pair_eq"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 61 | val atomize_all = thm "HOL.atomize_all"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 62 | val atomize_imp = thm "HOL.atomize_imp"; | 
| 17960 | 63 | val meta_allE = thm "Pure.meta_allE"; | 
| 64 | val prop_subst = thm "prop_subst"; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 65 | val Pair_sel_convs = [fst_conv,snd_conv]; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 66 | val K_record_apply = thm "Record.K_record_apply"; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 67 | val K_comp_convs = [o_apply,K_record_apply] | 
| 11832 | 68 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 69 | (** name components **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 70 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 71 | val rN = "r"; | 
| 15215 | 72 | val wN = "w"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 73 | val moreN = "more"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 74 | val schemeN = "_scheme"; | 
| 17261 | 75 | val ext_typeN = "_ext_type"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 76 | val extN ="_ext"; | 
| 15215 | 77 | val casesN = "_cases"; | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 78 | val ext_dest = "_sel"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 79 | val updateN = "_update"; | 
| 15215 | 80 | val updN = "_upd"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 81 | val makeN = "make"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 82 | val fields_selN = "fields"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 83 | val extendN = "extend"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 84 | val truncateN = "truncate"; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 85 | val KN = "Record.K_record"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 86 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 87 | (*see typedef_package.ML*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 88 | val RepN = "Rep_"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 89 | val AbsN = "Abs_"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 90 | |
| 4894 | 91 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 92 | |
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 93 | fun but_last xs = fst (split_last xs); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 94 | |
| 19748 | 95 | fun varifyT midx = | 
| 96 | let fun varify (a, S) = TVar ((a, midx + 1), S); | |
| 97 | in map_type_tfree varify end; | |
| 98 | ||
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 99 | fun the_plist (SOME (a,b)) = [a,b] | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 100 | | the_plist NONE = raise Option; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 101 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 102 | fun domain_type' T = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 103 | domain_type T handle Match => T; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 104 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 105 | fun range_type' T = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 106 | range_type T handle Match => T; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 107 | |
| 5698 | 108 | (* messages *) | 
| 109 | ||
| 110 | val quiet_mode = ref false; | |
| 111 | fun message s = if ! quiet_mode then () else writeln s; | |
| 112 | ||
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 113 | fun trace_thm str thm = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 114 | tracing (str ^ (Pretty.string_of (Display.pretty_thm thm))); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 115 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 116 | fun trace_thms str thms = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 117 | (tracing str; map (trace_thm "") thms); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 118 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 119 | fun trace_term str t = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 120 | tracing (str ^ (Display.raw_string_of_term t)); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 121 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 122 | (* timing *) | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 123 | |
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 124 | fun timeit_msg s x = if !timing then (warning s; timeit x) else x (); | 
| 16379 | 125 | fun timing_msg s = if !timing then warning s else (); | 
| 17261 | 126 | |
| 12255 | 127 | (* syntax *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 128 | |
| 12247 | 129 | fun prune n xs = Library.drop (n, xs); | 
| 11832 | 130 | fun prefix_base s = NameSpace.map_base (fn bname => s ^ bname); | 
| 131 | ||
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 132 | val Trueprop = HOLogic.mk_Trueprop; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 133 | fun All xs t = Term.list_all_free (xs, t); | 
| 4894 | 134 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 135 | infix 9 $$; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 136 | infix 0 :== ===; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 137 | infixr 0 ==>; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 138 | |
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 139 | val (op $$) = Term.list_comb; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 140 | val (op :==) = Logic.mk_defpair; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 141 | val (op ===) = Trueprop o HOLogic.mk_eq; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 142 | val (op ==>) = Logic.mk_implies; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 143 | |
| 11832 | 144 | (* morphisms *) | 
| 145 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 146 | fun mk_RepN name = suffix ext_typeN (prefix_base RepN name); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 147 | fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name); | 
| 11832 | 148 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 149 | fun mk_Rep name repT absT = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 150 | Const (suffix ext_typeN (prefix_base RepN name),absT --> repT); | 
| 11832 | 151 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 152 | fun mk_Abs name repT absT = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 153 | Const (mk_AbsN name,repT --> absT); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 154 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 155 | (* constructor *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 156 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 157 | fun mk_extC (name,T) Ts = (suffix extN name, Ts ---> T); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 158 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 159 | fun mk_ext (name,T) ts = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 160 | let val Ts = map fastype_of ts | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 161 | in list_comb (Const (mk_extC (name,T) Ts),ts) end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 162 | |
| 15215 | 163 | (* cases *) | 
| 164 | ||
| 165 | fun mk_casesC (name,T,vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT) | |
| 166 | ||
| 167 | fun mk_cases (name,T,vT) f = | |
| 17261 | 168 | let val Ts = binder_types (fastype_of f) | 
| 15215 | 169 | in Const (mk_casesC (name,T,vT) Ts) $ f end; | 
| 17261 | 170 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 171 | (* selector *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 172 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 173 | fun mk_selC sT (c,T) = (c,sT --> T); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 174 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 175 | fun mk_sel s (c,T) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 176 | let val sT = fastype_of s | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 177 | in Const (mk_selC sT (c,T)) $ s end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 178 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 179 | (* updates *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 180 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 181 | fun mk_updC sfx sT (c,T) = (suffix sfx c, (T --> T) --> sT --> sT); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 182 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 183 | fun mk_upd' sfx c v sT = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 184 | let val vT = domain_type (fastype_of v); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 185 | in Const (mk_updC sfx sT (c, vT)) $ v end; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 186 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 187 | fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 188 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 189 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 190 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 191 | fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 192 | (case try (unsuffix ext_typeN) c_ext_type of | 
| 15531 | 193 |         NONE => raise TYPE ("RecordPackage.dest_recT", [typ], [])
 | 
| 15570 | 194 | | SOME c => ((c, Ts), List.last Ts)) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 195 |   | dest_recT typ = raise TYPE ("RecordPackage.dest_recT", [typ], []);
 | 
| 5197 | 196 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 197 | fun is_recT T = | 
| 17261 | 198 | (case try dest_recT T of NONE => false | SOME _ => true); | 
| 11833 | 199 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 200 | fun dest_recTs T = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 201 | let val ((c, Ts), U) = dest_recT T | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 202 | in (c, Ts) :: dest_recTs U | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 203 | end handle TYPE _ => []; | 
| 14255 | 204 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 205 | fun last_extT T = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 206 | let val ((c, Ts), U) = dest_recT T | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 207 | in (case last_extT U of | 
| 15531 | 208 | NONE => SOME (c,Ts) | 
| 209 | | SOME l => SOME l) | |
| 210 | end handle TYPE _ => NONE | |
| 14255 | 211 | |
| 17261 | 212 | fun rec_id i T = | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 213 | let val rTs = dest_recTs T | 
| 15570 | 214 | val rTs' = if i < 0 then rTs else Library.take (i,rTs) | 
| 215 |   in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end;
 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 216 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 217 | (*** extend theory by record definition ***) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 218 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 219 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 220 | |
| 14255 | 221 | (* type record_info and parent_info *) | 
| 4867 
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 | type record_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 224 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 225 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 226 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 227 | extension: (string * typ list), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 228 | induct: thm | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 229 | }; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 230 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 231 | fun make_record_info args parent fields extension induct = | 
| 17261 | 232 |  {args = args, parent = parent, fields = fields, extension = extension,
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 233 | induct = induct}: record_info; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 234 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 235 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 236 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 237 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 238 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 239 | extension: (string * typ list), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 240 | induct: thm | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 241 | }; | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 242 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 243 | fun make_parent_info name fields extension induct = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 244 |  {name = name, fields = fields, extension = extension, induct = induct}: parent_info;
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 245 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 246 | (* data kind 'HOL/record' *) | 
| 5001 | 247 | |
| 7178 | 248 | type record_data = | 
| 249 |  {records: record_info Symtab.table,
 | |
| 250 | sel_upd: | |
| 251 |    {selectors: unit Symtab.table,
 | |
| 252 | updates: string Symtab.table, | |
| 253 | simpset: Simplifier.simpset}, | |
| 14255 | 254 | equalities: thm Symtab.table, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 255 | extinjects: thm list, | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 256 | extsplit: thm Symtab.table, (* maps extension name to split rule *) | 
| 17261 | 257 | splits: (thm*thm*thm*thm) Symtab.table, (* !!,!,EX - split-equalities,induct rule *) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 258 | extfields: (string*typ) list Symtab.table, (* maps extension to its fields *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 259 | fieldext: (string*typ list) Symtab.table (* maps field to its extension *) | 
| 14255 | 260 | }; | 
| 7178 | 261 | |
| 17261 | 262 | fun make_record_data | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 263 | records sel_upd equalities extinjects extsplit splits extfields fieldext = | 
| 17261 | 264 |  {records = records, sel_upd = sel_upd,
 | 
| 265 | equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits, | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 266 | extfields = extfields, fieldext = fieldext }: record_data; | 
| 7178 | 267 | |
| 16458 | 268 | structure RecordsData = TheoryDataFun | 
| 269 | (struct | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 270 | val name = "HOL/record"; | 
| 7178 | 271 | type T = record_data; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 272 | |
| 7178 | 273 | val empty = | 
| 274 | make_record_data Symtab.empty | |
| 275 |       {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss}
 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 276 | Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; | 
| 7178 | 277 | |
| 6556 | 278 | val copy = I; | 
| 16458 | 279 | val extend = I; | 
| 280 | fun merge _ | |
| 7178 | 281 |    ({records = recs1,
 | 
| 282 |      sel_upd = {selectors = sels1, updates = upds1, simpset = ss1},
 | |
| 14255 | 283 | equalities = equalities1, | 
| 17261 | 284 | extinjects = extinjects1, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 285 | extsplit = extsplit1, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 286 | splits = splits1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 287 | extfields = extfields1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 288 | fieldext = fieldext1}, | 
| 7178 | 289 |     {records = recs2,
 | 
| 290 |      sel_upd = {selectors = sels2, updates = upds2, simpset = ss2},
 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 291 | equalities = equalities2, | 
| 17261 | 292 | extinjects = extinjects2, | 
| 293 | extsplit = extsplit2, | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 294 | splits = splits2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 295 | extfields = extfields2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 296 | fieldext = fieldext2}) = | 
| 17261 | 297 | make_record_data | 
| 7178 | 298 | (Symtab.merge (K true) (recs1, recs2)) | 
| 299 |       {selectors = Symtab.merge (K true) (sels1, sels2),
 | |
| 300 | updates = Symtab.merge (K true) (upds1, upds2), | |
| 301 | simpset = Simplifier.merge_ss (ss1, ss2)} | |
| 14255 | 302 | (Symtab.merge Thm.eq_thm (equalities1, equalities2)) | 
| 15248 
b436486091a6
record_split_simp_tac now can get simp rules as parameter
 schirmer parents: 
15215diff
changeset | 303 | (gen_merge_lists Thm.eq_thm extinjects1 extinjects2) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 304 | (Symtab.merge Thm.eq_thm (extsplit1,extsplit2)) | 
| 17261 | 305 | (Symtab.merge (fn ((a,b,c,d),(w,x,y,z)) | 
| 306 | => Thm.eq_thm (a,w) andalso Thm.eq_thm (b,x) andalso | |
| 307 | Thm.eq_thm (c,y) andalso Thm.eq_thm (d,z)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 308 | (splits1, splits2)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 309 | (Symtab.merge (K true) (extfields1,extfields2)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 310 | (Symtab.merge (K true) (fieldext1,fieldext2)); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 311 | |
| 18858 | 312 |   fun print thy ({records = recs, ...}: record_data) =
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 313 | let | 
| 18858 | 314 | val prt_typ = Sign.pretty_typ thy; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 315 | |
| 15531 | 316 | fun pretty_parent NONE = [] | 
| 317 | | pretty_parent (SOME (Ts, name)) = | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 318 | [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; | 
| 
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 | fun pretty_field (c, T) = Pretty.block | 
| 18858 | 321 | [Pretty.str (Sign.extern_const thy c), Pretty.str " ::", | 
| 12129 | 322 | Pretty.brk 1, Pretty.quote (prt_typ T)]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 323 | |
| 12247 | 324 |       fun pretty_record (name, {args, parent, fields, ...}: record_info) =
 | 
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 325 | Pretty.block (Pretty.fbreaks (Pretty.block | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 326 | [prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 327 | pretty_parent parent @ map pretty_field fields)); | 
| 12129 | 328 | in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end; | 
| 16458 | 329 | end); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 330 | |
| 5006 | 331 | val print_records = RecordsData.print; | 
| 332 | ||
| 16458 | 333 | |
| 7178 | 334 | (* access 'records' *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 335 | |
| 17412 | 336 | val get_record = Symtab.lookup o #records o RecordsData.get; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 337 | |
| 4890 | 338 | fun put_record name info thy = | 
| 7178 | 339 | let | 
| 17261 | 340 |     val {records, sel_upd, equalities, extinjects,extsplit,splits,extfields,fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 341 | RecordsData.get thy; | 
| 17412 | 342 | val data = make_record_data (Symtab.update (name, info) records) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 343 | sel_upd equalities extinjects extsplit splits extfields fieldext; | 
| 7178 | 344 | in RecordsData.put data thy end; | 
| 345 | ||
| 346 | (* access 'sel_upd' *) | |
| 347 | ||
| 16458 | 348 | val get_sel_upd = #sel_upd o RecordsData.get; | 
| 7178 | 349 | |
| 17510 | 350 | val is_selector = Symtab.defined o #selectors o get_sel_upd; | 
| 17412 | 351 | val get_updates = Symtab.lookup o #updates o get_sel_upd; | 
| 17892 | 352 | fun get_simpset thy = Simplifier.theory_context thy (#simpset (get_sel_upd thy)); | 
| 7178 | 353 | |
| 354 | fun put_sel_upd names simps thy = | |
| 355 | let | |
| 356 | val sels = map (rpair ()) names; | |
| 357 | val upds = map (suffix updateN) names ~~ names; | |
| 358 | ||
| 17261 | 359 |     val {records, sel_upd = {selectors, updates, simpset},
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 360 | equalities, extinjects, extsplit, splits, extfields,fieldext} = RecordsData.get thy; | 
| 7178 | 361 | val data = make_record_data records | 
| 362 |       {selectors = Symtab.extend (selectors, sels),
 | |
| 363 | updates = Symtab.extend (updates, upds), | |
| 364 | simpset = Simplifier.addsimps (simpset, simps)} | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 365 | equalities extinjects extsplit splits extfields fieldext; | 
| 7178 | 366 | in RecordsData.put data thy end; | 
| 367 | ||
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 368 | (* access 'equalities' *) | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 369 | |
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 370 | fun add_record_equalities name thm thy = | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 371 | let | 
| 17261 | 372 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 373 | RecordsData.get thy; | 
| 17261 | 374 | val data = make_record_data records sel_upd | 
| 17412 | 375 | (Symtab.update_new (name, thm) equalities) extinjects extsplit | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 376 | splits extfields fieldext; | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 377 | in RecordsData.put data thy end; | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 378 | |
| 17412 | 379 | val get_equalities =Symtab.lookup o #equalities o RecordsData.get; | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 380 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 381 | (* access 'extinjects' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 382 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 383 | fun add_extinjects thm thy = | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 384 | let | 
| 17261 | 385 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 386 | RecordsData.get thy; | 
| 17261 | 387 | val data = make_record_data records sel_upd equalities (extinjects@[thm]) extsplit | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 388 | splits extfields fieldext; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 389 | in RecordsData.put data thy end; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 390 | |
| 18858 | 391 | fun get_extinjects thy = #extinjects (RecordsData.get thy); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 392 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 393 | (* access 'extsplit' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 394 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 395 | fun add_extsplit name thm thy = | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 396 | let | 
| 17261 | 397 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 398 | RecordsData.get thy; | 
| 17261 | 399 | val data = make_record_data records sel_upd | 
| 17412 | 400 | equalities extinjects (Symtab.update_new (name, thm) extsplit) splits | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 401 | extfields fieldext; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 402 | in RecordsData.put data thy end; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 403 | |
| 17412 | 404 | val get_extsplit = Symtab.lookup o #extsplit o RecordsData.get; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 405 | |
| 14255 | 406 | (* access 'splits' *) | 
| 407 | ||
| 408 | fun add_record_splits name thmP thy = | |
| 409 | let | |
| 17261 | 410 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 411 | RecordsData.get thy; | 
| 17261 | 412 | val data = make_record_data records sel_upd | 
| 17412 | 413 | equalities extinjects extsplit (Symtab.update_new (name, thmP) splits) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 414 | extfields fieldext; | 
| 14255 | 415 | in RecordsData.put data thy end; | 
| 416 | ||
| 17412 | 417 | val get_splits = Symtab.lookup o #splits o RecordsData.get; | 
| 14255 | 418 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 419 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 420 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 421 | (* extension of a record name *) | 
| 17261 | 422 | val get_extension = | 
| 17412 | 423 | Option.map #extension oo (Symtab.lookup o #records o RecordsData.get); | 
| 17261 | 424 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 425 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 426 | (* access 'extfields' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 427 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 428 | fun add_extfields name fields thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 429 | let | 
| 17261 | 430 |     val {records, sel_upd, equalities, extinjects, extsplit,splits, extfields, fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 431 | RecordsData.get thy; | 
| 17261 | 432 | val data = make_record_data records sel_upd | 
| 433 | equalities extinjects extsplit splits | |
| 17412 | 434 | (Symtab.update_new (name, fields) extfields) fieldext; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 435 | in RecordsData.put data thy end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 436 | |
| 17412 | 437 | val get_extfields = Symtab.lookup o #extfields o RecordsData.get; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 438 | |
| 18858 | 439 | fun get_extT_fields thy T = | 
| 15059 | 440 | let | 
| 441 | val ((name,Ts),moreT) = dest_recT T; | |
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 442 | val recname = let val (nm::recn::rst) = rev (NameSpace.explode name) | 
| 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 443 | in NameSpace.implode (rev (nm::rst)) end; | 
| 15059 | 444 | val midx = maxidx_of_typs (moreT::Ts); | 
| 19748 | 445 | val varifyT = varifyT midx; | 
| 18858 | 446 |     val {records,extfields,...} = RecordsData.get thy;
 | 
| 18931 | 447 | val (flds,(more,_)) = split_last (Symtab.lookup_list extfields name); | 
| 17412 | 448 | val args = map varifyT (snd (#extension (the (Symtab.lookup records recname)))); | 
| 15058 | 449 | |
| 19748 | 450 | val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) (Vartab.empty); | 
| 15059 | 451 | val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds; | 
| 452 | in (flds',(more,moreT)) end; | |
| 15058 | 453 | |
| 18858 | 454 | fun get_recT_fields thy T = | 
| 17261 | 455 | let | 
| 18858 | 456 | val (root_flds,(root_more,root_moreT)) = get_extT_fields thy T; | 
| 17261 | 457 | val (rest_flds,rest_more) = | 
| 18858 | 458 | if is_recT root_moreT then get_recT_fields thy root_moreT | 
| 15059 | 459 | else ([],(root_more,root_moreT)); | 
| 460 | in (root_flds@rest_flds,rest_more) end; | |
| 461 | ||
| 15058 | 462 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 463 | (* access 'fieldext' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 464 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 465 | fun add_fieldext extname_types fields thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 466 | let | 
| 17261 | 467 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 468 | RecordsData.get thy; | 
| 17261 | 469 | val fieldext' = | 
| 17412 | 470 | fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext; | 
| 17261 | 471 | val data=make_record_data records sel_upd equalities extinjects extsplit | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 472 | splits extfields fieldext'; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 473 | in RecordsData.put data thy end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 474 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 475 | |
| 17412 | 476 | val get_fieldext = Symtab.lookup o #fieldext o RecordsData.get; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 477 | |
| 21962 | 478 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 479 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 480 | |
| 15531 | 481 | fun add_parents thy NONE parents = parents | 
| 482 | | add_parents thy (SOME (types, name)) parents = | |
| 12247 | 483 | let | 
| 484 | val sign = Theory.sign_of thy; | |
| 485 | fun err msg = error (msg ^ " parent record " ^ quote name); | |
| 12255 | 486 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 487 |         val {args, parent, fields, extension, induct} =
 | 
| 15531 | 488 | (case get_record thy name of SOME info => info | NONE => err "Unknown"); | 
| 12247 | 489 | val _ = if length types <> length args then err "Bad number of arguments for" else (); | 
| 12255 | 490 | |
| 12247 | 491 | fun bad_inst ((x, S), T) = | 
| 15531 | 492 | if Sign.of_sort sign (T, S) then NONE else SOME x | 
| 15570 | 493 | val bads = List.mapPartial bad_inst (args ~~ types); | 
| 21962 | 494 |         val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
 | 
| 12255 | 495 | |
| 12247 | 496 | val inst = map fst args ~~ types; | 
| 17377 | 497 | val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); | 
| 15570 | 498 | val parent' = Option.map (apfst (map subst)) parent; | 
| 12247 | 499 | val fields' = map (apsnd subst) fields; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 500 | val extension' = apsnd (map subst) extension; | 
| 12247 | 501 | in | 
| 12255 | 502 | add_parents thy parent' | 
| 21962 | 503 | (make_parent_info name fields' extension' induct :: parents) | 
| 12247 | 504 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 505 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 506 | |
| 21962 | 507 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 508 | (** concrete syntax for records **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 509 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 510 | (* parse translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 511 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 512 | fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 513 | if c = mark then Syntax.const (suffix sfx name) $ (Syntax.const KN $ arg) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 514 |       else raise TERM ("gen_field_tr: " ^ mark, [t])
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 515 |   | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 516 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 517 | fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 518 | if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 519 | else [gen_field_tr mark sfx tm] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 520 | | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 521 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 522 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 523 | fun record_update_tr [t, u] = | 
| 21078 | 524 | Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 525 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 526 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 527 | fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 528 | | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 529 |   | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 530 | (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 531 |   | update_name_tr ts = raise TERM ("update_name_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 532 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 533 | fun dest_ext_field mark (t as (Const (c,_) $ Const (name,_) $ arg)) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 534 |      if c = mark then (name,arg) else raise TERM ("dest_ext_field: " ^ mark, [t])
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 535 |   | dest_ext_field _ t = raise TERM ("dest_ext_field", [t])
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 536 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 537 | fun dest_ext_fields sep mark (trm as (Const (c,_) $ t $ u)) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 538 | if c = sep then dest_ext_field mark t::dest_ext_fields sep mark u | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 539 | else [dest_ext_field mark trm] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 540 | | dest_ext_fields _ mark t = [dest_ext_field mark t] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 541 | |
| 21772 | 542 | fun gen_ext_fields_tr sep mark sfx more ctxt t = | 
| 17261 | 543 | let | 
| 21772 | 544 | val thy = ProofContext.theory_of ctxt; | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 545 | val msg = "error in record input: "; | 
| 17261 | 546 | val fieldargs = dest_ext_fields sep mark t; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 547 | fun splitargs (field::fields) ((name,arg)::fargs) = | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 548 | if can (unsuffix name) field | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 549 | then let val (args,rest) = splitargs fields fargs | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 550 | in (arg::args,rest) end | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 551 | else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 552 | | splitargs [] (fargs as (_::_)) = ([],fargs) | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 553 | | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 554 | | splitargs _ _ = ([],[]); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 555 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 556 | fun mk_ext (fargs as (name,arg)::_) = | 
| 18858 | 557 | (case get_fieldext thy (Sign.intern_const thy name) of | 
| 558 | SOME (ext,_) => (case get_extfields thy ext of | |
| 17261 | 559 | SOME flds | 
| 560 | => let val (args,rest) = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 561 | splitargs (map fst (but_last flds)) fargs; | 
| 17261 | 562 | val more' = mk_ext rest; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 563 | in list_comb (Syntax.const (suffix sfx ext),args@[more']) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 564 | end | 
| 15531 | 565 | | NONE => raise TERM(msg ^ "no fields defined for " | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 566 | ^ ext,[t])) | 
| 15531 | 567 | | NONE => raise TERM (msg ^ name ^" is no proper field",[t])) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 568 | | mk_ext [] = more | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 569 | |
| 17261 | 570 | in mk_ext fieldargs end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 571 | |
| 21772 | 572 | fun gen_ext_type_tr sep mark sfx more ctxt t = | 
| 17261 | 573 | let | 
| 21772 | 574 | val thy = ProofContext.theory_of ctxt; | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 575 | val msg = "error in record-type input: "; | 
| 17261 | 576 | val fieldargs = dest_ext_fields sep mark t; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 577 | fun splitargs (field::fields) ((name,arg)::fargs) = | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 578 | if can (unsuffix name) field | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 579 | then let val (args,rest) = splitargs fields fargs | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 580 | in (arg::args,rest) end | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 581 | else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 582 | | splitargs [] (fargs as (_::_)) = ([],fargs) | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 583 | | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 584 | | splitargs _ _ = ([],[]); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 585 | |
| 18858 | 586 | fun get_sort xs n = AList.lookup (op =) xs n |> the_default (Sign.defaultS thy); | 
| 17261 | 587 | |
| 18858 | 588 | fun to_type t = Sign.certify_typ thy | 
| 589 | (Sign.intern_typ thy | |
| 15957 | 590 | (Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t)); | 
| 16934 | 591 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 592 | fun mk_ext (fargs as (name,arg)::_) = | 
| 18858 | 593 | (case get_fieldext thy (Sign.intern_const thy name) of | 
| 17261 | 594 | SOME (ext,alphas) => | 
| 18858 | 595 | (case get_extfields thy ext of | 
| 17261 | 596 | SOME flds | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 597 | => (let | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 598 | val flds' = but_last flds; | 
| 17261 | 599 | val types = map snd flds'; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 600 | val (args,rest) = splitargs (map fst flds') fargs; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 601 | val argtypes = map to_type args; | 
| 19748 | 602 | val midx = fold (fn T => fn i => Int.max (maxidx_of_typ T, i)) | 
| 603 | argtypes 0; | |
| 604 | val varifyT = varifyT midx; | |
| 605 | val vartypes = map varifyT types; | |
| 606 | ||
| 607 | val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes) | |
| 608 | Vartab.empty; | |
| 17261 | 609 | val alphas' = map ((Syntax.term_of_typ (! Syntax.show_sorts)) o | 
| 19748 | 610 | Envir.norm_type subst o varifyT) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 611 | (but_last alphas); | 
| 17261 | 612 | |
| 613 | val more' = mk_ext rest; | |
| 614 | in list_comb (Syntax.const (suffix sfx ext),alphas'@[more']) | |
| 19750 | 615 | end handle TYPE_MATCH => raise | 
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 616 | TERM (msg ^ "type is no proper record (extension)", [t])) | 
| 15531 | 617 | | NONE => raise TERM (msg ^ "no fields defined for " ^ ext,[t])) | 
| 618 | | NONE => raise TERM (msg ^ name ^" is no proper field",[t])) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 619 | | mk_ext [] = more | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 620 | |
| 17261 | 621 | in mk_ext fieldargs end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 622 | |
| 21772 | 623 | fun gen_adv_record_tr sep mark sfx unit ctxt [t] = | 
| 624 | gen_ext_fields_tr sep mark sfx unit ctxt t | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 625 |   | gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 626 | |
| 21772 | 627 | fun gen_adv_record_scheme_tr sep mark sfx ctxt [t, more] = | 
| 628 | gen_ext_fields_tr sep mark sfx more ctxt t | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 629 |   | gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 630 | |
| 21772 | 631 | fun gen_adv_record_type_tr sep mark sfx unit ctxt [t] = | 
| 632 | gen_ext_type_tr sep mark sfx unit ctxt t | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 633 |   | gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 634 | |
| 21772 | 635 | fun gen_adv_record_type_scheme_tr sep mark sfx ctxt [t, more] = | 
| 636 | gen_ext_type_tr sep mark sfx more ctxt t | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 637 |   | gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 638 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 639 | val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 640 | val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 641 | |
| 17261 | 642 | val adv_record_type_tr = | 
| 643 | gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 644 | (Syntax.term_of_typ false (HOLogic.unitT)); | 
| 17261 | 645 | val adv_record_type_scheme_tr = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 646 | gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 647 | |
| 15215 | 648 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 649 | val parse_translation = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 650 |  [("_record_update", record_update_tr),
 | 
| 17261 | 651 |   ("_update_name", update_name_tr)];
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 652 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 653 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 654 | val adv_parse_translation = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 655 |  [("_record",adv_record_tr),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 656 |   ("_record_scheme",adv_record_scheme_tr),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 657 |   ("_record_type",adv_record_type_tr),
 | 
| 17261 | 658 |   ("_record_type_scheme",adv_record_type_scheme_tr)];
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 659 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 660 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 661 | (* print translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 662 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 663 | val print_record_type_abbr = ref true; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 664 | val print_record_type_as_fields = ref true; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 665 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 666 | fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ (Const ("K_record",_)$t) $ u) =
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 667 | (case try (unsuffix sfx) name_field of | 
| 15531 | 668 | SOME name => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 669 | apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u) | 
| 15531 | 670 | | NONE => ([], tm)) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 671 | | gen_field_upds_tr' _ _ tm = ([], tm); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 672 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 673 | fun record_update_tr' tm = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 674 | let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 675 | if null ts then raise Match | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 676 | else Syntax.const "_record_update" $ u $ | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 677 | foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 678 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 679 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 680 | fun gen_field_tr' sfx tr' name = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 681 | let val name_sfx = suffix sfx name | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 682 | in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 683 | |
| 21772 | 684 | fun record_tr' sep mark record record_scheme unit ctxt t = | 
| 17261 | 685 | let | 
| 21772 | 686 | val thy = ProofContext.theory_of ctxt; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 687 | fun field_lst t = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 688 | (case strip_comb t of | 
| 17600 | 689 | (Const (ext,_),args as (_::_)) | 
| 18858 | 690 | => (case try (unsuffix extN) (Sign.intern_const thy ext) of | 
| 17261 | 691 | SOME ext' | 
| 18858 | 692 | => (case get_extfields thy ext' of | 
| 17261 | 693 | SOME flds | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 694 | => (let | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 695 | val (f::fs) = but_last (map fst flds); | 
| 18858 | 696 | val flds' = Sign.extern_const thy f :: map NameSpace.base fs; | 
| 17261 | 697 | val (args',more) = split_last args; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 698 | in (flds'~~args')@field_lst more end | 
| 19841 | 699 |                          handle Library.UnequalLengths => [("",t)])
 | 
| 15531 | 700 |                    | NONE => [("",t)])
 | 
| 701 |              | NONE => [("",t)])
 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 702 |        | _ => [("",t)])
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 703 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 704 | val (flds,(_,more)) = split_last (field_lst t); | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 705 | val _ = if null flds then raise Match else (); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 706 | val flds' = map (fn (n,t)=>Syntax.const mark$Syntax.const n$t) flds; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 707 | val flds'' = foldr1 (fn (x,y) => Syntax.const sep$x$y) flds'; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 708 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 709 | in if unit more | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 710 | then Syntax.const record$flds'' | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 711 | else Syntax.const record_scheme$flds''$more | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 712 | end | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 713 | |
| 17261 | 714 | fun gen_record_tr' name = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 715 | let val name_sfx = suffix extN name; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 716 |       val unit = (fn Const ("Unity",_) => true | _ => false);
 | 
| 21772 | 717 | fun tr' ctxt ts = record_tr' "_fields" "_field" "_record" "_record_scheme" unit ctxt | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 718 | (list_comb (Syntax.const name_sfx,ts)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 719 | in (name_sfx,tr') | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 720 | end | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 721 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 722 | fun print_translation names = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 723 | map (gen_field_tr' updateN record_update_tr') names; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 724 | |
| 19748 | 725 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 726 | (* record_type_abbr_tr' tries to reconstruct the record name type abbreviation from *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 727 | (* the (nested) extension types. *) | 
| 21772 | 728 | fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt tm = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 729 | let | 
| 21772 | 730 | val thy = ProofContext.theory_of ctxt; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 731 | (* tm is term representation of a (nested) field type. We first reconstruct the *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 732 | (* type from tm so that we can continue on the type level rather then the term level.*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 733 | |
| 18858 | 734 | fun get_sort xs n = AList.lookup (op =) xs n |> the_default (Sign.defaultS thy); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 735 | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 736 | (* WORKAROUND: | 
| 17261 | 737 | * If a record type occurs in an error message of type inference there | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 738 | * may be some internal frees donoted by ??: | 
| 17261 | 739 |        * (Const "_tfree",_)$Free ("??'a",_).
 | 
| 740 | ||
| 741 |        * This will unfortunately be translated to Type ("??'a",[]) instead of
 | |
| 742 |        * TFree ("??'a",_) by typ_of_term, which will confuse unify below.
 | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 743 | * fixT works around. | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 744 | *) | 
| 17261 | 745 | fun fixT (T as Type (x,[])) = | 
| 18858 | 746 | if String.isPrefix "??'" x then TFree (x,Sign.defaultS thy) else T | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 747 | | fixT (Type (x,xs)) = Type (x,map fixT xs) | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 748 | | fixT T = T; | 
| 17261 | 749 | |
| 18858 | 750 | val T = fixT (Sign.intern_typ thy | 
| 17261 | 751 | (Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts tm)) I tm)); | 
| 19748 | 752 | val midx = maxidx_of_typ T; | 
| 753 | val varifyT = varifyT midx; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 754 | |
| 17261 | 755 | fun mk_type_abbr subst name alphas = | 
| 19748 | 756 | let val abbrT = Type (name, map (fn a => varifyT (TFree (a, Sign.defaultS thy))) alphas); | 
| 17261 | 757 | in Syntax.term_of_typ (! Syntax.show_sorts) | 
| 18858 | 758 | (Sign.extern_typ thy (Envir.norm_type subst abbrT)) end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 759 | |
| 19748 | 760 | fun match rT T = (Sign.typ_match thy (varifyT rT,T) | 
| 761 | Vartab.empty); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 762 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 763 | in if !print_record_type_abbr | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 764 | then (case last_extT T of | 
| 17261 | 765 | SOME (name,_) | 
| 766 | => if name = lastExt | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 767 | then | 
| 17261 | 768 | (let | 
| 19748 | 769 | val subst = match schemeT T | 
| 17261 | 770 | in | 
| 19748 | 771 | if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree(zeta,Sign.defaultS thy)))) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 772 | then mk_type_abbr subst abbr alphas | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 773 | else mk_type_abbr subst (suffix schemeN abbr) (alphas@[zeta]) | 
| 21772 | 774 | end handle TYPE_MATCH => default_tr' ctxt tm) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 775 | else raise Match (* give print translation of specialised record a chance *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 776 | | _ => raise Match) | 
| 21772 | 777 | else default_tr' ctxt tm | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 778 | end | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 779 | |
| 21772 | 780 | fun record_type_tr' sep mark record record_scheme ctxt t = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 781 | let | 
| 21772 | 782 | val thy = ProofContext.theory_of ctxt; | 
| 18858 | 783 | fun get_sort xs n = AList.lookup (op =) xs n |> the_default (Sign.defaultS thy); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 784 | |
| 18858 | 785 | val T = Sign.intern_typ thy (Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t) | 
| 19748 | 786 | val varifyT = varifyT (Term.maxidx_of_typ T) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 787 | |
| 18858 | 788 | fun term_of_type T = Syntax.term_of_typ (!Syntax.show_sorts) (Sign.extern_typ thy T); | 
| 17261 | 789 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 790 | fun field_lst T = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 791 | (case T of | 
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 792 | Type (ext, args) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 793 | => (case try (unsuffix ext_typeN) ext of | 
| 17261 | 794 | SOME ext' | 
| 18858 | 795 | => (case get_extfields thy ext' of | 
| 17261 | 796 | SOME flds | 
| 18858 | 797 | => (case get_fieldext thy (fst (hd flds)) of | 
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 798 | SOME (_, alphas) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 799 | => (let | 
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 800 | val (f :: fs) = but_last flds; | 
| 18858 | 801 | val flds' = apfst (Sign.extern_const thy) f | 
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 802 | :: map (apfst NameSpace.base) fs; | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 803 | val (args', more) = split_last args; | 
| 19748 | 804 | val alphavars = map varifyT (but_last alphas); | 
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 805 | val subst = fold2 (curry (Sign.typ_match thy)) | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 806 | alphavars args' Vartab.empty; | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 807 | val flds'' = (map o apsnd) | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 808 | (Envir.norm_type subst o varifyT) flds'; | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 809 | in flds'' @ field_lst more end | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 810 |                               handle TYPE_MATCH => [("", T)]
 | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 811 |                                   | Library.UnequalLengths => [("", T)])
 | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 812 |                          | NONE => [("", T)])
 | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 813 |                    | NONE => [("", T)])
 | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 814 |              | NONE => [("", T)])
 | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 815 |         | _ => [("", T)])
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 816 | |
| 22219 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 817 | val (flds, (_, moreT)) = split_last (field_lst T); | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 818 | val flds' = map (fn (n, T) => Syntax.const mark $ Syntax.const n $ term_of_type T) flds; | 
| 
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
 haftmann parents: 
21962diff
changeset | 819 | val flds'' = foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds' handle Empty => raise Match; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 820 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 821 | in if not (!print_record_type_as_fields) orelse null flds then raise Match | 
| 17261 | 822 | else if moreT = HOLogic.unitT | 
| 823 | then Syntax.const record$flds'' | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 824 | else Syntax.const record_scheme$flds''$term_of_type moreT | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 825 | end | 
| 17261 | 826 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 827 | |
| 17261 | 828 | fun gen_record_type_tr' name = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 829 | let val name_sfx = suffix ext_typeN name; | 
| 21772 | 830 | fun tr' ctxt ts = record_type_tr' "_field_types" "_field_type" | 
| 831 | "_record_type" "_record_type_scheme" ctxt | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 832 | (list_comb (Syntax.const name_sfx,ts)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 833 | in (name_sfx,tr') | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 834 | end | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 835 | |
| 17261 | 836 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 837 | fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 838 | let val name_sfx = suffix ext_typeN name; | 
| 17261 | 839 | val default_tr' = record_type_tr' "_field_types" "_field_type" | 
| 840 | "_record_type" "_record_type_scheme" | |
| 21772 | 841 | fun tr' ctxt ts = | 
| 842 | record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt | |
| 18858 | 843 | (list_comb (Syntax.const name_sfx,ts)) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 844 | in (name_sfx, tr') end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 845 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 846 | (** record simprocs **) | 
| 14358 | 847 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 848 | val record_quick_and_dirty_sensitive = ref false; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 849 | |
| 15215 | 850 | |
| 18858 | 851 | fun quick_and_dirty_prove stndrd thy asms prop tac = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 852 | if !record_quick_and_dirty_sensitive andalso !quick_and_dirty | 
| 20049 | 853 | then Goal.prove (ProofContext.init thy) [] [] | 
| 854 | (Logic.list_implies (map Logic.varify asms,Logic.varify prop)) | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 855 | (K (SkipProof.cheat_tac HOL.thy)) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 856 | (* standard can take quite a while for large records, thats why | 
| 17261 | 857 | * we varify the proposition manually here.*) | 
| 20049 | 858 | else let val prf = Goal.prove (ProofContext.init thy) [] asms prop tac; | 
| 17261 | 859 | in if stndrd then standard prf else prf end; | 
| 15215 | 860 | |
| 17261 | 861 | fun quick_and_dirty_prf noopt opt () = | 
| 862 | if !record_quick_and_dirty_sensitive andalso !quick_and_dirty | |
| 15215 | 863 | then noopt () | 
| 864 | else opt (); | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 865 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 866 | local | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 867 | fun abstract_over_fun_app (Abs (f,fT,t)) = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 868 | let | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 869 | val (f',t') = Term.dest_abs (f,fT,t); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 870 | val T = domain_type fT; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 871 |      val (x,T') = hd (Term.variant_frees t' [("x",T)]);
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 872 | val f_x = Free (f',fT)$(Free (x,T')); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 873 | fun is_constr (Const (c,_)$_) = can (unsuffix extN) c | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 874 | | is_constr _ = false; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 875 | fun subst (t as u$w) = if Free (f',fT)=u | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 876 | then if is_constr w then f_x | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 877 |                                  else raise TERM ("abstract_over_fun_app",[t])
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 878 | else subst u$subst w | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 879 | | subst (Abs (x,T,t)) = (Abs (x,T,subst t)) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 880 | | subst t = t | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 881 | val t'' = abstract_over (f_x,subst t'); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 882 | val vars = strip_qnt_vars "all" t''; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 883 | val bdy = strip_qnt_body "all" t''; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 884 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 885 | in list_abs ((x,T')::vars,bdy) end | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 886 |   | abstract_over_fun_app t = raise TERM ("abstract_over_fun_app",[t]);
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 887 | (* Generates a theorem of the kind: | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 888 | * !!f x*. PROP P (f ( r x* ) x* == !!r x*. PROP P r x* | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 889 | *) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 890 | fun mk_fun_apply_eq (Abs (f, fT, t)) thy = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 891 | let | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 892 | val rT = domain_type fT; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 893 | val vars = Term.strip_qnt_vars "all" t; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 894 | val Ts = map snd vars; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 895 | val n = length vars; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 896 | fun app_bounds 0 t = t$Bound 0 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 897 | | app_bounds n t = if n > 0 then app_bounds (n-1) (t$Bound n) else t | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 898 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 899 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 900 |     val [P,r] = Term.variant_frees t [("P",rT::Ts--->Term.propT),("r",Ts--->rT)];
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 901 | val prop = Logic.mk_equals | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 902 | (list_all ((f,fT)::vars, | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 903 | app_bounds (n - 1) ((Free P)$(Bound n$app_bounds (n-1) (Free r)))), | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 904 | list_all ((fst r,rT)::vars, | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 905 | app_bounds (n - 1) ((Free P)$Bound n))); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 906 | val prove_standard = quick_and_dirty_prove true thy; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 907 | val thm = prove_standard [] prop (fn prems => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 908 | EVERY [rtac equal_intr_rule 1, | 
| 21687 | 909 | Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1, | 
| 910 | Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1]); | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 911 | in thm end | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 912 |   | mk_fun_apply_eq t thy = raise TERM ("mk_fun_apply_eq",[t]);
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 913 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 914 | in | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 915 | (* During proof of theorems produced by record_simproc you can end up in | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 916 | * situations like "!!f ... . ... f r ..." where f is an extension update function. | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 917 | * In order to split "f r" we transform this to "!!r ... . ... r ..." so that the | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 918 | * usual split rules for extensions can apply. | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 919 | *) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 920 | val record_split_f_more_simproc = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 921 | Simplifier.simproc HOL.thy "record_split_f_more_simp" ["x"] | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 922 | (fn thy => fn _ => fn t => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 923 |       (case t of (Const ("all", Type (_, [Type (_, [Type("fun",[T,T']), _]), _])))$
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 924 | (trm as Abs _) => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 925 | (case rec_id (~1) T of | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 926 | "" => NONE | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 927 | | n => if T=T' | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 928 | then (let | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 929 | val P=cterm_of thy (abstract_over_fun_app trm); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 930 | val thm = mk_fun_apply_eq trm thy; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 931 | val PV = cterm_of thy (hd (term_vars (prop_of thm))); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 932 | val thm' = cterm_instantiate [(PV,P)] thm; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 933 | in SOME thm' end handle TERM _ => NONE) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 934 | else NONE) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 935 | | _ => NONE)) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 936 | end | 
| 14255 | 937 | |
| 18858 | 938 | fun prove_split_simp thy ss T prop = | 
| 17261 | 939 | let | 
| 18858 | 940 |     val {sel_upd={simpset,...},extsplit,...} = RecordsData.get thy;
 | 
| 17261 | 941 | val extsplits = | 
| 17412 | 942 | Library.foldl (fn (thms,(n,_)) => the_list (Symtab.lookup extsplit n) @ thms) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 943 | ([],dest_recTs T); | 
| 18858 | 944 | val thms = (case get_splits thy (rec_id (~1) T) of | 
| 17261 | 945 | SOME (all_thm,_,_,_) => | 
| 15203 | 946 | all_thm::(case extsplits of [thm] => [] | _ => extsplits) | 
| 947 | (* [thm] is the same as all_thm *) | |
| 17261 | 948 | | NONE => extsplits) | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 949 | val thms'=K_comp_convs@thms; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 950 | val ss' = (Simplifier.inherit_context ss simpset | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 951 | addsimps thms' | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 952 | addsimprocs [record_split_f_more_simproc]); | 
| 16973 | 953 | in | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 954 | quick_and_dirty_prove true thy [] prop (fn _ => simp_tac ss' 1) | 
| 15203 | 955 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 956 | |
| 15215 | 957 | |
| 15059 | 958 | local | 
| 16822 | 959 | fun eq (s1:string) (s2:string) = (s1 = s2); | 
| 960 | fun has_field extfields f T = | |
| 18931 | 961 | exists (fn (eN,_) => exists (eq f o fst) (Symtab.lookup_list extfields eN)) | 
| 16822 | 962 | (dest_recTs T); | 
| 15059 | 963 | in | 
| 14255 | 964 | (* record_simproc *) | 
| 965 | (* Simplifies selections of an record update: | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 966 | * (1) S (S_update k r) = k (S r) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 967 | * (2) S (X_update k r) = S r | 
| 14255 | 968 | * The simproc skips multiple updates at once, eg: | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 969 | * S (X_update x (Y_update y (S_update k r))) = k (S r) | 
| 14255 | 970 | * But be careful in (2) because of the extendibility of records. | 
| 971 | * - If S is a more-selector we have to make sure that the update on component | |
| 972 | * X does not affect the selected subrecord. | |
| 973 | * - If X is a more-selector we have to make sure that S is not in the updated | |
| 17261 | 974 | * subrecord. | 
| 14255 | 975 | *) | 
| 13462 | 976 | val record_simproc = | 
| 17616 
f526e2b9fe9e
simprocs: pattern now "x" (the proc is supposed to discriminate faster than Pattern.match);
 wenzelm parents: 
17600diff
changeset | 977 | Simplifier.simproc HOL.thy "record_simp" ["x"] | 
| 18858 | 978 | (fn thy => fn ss => fn t => | 
| 16872 | 979 | (case t of (sel as Const (s, Type (_,[domS,rangeS])))$ | 
| 980 | ((upd as Const (u,Type(_,[_,Type (_,[rT,_])]))) $ k $ r)=> | |
| 18858 | 981 | if is_selector thy s then | 
| 982 | (case get_updates thy u of SOME u_name => | |
| 13462 | 983 | let | 
| 18858 | 984 |               val {sel_upd={updates,...},extfields,...} = RecordsData.get thy;
 | 
| 17261 | 985 | |
| 16872 | 986 | fun mk_eq_terms ((upd as Const (u,Type(_,[kT,_]))) $ k $ r) = | 
| 17412 | 987 | (case Symtab.lookup updates u of | 
| 15531 | 988 | NONE => NONE | 
| 17261 | 989 | | SOME u_name | 
| 14255 | 990 | => if u_name = s | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 991 | then (case mk_eq_terms r of | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 992 | NONE => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 993 | let | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 994 |                                    val rv = ("r",rT)
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 995 | val rb = Bound 0 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 996 |                                    val kv = ("k",kT)
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 997 | val kb = Bound 1 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 998 | in SOME (upd$kb$rb,kb$(sel$rb),[kv,rv]) end | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 999 | | SOME (trm,trm',vars) => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1000 | let | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1001 |                                    val kv = ("k",kT)
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1002 | val kb = Bound (length vars) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1003 | in SOME (upd$kb$trm,kb$trm',kv::vars) end) | 
| 16822 | 1004 | else if has_field extfields u_name rangeS | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1005 | orelse has_field extfields s (domain_type kT) | 
| 15531 | 1006 | then NONE | 
| 17261 | 1007 | else (case mk_eq_terms r of | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1008 | SOME (trm,trm',vars) | 
| 17261 | 1009 | => let | 
| 1010 |                                           val kv = ("k",kT)
 | |
| 14255 | 1011 | val kb = Bound (length vars) | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1012 | in SOME (upd$kb$trm,trm',kv::vars) end | 
| 15531 | 1013 | | NONE | 
| 17261 | 1014 | => let | 
| 1015 |                                           val rv = ("r",rT)
 | |
| 14255 | 1016 | val rb = Bound 0 | 
| 16872 | 1017 |                                           val kv = ("k",kT)
 | 
| 17261 | 1018 | val kb = Bound 1 | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1019 | in SOME (upd$kb$rb,sel$rb,[kv,rv]) end)) | 
| 17261 | 1020 | | mk_eq_terms r = NONE | 
| 13462 | 1021 | in | 
| 17261 | 1022 | (case mk_eq_terms (upd$k$r) of | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1023 | SOME (trm,trm',vars) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1024 | => SOME (prove_split_simp thy ss domS | 
| 16872 | 1025 | (list_all(vars,(equals rangeS$(sel$trm)$trm')))) | 
| 15531 | 1026 | | NONE => NONE) | 
| 13462 | 1027 | end | 
| 15531 | 1028 | | NONE => NONE) | 
| 17510 | 1029 | else NONE | 
| 15531 | 1030 | | _ => NONE)); | 
| 7178 | 1031 | |
| 17261 | 1032 | (* record_upd_simproc *) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1033 | (* simplify multiple updates: | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1034 | * (1) "N_update y (M_update g (N_update x (M_update f r))) = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1035 | (N_update (y o x) (M_update (g o f) r))" | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1036 | * (2) "r(|M:= M r|) = r" | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1037 | * For (2) special care of "more" updates has to be taken: | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1038 | * r(|more := m; A := A r|) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1039 | * If A is contained in the fields of m we cannot remove the update A := A r! | 
| 17261 | 1040 | * (But r(|more := r; A := A (r(|more := r|))|) = r(|more := r|) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1041 | *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1042 | val record_upd_simproc = | 
| 17616 
f526e2b9fe9e
simprocs: pattern now "x" (the proc is supposed to discriminate faster than Pattern.match);
 wenzelm parents: 
17600diff
changeset | 1043 | Simplifier.simproc HOL.thy "record_upd_simp" ["x"] | 
| 18858 | 1044 | (fn thy => fn ss => fn t => | 
| 16872 | 1045 | (case t of ((upd as Const (u, Type(_,[_,Type(_,[rT,_])]))) $ k $ r) => | 
| 17261 | 1046 |          let datatype ('a,'b) calc = Init of 'b | Inter of 'a
 | 
| 18858 | 1047 |              val {sel_upd={selectors,updates,...},extfields,...} = RecordsData.get thy;
 | 
| 17261 | 1048 | |
| 1049 | (*fun mk_abs_var x t = (x, fastype_of t);*) | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1050 | fun sel_name u = NameSpace.base (unsuffix updateN u); | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1051 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1052 | fun seed s (upd as Const (more,Type(_,[mT,_]))$ k $ r) = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1053 | if has_field extfields s (domain_type' mT) then upd else seed s r | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1054 | | seed _ r = r; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1055 | |
| 17261 | 1056 | fun grow u uT k kT vars (sprout,skeleton) = | 
| 1057 | if sel_name u = moreN | |
| 16872 | 1058 |                    then let val kv = ("k", kT);
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1059 | val kb = Bound (length vars); | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1060 | in ((Const (u,uT)$k$sprout,Const (u,uT)$kb$skeleton),kv::vars) end | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1061 | else ((sprout,skeleton),vars); | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1062 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1063 | fun is_upd_same (sprout,skeleton) u | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1064 |                                 ((K_rec as Const ("Record.K_record",_))$
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1065 | ((sel as Const (s,_))$r)) = | 
| 17261 | 1066 | if (unsuffix updateN u) = s andalso (seed s sprout) = r | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1067 | then SOME (K_rec,sel,seed s skeleton) | 
| 15531 | 1068 | else NONE | 
| 1069 | | is_upd_same _ _ _ = NONE | |
| 17261 | 1070 | |
| 16872 | 1071 |              fun init_seed r = ((r,Bound 0), [("r", rT)]);
 | 
| 17261 | 1072 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1073 | fun add (n:string) f fmaps = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1074 | (case AList.lookup (op =) fmaps n of | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1075 | NONE => AList.update (op =) (n,[f]) fmaps | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1076 | | SOME fs => AList.update (op =) (n,f::fs) fmaps) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1077 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1078 | fun comps (n:string) T fmaps = | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1079 | (case AList.lookup (op =) fmaps n of | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1080 | SOME fs => | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1081 |                    foldr1 (fn (f,g) => Const ("Fun.comp",(T-->T)-->(T-->T)-->(T-->T))$f$g) fs
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1082 |                 | NONE => error ("record_upd_simproc.comps"))
 | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1083 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1084 | (* mk_updterm returns either | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1085 | * - Init (orig-term, orig-term-skeleton, vars) if no optimisation can be made, | 
| 17261 | 1086 | * where vars are the bound variables in the skeleton | 
| 1087 | * - Inter (orig-term-skeleton,simplified-term-skeleton, | |
| 16872 | 1088 | * vars, (term-sprout, skeleton-sprout)) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1089 | * where "All vars. orig-term-skeleton = simplified-term-skeleton" is | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1090 | * the desired simplification rule, | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1091 | * the sprouts accumulate the "more-updates" on the way from the seed | 
| 17261 | 1092 | * to the outermost update. It is only relevant to calculate the | 
| 1093 | * possible simplification for (2) | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1094 | * The algorithm first walks down the updates to the seed-record while | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1095 | * memorising the updates in the already-table. While walking up the | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1096 | * updates again, the optimised term is constructed. | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1097 | *) | 
| 17261 | 1098 | fun mk_updterm upds already | 
| 16872 | 1099 | (t as ((upd as Const (u,uT as (Type (_,[kT,_])))) $ k $ r)) = | 
| 17261 | 1100 | if Symtab.defined upds u | 
| 1101 | then let | |
| 1102 | fun rest already = mk_updterm upds already | |
| 1103 | in if u mem_string already | |
| 1104 | then (case (rest already r) of | |
| 1105 | Init ((sprout,skel),vars) => | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1106 | let | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1107 | val n = sel_name u; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1108 | val kv = (n, kT); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1109 | val kb = Bound (length vars); | 
| 16872 | 1110 | val (sprout',vars')= grow u uT k kT (kv::vars) (sprout,skel); | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1111 | in Inter (upd$kb$skel,skel,vars',add n kb [],sprout') end | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1112 | | Inter (trm,trm',vars,fmaps,sprout) => | 
| 17261 | 1113 | let | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1114 | val n = sel_name u; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1115 | val kv = (n, kT); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1116 | val kb = Bound (length vars); | 
| 16872 | 1117 | val (sprout',vars') = grow u uT k kT (kv::vars) sprout; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1118 | in Inter(upd$kb$trm,trm',kv::vars',add n kb fmaps,sprout') | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1119 | end) | 
| 17261 | 1120 | else | 
| 1121 | (case rest (u::already) r of | |
| 1122 | Init ((sprout,skel),vars) => | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1123 | (case is_upd_same (sprout,skel) u k of | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1124 | SOME (K_rec,sel,skel') => | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1125 | let | 
| 17261 | 1126 | val (sprout',vars') = grow u uT k kT vars (sprout,skel); | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1127 | in Inter(upd$(K_rec$(sel$skel'))$skel,skel,vars',[],sprout') | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1128 | end | 
| 17261 | 1129 | | NONE => | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1130 | let | 
| 17261 | 1131 | val kv = (sel_name u, kT); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1132 | val kb = Bound (length vars); | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1133 | in Init ((upd$k$sprout,upd$kb$skel),kv::vars) end) | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1134 | | Inter (trm,trm',vars,fmaps,sprout) => | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1135 | (case is_upd_same sprout u k of | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1136 | SOME (K_rec,sel,skel) => | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1137 | let | 
| 16872 | 1138 | val (sprout',vars') = grow u uT k kT vars sprout | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1139 | in Inter(upd$(K_rec$(sel$skel))$trm,trm',vars',fmaps,sprout') | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1140 | end | 
| 15531 | 1141 | | NONE => | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1142 | let | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1143 | val n = sel_name u | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1144 | val T = domain_type kT | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1145 | val kv = (n, kT) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1146 | val kb = Bound (length vars) | 
| 16872 | 1147 | val (sprout',vars') = grow u uT k kT (kv::vars) sprout | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1148 | val fmaps' = add n kb fmaps | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1149 | in Inter (upd$kb$trm,upd$comps n T fmaps'$trm' | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1150 | ,vars',fmaps',sprout') end)) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1151 | end | 
| 17261 | 1152 | else Init (init_seed t) | 
| 1153 | | mk_updterm _ _ t = Init (init_seed t); | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1154 | |
| 17261 | 1155 | in (case mk_updterm updates [] t of | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1156 | Inter (trm,trm',vars,_,_) | 
| 18858 | 1157 | => SOME (prove_split_simp thy ss rT | 
| 16872 | 1158 | (list_all(vars,(equals rT$trm$trm')))) | 
| 15531 | 1159 | | _ => NONE) | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1160 | end | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1161 | | _ => NONE)) | 
| 15059 | 1162 | end | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1163 | |
| 14255 | 1164 | (* record_eq_simproc *) | 
| 1165 | (* looks up the most specific record-equality. | |
| 1166 | * Note on efficiency: | |
| 1167 | * Testing equality of records boils down to the test of equality of all components. | |
| 1168 | * Therefore the complexity is: #components * complexity for single component. | |
| 1169 | * Especially if a record has a lot of components it may be better to split up | |
| 1170 | * the record first and do simplification on that (record_split_simp_tac). | |
| 1171 | * e.g. r(|lots of updates|) = x | |
| 1172 | * | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1173 | * record_eq_simproc record_split_simp_tac | 
| 17261 | 1174 | * Complexity: #components * #updates #updates | 
| 1175 | * | |
| 14255 | 1176 | *) | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1177 | val record_eq_simproc = | 
| 17510 | 1178 | Simplifier.simproc HOL.thy "record_eq_simp" ["r = s"] | 
| 18858 | 1179 | (fn thy => fn _ => fn t => | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1180 |       (case t of Const ("op =", Type (_, [T, _])) $ _ $ _ =>
 | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1181 | (case rec_id (~1) T of | 
| 15531 | 1182 | "" => NONE | 
| 18858 | 1183 | | name => (case get_equalities thy name of | 
| 15531 | 1184 | NONE => NONE | 
| 1185 | | SOME thm => SOME (thm RS Eq_TrueI))) | |
| 1186 | | _ => NONE)); | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1187 | |
| 14255 | 1188 | (* record_split_simproc *) | 
| 17261 | 1189 | (* splits quantified occurrences of records, for which P holds. P can peek on the | 
| 14255 | 1190 | * subterm starting at the quantified occurrence of the record (including the quantifier) | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1191 | * P t = 0: do not split | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1192 | * P t = ~1: completely split | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1193 | * P t > 0: split up to given bound of record extensions | 
| 14255 | 1194 | *) | 
| 1195 | fun record_split_simproc P = | |
| 17616 
f526e2b9fe9e
simprocs: pattern now "x" (the proc is supposed to discriminate faster than Pattern.match);
 wenzelm parents: 
17600diff
changeset | 1196 | Simplifier.simproc HOL.thy "record_split_simp" ["x"] | 
| 18858 | 1197 | (fn thy => fn _ => fn t => | 
| 14255 | 1198 | (case t of (Const (quantifier, Type (_, [Type (_, [T, _]), _])))$trm => | 
| 1199 | if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex" | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1200 | then (case rec_id (~1) T of | 
| 15531 | 1201 | "" => NONE | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1202 | | name | 
| 17261 | 1203 | => let val split = P t | 
| 1204 | in if split <> 0 then | |
| 18858 | 1205 | (case get_splits thy (rec_id split T) of | 
| 15531 | 1206 | NONE => NONE | 
| 17261 | 1207 | | SOME (all_thm, All_thm, Ex_thm,_) | 
| 15531 | 1208 | => SOME (case quantifier of | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1209 | "all" => all_thm | 
| 21546 | 1210 | | "All" => All_thm RS eq_reflection | 
| 1211 | | "Ex" => Ex_thm RS eq_reflection | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1212 | | _ => error "record_split_simproc")) | 
| 15531 | 1213 | else NONE | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1214 | end) | 
| 15531 | 1215 | else NONE | 
| 1216 | | _ => NONE)) | |
| 7178 | 1217 | |
| 14427 | 1218 | val record_ex_sel_eq_simproc = | 
| 17510 | 1219 | Simplifier.simproc HOL.thy "record_ex_sel_eq_simproc" ["Ex t"] | 
| 18858 | 1220 | (fn thy => fn ss => fn t => | 
| 17261 | 1221 | let | 
| 16973 | 1222 | fun prove prop = | 
| 18858 | 1223 | quick_and_dirty_prove true thy [] prop | 
| 1224 | (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy) | |
| 16973 | 1225 | addsimps simp_thms addsimprocs [record_split_simproc (K ~1)]) 1); | 
| 14959 | 1226 | |
| 1227 | fun mkeq (lr,Teq,(sel,Tsel),x) i = | |
| 18858 | 1228 | if is_selector thy sel then | 
| 17261 | 1229 | let val x' = if not (loose_bvar1 (x,0)) | 
| 1230 |                               then Free ("x" ^ string_of_int i, range_type Tsel)
 | |
| 14959 | 1231 |                               else raise TERM ("",[x]);
 | 
| 1232 | val sel' = Const (sel,Tsel)$Bound 0; | |
| 1233 | val (l,r) = if lr then (sel',x') else (x',sel'); | |
| 1234 |                   in Const ("op =",Teq)$l$r end
 | |
| 17510 | 1235 |               else raise TERM ("",[Const (sel,Tsel)]);
 | 
| 14959 | 1236 | |
| 17261 | 1237 |          fun dest_sel_eq (Const ("op =",Teq)$(Const (sel,Tsel)$Bound 0)$X) =
 | 
| 14959 | 1238 | (true,Teq,(sel,Tsel),X) | 
| 1239 |            | dest_sel_eq (Const ("op =",Teq)$X$(Const (sel,Tsel)$Bound 0)) =
 | |
| 1240 | (false,Teq,(sel,Tsel),X) | |
| 1241 |            | dest_sel_eq _ = raise TERM ("",[]);
 | |
| 1242 | ||
| 17261 | 1243 | in | 
| 1244 | (case t of | |
| 14959 | 1245 |            (Const ("Ex",Tex)$Abs(s,T,t)) =>
 | 
| 16872 | 1246 | (let val eq = mkeq (dest_sel_eq t) 0; | 
| 14959 | 1247 |                  val prop = list_all ([("r",T)],
 | 
| 1248 |                               Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
 | |
| 1249 | HOLogic.true_const)); | |
| 15531 | 1250 | in SOME (prove prop) end | 
| 16872 | 1251 | handle TERM _ => NONE) | 
| 17261 | 1252 | | _ => NONE) | 
| 14427 | 1253 | end) | 
| 1254 | ||
| 5698 | 1255 | |
| 17261 | 1256 | |
| 6358 | 1257 | |
| 14255 | 1258 | local | 
| 1259 | val inductive_atomize = thms "induct_atomize"; | |
| 18464 | 1260 | val inductive_rulify = thms "induct_rulify"; | 
| 14255 | 1261 | in | 
| 1262 | (* record_split_simp_tac *) | |
| 17261 | 1263 | (* splits (and simplifies) all records in the goal for which P holds. | 
| 14255 | 1264 | * For quantified occurrences of a record | 
| 1265 | * P can peek on the whole subterm (including the quantifier); for free variables P | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1266 | * can only peek on the variable itself. | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1267 | * P t = 0: do not split | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1268 | * P t = ~1: completely split | 
| 17261 | 1269 | * P t > 0: split up to given bound of record extensions | 
| 14255 | 1270 | *) | 
| 15248 
b436486091a6
record_split_simp_tac now can get simp rules as parameter
 schirmer parents: 
15215diff
changeset | 1271 | fun record_split_simp_tac thms P i st = | 
| 14255 | 1272 | let | 
| 18858 | 1273 | val thy = Thm.theory_of_thm st; | 
| 14255 | 1274 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1275 | val has_rec = exists_Const | 
| 14255 | 1276 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 17261 | 1277 | (s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T | 
| 14255 | 1278 | | _ => false); | 
| 1279 | ||
| 18011 
685d95c793ff
cleaned up nth, nth_update, nth_map and nth_string functions
 haftmann parents: 
17960diff
changeset | 1280 | val goal = nth (Thm.prems_of st) (i - 1); | 
| 15570 | 1281 | val frees = List.filter (is_recT o type_of) (term_frees goal); | 
| 14255 | 1282 | |
| 17261 | 1283 | fun mk_split_free_tac free induct_thm i = | 
| 18858 | 1284 | let val cfree = cterm_of thy free; | 
| 14255 | 1285 | val (_$(_$r)) = concl_of induct_thm; | 
| 18858 | 1286 | val crec = cterm_of thy r; | 
| 14255 | 1287 | val thm = cterm_instantiate [(crec,cfree)] induct_thm; | 
| 1288 | in EVERY [simp_tac (HOL_basic_ss addsimps inductive_atomize) i, | |
| 1289 | rtac thm i, | |
| 18464 | 1290 | simp_tac (HOL_basic_ss addsimps inductive_rulify) i] | 
| 17261 | 1291 | end; | 
| 14255 | 1292 | |
| 17261 | 1293 | fun split_free_tac P i (free as Free (n,T)) = | 
| 1294 | (case rec_id (~1) T of | |
| 15531 | 1295 | "" => NONE | 
| 17261 | 1296 | | name => let val split = P free | 
| 1297 | in if split <> 0 then | |
| 18858 | 1298 | (case get_splits thy (rec_id split T) of | 
| 15531 | 1299 | NONE => NONE | 
| 1300 | | SOME (_,_,_,induct_thm) | |
| 1301 | => SOME (mk_split_free_tac free induct_thm i)) | |
| 1302 | else NONE | |
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1303 | end) | 
| 15531 | 1304 | | split_free_tac _ _ _ = NONE; | 
| 14255 | 1305 | |
| 15570 | 1306 | val split_frees_tacs = List.mapPartial (split_free_tac P i) frees; | 
| 17261 | 1307 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1308 | val simprocs = if has_rec goal then [record_split_simproc P] else []; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1309 | val thms' = K_comp_convs@thms | 
| 17261 | 1310 | in st |> ((EVERY split_frees_tacs) | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1311 | THEN (Simplifier.full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i)) | 
| 15570 | 1312 | end handle Empty => Seq.empty; | 
| 14255 | 1313 | end; | 
| 1314 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1315 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1316 | (* record_split_tac *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1317 | (* splits all records in the goal, which are quantified by ! or !!. *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1318 | fun record_split_tac i st = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1319 | let | 
| 18858 | 1320 | val thy = Thm.theory_of_thm st; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1321 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1322 | val has_rec = exists_Const | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1323 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 17261 | 1324 | (s = "all" orelse s = "All") andalso is_recT T | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1325 | | _ => false); | 
| 17261 | 1326 | |
| 18011 
685d95c793ff
cleaned up nth, nth_update, nth_map and nth_string functions
 haftmann parents: 
17960diff
changeset | 1327 | val goal = nth (Thm.prems_of st) (i - 1); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1328 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1329 | fun is_all t = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1330 | (case t of (Const (quantifier, _)$_) => | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1331 | if quantifier = "All" orelse quantifier = "all" then ~1 else 0 | 
| 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 1332 | | _ => 0); | 
| 17261 | 1333 | |
| 1334 | in if has_rec goal | |
| 1335 | then Simplifier.full_simp_tac | |
| 1336 | (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i st | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1337 | else Seq.empty | 
| 15570 | 1338 | end handle Subscript => Seq.empty; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1339 | |
| 6358 | 1340 | (* wrapper *) | 
| 1341 | ||
| 5707 | 1342 | val record_split_name = "record_split_tac"; | 
| 1343 | val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac); | |
| 5698 | 1344 | |
| 16330 | 1345 | |
| 1346 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1347 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1348 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1349 | (* prepare arguments *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1350 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1351 | fun read_raw_parent sign s = | 
| 16330 | 1352 | (case Sign.read_typ_abbrev (sign, K NONE) s handle TYPE (msg, _, _) => error msg of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1353 | Type (name, Ts) => (Ts, name) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1354 |   | _ => error ("Bad parent record specification: " ^ quote s));
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1355 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1356 | fun read_typ sign (env, s) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1357 | let | 
| 17377 | 1358 | fun def_sort (x, ~1) = AList.lookup (op =) env x | 
| 15531 | 1359 | | def_sort _ = NONE; | 
| 5060 | 1360 | 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 | 1361 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1362 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1363 | fun cert_typ sign (env, raw_T) = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1364 | 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 | 1365 | in (Term.add_typ_tfrees (T, env), T) end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1366 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1367 | (* attributes *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1368 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1369 | fun case_names_fields x = RuleCases.case_names ["fields"] x; | 
| 18728 | 1370 | fun induct_type_global name = [case_names_fields, InductAttrib.induct_type name]; | 
| 1371 | fun cases_type_global name = [case_names_fields, InductAttrib.cases_type name]; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1372 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1373 | (* tactics *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1374 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1375 | fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps)); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1376 | |
| 17261 | 1377 | (* do case analysis / induction according to rule on last parameter of ith subgoal | 
| 1378 | * (or on s if there are no parameters); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1379 | * Instatiation of record variable (and predicate) in rule is calculated to | 
| 17261 | 1380 | * avoid problems with higher order unification. | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1381 | *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1382 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1383 | fun try_param_tac s rule i st = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1384 | let | 
| 17510 | 1385 | val cert = cterm_of (Thm.theory_of_thm st); | 
| 18011 
685d95c793ff
cleaned up nth, nth_update, nth_map and nth_string functions
 haftmann parents: 
17960diff
changeset | 1386 | val g = nth (prems_of st) (i - 1); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1387 | val params = Logic.strip_params g; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1388 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g); | 
| 18145 | 1389 | val rule' = Thm.lift_rule (Thm.cprem_of st i) rule; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1390 | val (P, ys) = strip_comb (HOLogic.dest_Trueprop | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1391 | (Logic.strip_assums_concl (prop_of rule'))); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1392 | (* ca indicates if rule is a case analysis or induction rule *) | 
| 15570 | 1393 | val (x, ca) = (case rev (Library.drop (length params, ys)) of | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1394 | [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1395 | (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1396 | | [x] => (head_of x, false)); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1397 | val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of | 
| 17377 | 1398 | [] => (case AList.lookup (op =) (map dest_Free (term_frees (prop_of st))) s of | 
| 15531 | 1399 | NONE => sys_error "try_param_tac: no such variable" | 
| 1400 | | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1401 | (x, Free (s, T))]) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1402 | | (_, T) :: _ => [(P, list_abs (params, if ca then concl | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1403 | else incr_boundvars 1 (Abs (s, T, concl)))), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1404 | (x, list_abs (params, Bound 0))])) rule' | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1405 | in compose_tac (false, rule'', nprems_of rule) i st end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1406 | |
| 15215 | 1407 | |
| 1408 | (* !!x1 ... xn. ... ==> EX x1 ... xn. P x1 ... xn; | |
| 1409 | instantiates x1 ... xn with parameters x1 ... xn *) | |
| 1410 | fun ex_inst_tac i st = | |
| 1411 | let | |
| 18858 | 1412 | val thy = Thm.theory_of_thm st; | 
| 18011 
685d95c793ff
cleaned up nth, nth_update, nth_map and nth_string functions
 haftmann parents: 
17960diff
changeset | 1413 | val g = nth (prems_of st) (i - 1); | 
| 15215 | 1414 | val params = Logic.strip_params g; | 
| 18145 | 1415 | val exI' = Thm.lift_rule (Thm.cprem_of st i) exI; | 
| 15215 | 1416 | val (_$(_$x)) = Logic.strip_assums_concl (hd (prems_of exI')); | 
| 18858 | 1417 | val cx = cterm_of thy (fst (strip_comb x)); | 
| 15215 | 1418 | |
| 17261 | 1419 | in Seq.single (Library.foldl (fn (st,v) => | 
| 1420 | Seq.hd | |
| 1421 | (compose_tac (false, cterm_instantiate | |
| 18858 | 1422 | [(cx,cterm_of thy (list_abs (params,Bound v)))] exI',1) | 
| 15215 | 1423 | i st)) (st,((length params) - 1) downto 0)) | 
| 1424 | end; | |
| 1425 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1426 | fun extension_typedef name repT alphas thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1427 | let | 
| 20484 | 1428 | fun get_thms thy name = | 
| 1429 | let | |
| 1430 |         val SOME { Abs_induct = abs_induct,
 | |
| 1431 | Abs_inject=abs_inject, Abs_inverse = abs_inverse, ...} = TypedefPackage.get_info thy name; | |
| 21708 | 1432 | val rewrite_rule = MetaSimplifier.rewrite_rule [rec_UNIV_I, rec_True_simp]; | 
| 20484 | 1433 | in map rewrite_rule [abs_inject, abs_inverse, abs_induct] end; | 
| 1434 | in | |
| 1435 | thy | |
| 1436 | |> TypecopyPackage.add_typecopy (suffix ext_typeN (Sign.base_name name), alphas) repT NONE | |
| 1437 | |-> (fn (name, _) => `(fn thy => get_thms thy name)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1438 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1439 | |
| 17261 | 1440 | fun mixit convs refls = | 
| 15215 | 1441 | let fun f ((res,lhs,rhs),refl) = ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs); | 
| 15570 | 1442 | in #1 (Library.foldl f (([],[],convs),refls)) end; | 
| 15215 | 1443 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1444 | |
| 17261 | 1445 | fun extension_definition full name fields names alphas zeta moreT more vars thy = | 
| 1446 | let | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1447 | val base = Sign.base_name; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1448 | val fieldTs = (map snd fields); | 
| 15215 | 1449 | val alphas_zeta = alphas@[zeta]; | 
| 1450 | val alphas_zetaTs = map (fn n => TFree (n, HOLogic.typeS)) alphas_zeta; | |
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1451 | val vT = TFree (Name.variant alphas_zeta "'v", HOLogic.typeS); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1452 | val extT_name = suffix ext_typeN name | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1453 | val extT = Type (extT_name, alphas_zetaTs); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1454 | val repT = foldr1 HOLogic.mk_prodT (fieldTs@[moreT]); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1455 | val fields_more = fields@[(full moreN,moreT)]; | 
| 15215 | 1456 | val fields_moreTs = fieldTs@[moreT]; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1457 | val bfields_more = map (apfst base) fields_more; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1458 | val r = Free (rN,extT) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1459 | val len = length fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1460 | val idxms = 0 upto len; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1461 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1462 | (* prepare declarations and definitions *) | 
| 17261 | 1463 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1464 | (*fields constructor*) | 
| 15215 | 1465 | val ext_decl = (mk_extC (name,extT) fields_moreTs); | 
| 17261 | 1466 | (* | 
| 1467 | val ext_spec = Const ext_decl :== | |
| 1468 | (foldr (uncurry lambda) | |
| 1469 | (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more]))) (vars@[more])) | |
| 1470 | *) | |
| 1471 | val ext_spec = list_comb (Const ext_decl,vars@[more]) :== | |
| 15215 | 1472 | (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more]))); | 
| 1473 | ||
| 17261 | 1474 | fun mk_ext args = list_comb (Const ext_decl, args); | 
| 1475 | ||
| 1476 | (*destructors*) | |
| 16379 | 1477 | val _ = timing_msg "record extension preparing definitions"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1478 | val dest_decls = map (mk_selC extT o (apfst (suffix ext_dest))) bfields_more; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1479 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1480 | fun mk_dest_spec (i, (c,T)) = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1481 | let val snds = (funpow i HOLogic.mk_snd (mk_Rep name repT extT $ r)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1482 | in Const (mk_selC extT (suffix ext_dest c,T)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1483 | :== (lambda r (if i=len then snds else HOLogic.mk_fst snds)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1484 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1485 | val dest_specs = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1486 | ListPair.map mk_dest_spec (idxms, fields_more); | 
| 17261 | 1487 | |
| 16379 | 1488 | (*updates*) | 
| 15215 | 1489 | val upd_decls = map (mk_updC updN extT) bfields_more; | 
| 1490 | fun mk_upd_spec (c,T) = | |
| 1491 | let | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1492 | val args = map (fn (n,nT) => if n=c then Free (base c,T --> T)$ | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1493 | (mk_sel r (suffix ext_dest n,nT)) | 
| 17261 | 1494 | else (mk_sel r (suffix ext_dest n,nT))) | 
| 15215 | 1495 | fields_more; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1496 | in Const (mk_updC updN extT (c,T))$(Free (base c,T --> T))$r | 
| 15215 | 1497 | :== mk_ext args | 
| 1498 | end; | |
| 1499 | val upd_specs = map mk_upd_spec fields_more; | |
| 17261 | 1500 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1501 | (* 1st stage: defs_thy *) | 
| 16379 | 1502 | fun mk_defs () = | 
| 17261 | 1503 | thy | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1504 | |> extension_typedef name repT (alphas@[zeta]) | 
| 18377 | 1505 | ||> Theory.add_consts_i | 
| 15215 | 1506 | (map Syntax.no_syn ((apfst base ext_decl)::dest_decls@upd_decls)) | 
| 18377 | 1507 | ||>> PureThy.add_defs_i false (map Thm.no_attributes (ext_spec::dest_specs)) | 
| 1508 | ||>> PureThy.add_defs_i false (map Thm.no_attributes upd_specs) | |
| 1509 | |> swap | |
| 16379 | 1510 | val (defs_thy, (([abs_inject, abs_inverse, abs_induct],ext_def::dest_defs),upd_defs)) = | 
| 1511 | timeit_msg "record extension type/selector/update defs:" mk_defs; | |
| 17261 | 1512 | |
| 1513 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1514 | (* prepare propositions *) | 
| 16379 | 1515 | val _ = timing_msg "record extension preparing propositions"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1516 | val vars_more = vars@[more]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1517 | val named_vars_more = (names@[full moreN])~~vars_more; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1518 | val variants = map (fn (Free (x,_))=>x) vars_more; | 
| 15215 | 1519 | val ext = mk_ext vars_more; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1520 | val s = Free (rN, extT); | 
| 15215 | 1521 | val w = Free (wN, extT); | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1522 | val P = Free (Name.variant variants "P", extT-->HOLogic.boolT); | 
| 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1523 | val C = Free (Name.variant variants "C", HOLogic.boolT); | 
| 17261 | 1524 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1525 | val inject_prop = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1526 | let val vars_more' = map (fn (Free (x,T)) => Free (x ^ "'",T)) vars_more; | 
| 17261 | 1527 | in All (map dest_Free (vars_more@vars_more')) | 
| 1528 | ((HOLogic.eq_const extT $ | |
| 1529 | mk_ext vars_more$mk_ext vars_more') | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1530 | === | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1531 | foldr1 HOLogic.mk_conj (map HOLogic.mk_eq (vars_more ~~ vars_more'))) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1532 | end; | 
| 17261 | 1533 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1534 | val induct_prop = | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1535 | (All (map dest_Free vars_more) (Trueprop (P $ ext)), Trueprop (P $ s)); | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1536 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1537 | val cases_prop = | 
| 17261 | 1538 | (All (map dest_Free vars_more) | 
| 1539 | (Trueprop (HOLogic.mk_eq (s,ext)) ==> Trueprop C)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1540 | ==> Trueprop C; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1541 | |
| 17261 | 1542 | (*destructors*) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1543 | val dest_conv_props = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1544 | map (fn (c, x as Free (_,T)) => mk_sel ext (suffix ext_dest c,T) === x) named_vars_more; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1545 | |
| 15215 | 1546 | (*updates*) | 
| 1547 | fun mk_upd_prop (i,(c,T)) = | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1548 | let val x' = Free (Name.variant variants (base c ^ "'"),T --> T) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1549 | val args' = nth_map i (K (x'$nth vars_more i)) vars_more | 
| 15215 | 1550 | in mk_upd updN c x' ext === mk_ext args' end; | 
| 1551 | val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more); | |
| 1552 | ||
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1553 | val surjective_prop = | 
| 17261 | 1554 | let val args = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1555 | map (fn (c, Free (_,T)) => mk_sel s (suffix ext_dest c,T)) named_vars_more; | 
| 15215 | 1556 | in s === mk_ext args end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1557 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1558 | val split_meta_prop = | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1559 | let val P = Free (Name.variant variants "P", extT-->Term.propT) in | 
| 17261 | 1560 | Logic.mk_equals | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1561 | (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext)) | 
| 17261 | 1562 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1563 | |
| 17510 | 1564 | fun prove stndrd = quick_and_dirty_prove stndrd defs_thy; | 
| 1565 | val prove_standard = quick_and_dirty_prove true defs_thy; | |
| 15215 | 1566 | fun prove_simp stndrd simps = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1567 | let val tac = simp_all_tac HOL_ss simps | 
| 15215 | 1568 | in fn prop => prove stndrd [] prop (K tac) end; | 
| 17261 | 1569 | |
| 15215 | 1570 | fun inject_prf () = (prove_simp true [ext_def,abs_inject,Pair_eq] inject_prop); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1571 | val inject = timeit_msg "record extension inject proof:" inject_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1572 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1573 | fun induct_prf () = | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1574 | let val (assm, concl) = induct_prop | 
| 20248 | 1575 |       in prove_standard [assm] concl (fn {prems, ...} =>
 | 
| 17261 | 1576 | EVERY [try_param_tac rN abs_induct 1, | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1577 | simp_tac (HOL_ss addsimps [split_paired_all]) 1, | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1578 | resolve_tac (map (rewrite_rule [ext_def]) prems) 1]) | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1579 | end; | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1580 | val induct = timeit_msg "record extension induct proof:" induct_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1581 | |
| 15215 | 1582 | fun cases_prf_opt () = | 
| 17261 | 1583 | let | 
| 15215 | 1584 | val (_$(Pvar$_)) = concl_of induct; | 
| 17261 | 1585 | val ind = cterm_instantiate | 
| 18858 | 1586 | [(cterm_of defs_thy Pvar, cterm_of defs_thy | 
| 15215 | 1587 | (lambda w (HOLogic.imp$HOLogic.mk_eq(r,w)$C)))] | 
| 1588 | induct; | |
| 1589 | in standard (ObjectLogic.rulify (mp OF [ind, refl])) end; | |
| 1590 | ||
| 1591 | fun cases_prf_noopt () = | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1592 | prove_standard [] cases_prop (fn prems => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1593 | EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1594 | try_param_tac rN induct 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1595 | rtac impI 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1596 | REPEAT (etac allE 1), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1597 | etac mp 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1598 | rtac refl 1]) | 
| 15215 | 1599 | |
| 1600 | val cases_prf = quick_and_dirty_prf cases_prf_noopt cases_prf_opt; | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1601 | val cases = timeit_msg "record extension cases proof:" cases_prf; | 
| 17261 | 1602 | |
| 1603 | fun dest_convs_prf () = map (prove_simp false | |
| 15215 | 1604 | ([ext_def,abs_inverse]@Pair_sel_convs@dest_defs)) dest_conv_props; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1605 | val dest_convs = timeit_msg "record extension dest_convs proof:" dest_convs_prf; | 
| 15215 | 1606 | fun dest_convs_standard_prf () = map standard dest_convs; | 
| 1607 | ||
| 17261 | 1608 | val dest_convs_standard = | 
| 1609 | timeit_msg "record extension dest_convs_standard proof:" dest_convs_standard_prf; | |
| 1610 | ||
| 1611 | fun upd_convs_prf_noopt () = map (prove_simp true (dest_convs_standard@upd_defs)) | |
| 15215 | 1612 | upd_conv_props; | 
| 1613 | fun upd_convs_prf_opt () = | |
| 17261 | 1614 | let | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1615 | |
| 17261 | 1616 | fun mkrefl (c,T) = Thm.reflexive | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1617 | (cterm_of defs_thy (Free (Name.variant variants (base c ^ "'"),T-->T))); | 
| 15215 | 1618 | val refls = map mkrefl fields_more; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1619 | val dest_convs' = map mk_meta_eq dest_convs; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1620 | val map_eqs = map (uncurry Thm.combination) (refls ~~ dest_convs'); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1621 | |
| 18858 | 1622 | val constr_refl = Thm.reflexive (cterm_of defs_thy (head_of ext)); | 
| 17261 | 1623 | |
| 15215 | 1624 | fun mkthm (udef,(fld_refl,thms)) = | 
| 15570 | 1625 | let val bdyeq = Library.foldl (uncurry Thm.combination) (constr_refl,thms); | 
| 15215 | 1626 | (* (|N=N (|N=N,M=M,K=K,more=more|) | 
| 1627 | M=M (|N=N,M=M,K=K,more=more|) | |
| 1628 | K=K' | |
| 1629 | more = more (|N=N,M=M,K=K,more=more|) = | |
| 1630 | (|N=N,M=M,K=K',more=more|) | |
| 1631 | *) | |
| 1632 | val (_$(_$v$r)$_) = prop_of udef; | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1633 | val (_$(v'$_)$_) = prop_of fld_refl; | 
| 17261 | 1634 | val udef' = cterm_instantiate | 
| 18858 | 1635 | [(cterm_of defs_thy v,cterm_of defs_thy v'), | 
| 1636 | (cterm_of defs_thy r,cterm_of defs_thy ext)] udef; | |
| 17261 | 1637 | in standard (Thm.transitive udef' bdyeq) end; | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1638 | in map mkthm (rev upd_defs ~~ (mixit dest_convs' map_eqs)) end; | 
| 17261 | 1639 | |
| 15215 | 1640 | val upd_convs_prf = quick_and_dirty_prf upd_convs_prf_noopt upd_convs_prf_opt; | 
| 1641 | ||
| 17261 | 1642 | val upd_convs = | 
| 1643 | timeit_msg "record extension upd_convs proof:" upd_convs_prf; | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1644 | |
| 17261 | 1645 | fun surjective_prf () = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1646 | prove_standard [] surjective_prop (fn prems => | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1647 | (EVERY [try_param_tac rN induct 1, | 
| 15215 | 1648 | simp_tac (HOL_basic_ss addsimps dest_convs_standard) 1])); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1649 | val surjective = timeit_msg "record extension surjective proof:" surjective_prf; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1650 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1651 | fun split_meta_prf () = | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1652 | prove_standard [] split_meta_prop (fn prems => | 
| 21687 | 1653 | EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1, | 
| 17960 | 1654 | etac meta_allE 1, atac 1, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1655 | rtac (prop_subst OF [surjective]) 1, | 
| 17960 | 1656 | REPEAT (etac meta_allE 1), atac 1]); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1657 | val split_meta = timeit_msg "record extension split_meta proof:" split_meta_prf; | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1658 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1659 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1660 | val (([inject',induct',cases',surjective',split_meta'], | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1661 | [dest_convs',upd_convs']), | 
| 18377 | 1662 | thm_thy) = | 
| 17261 | 1663 | defs_thy | 
| 1664 | |> (PureThy.add_thms o map Thm.no_attributes) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1665 |            [("ext_inject", inject),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1666 |             ("ext_induct", induct),
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1667 |             ("ext_cases", cases),
 | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1668 |             ("ext_surjective", surjective),
 | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1669 |             ("ext_split", split_meta)]
 | 
| 18377 | 1670 | ||>> (PureThy.add_thmss o map Thm.no_attributes) | 
| 17261 | 1671 |               [("dest_convs",dest_convs_standard),("upd_convs",upd_convs)]
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1672 | |
| 15215 | 1673 | in (thm_thy,extT,induct',inject',dest_convs',split_meta',upd_convs') | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1674 | end; | 
| 17261 | 1675 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1676 | fun chunks [] [] = [] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1677 | | chunks [] xs = [xs] | 
| 15570 | 1678 | | chunks (l::ls) xs = Library.take (l,xs)::chunks ls (Library.drop (l,xs)); | 
| 17261 | 1679 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1680 | fun chop_last [] = error "last: list should not be empty" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1681 | | chop_last [x] = ([],x) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1682 | | chop_last (x::xs) = let val (tl,l) = chop_last xs in (x::tl,l) end; | 
| 17261 | 1683 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1684 | fun subst_last s [] = error "subst_last: list should not be empty" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1685 | | subst_last s ([x]) = [s] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1686 | | subst_last s (x::xs) = (x::subst_last s xs); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1687 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1688 | (* mk_recordT builds up the record type from the current extension tpye extT and a list | 
| 17261 | 1689 | * of parent extensions, starting with the root of the record hierarchy | 
| 1690 | *) | |
| 21078 | 1691 | fun mk_recordT extT = | 
| 1692 | fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1693 | |
| 15215 | 1694 | |
| 1695 | ||
| 1696 | fun obj_to_meta_all thm = | |
| 1697 | let | |
| 17261 | 1698 | fun E thm = case (SOME (spec OF [thm]) handle THM _ => NONE) of | 
| 15531 | 1699 | SOME thm' => E thm' | 
| 1700 | | NONE => thm; | |
| 15215 | 1701 | val th1 = E thm; | 
| 1702 | val th2 = Drule.forall_intr_vars th1; | |
| 1703 | in th2 end; | |
| 1704 | ||
| 1705 | fun meta_to_obj_all thm = | |
| 1706 | let | |
| 1707 |     val {sign, prop, ...} = rep_thm thm;
 | |
| 1708 | val params = Logic.strip_params prop; | |
| 1709 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl prop); | |
| 1710 | val ct = cterm_of sign | |
| 1711 | (HOLogic.mk_Trueprop (HOLogic.list_all (params, concl))); | |
| 1712 | val thm' = Seq.hd (REPEAT (rtac allI 1) (Thm.trivial ct)); | |
| 1713 | in | |
| 1714 | Thm.implies_elim thm' thm | |
| 1715 | end; | |
| 1716 | ||
| 1717 | ||
| 1718 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1719 | (* record_definition *) | 
| 17261 | 1720 | fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy = | 
| 14702 | 1721 | (* smlnj needs type annotation of parents *) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1722 | let | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1723 | val sign = Theory.sign_of thy; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1724 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1725 | val alphas = map fst args; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1726 | val name = Sign.full_name sign bname; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1727 | val full = Sign.full_name_path sign bname; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1728 | val base = Sign.base_name; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1729 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1730 | val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1731 | |
| 15570 | 1732 | val parent_fields = List.concat (map #fields parents); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1733 | val parent_chunks = map (length o #fields) parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1734 | val parent_names = map fst parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1735 | val parent_types = map snd parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1736 | val parent_fields_len = length parent_fields; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1737 | val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map base parent_names); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1738 | val parent_vars = ListPair.map Free (parent_variants, parent_types); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1739 | val parent_len = length parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1740 | val parents_idx = (map #name parents) ~~ (0 upto (parent_len - 1)); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1741 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1742 | val fields = map (apfst full) bfields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1743 | val names = map fst fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1744 | val extN = full bname; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1745 | val types = map snd fields; | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1746 | val alphas_fields = foldr add_typ_tfree_names [] types; | 
| 17261 | 1747 | val alphas_ext = alphas inter alphas_fields; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1748 | val len = length fields; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1749 | val variants = Name.variant_list (moreN::rN::rN ^ "'"::wN::parent_variants) (map fst bfields); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1750 | val vars = ListPair.map Free (variants, types); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1751 | val named_vars = names ~~ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1752 | val idxs = 0 upto (len - 1); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1753 | val idxms = 0 upto len; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1754 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1755 | val all_fields = parent_fields @ fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1756 | val all_names = parent_names @ names; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1757 | val all_types = parent_types @ types; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1758 | val all_len = parent_fields_len + len; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1759 | val all_variants = parent_variants @ variants; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1760 | val all_vars = parent_vars @ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1761 | val all_named_vars = (parent_names ~~ parent_vars) @ named_vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1762 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1763 | |
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1764 | val zeta = Name.variant alphas "'z"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1765 | val moreT = TFree (zeta, HOLogic.typeS); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1766 | val more = Free (moreN, moreT); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1767 | val full_moreN = full moreN; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1768 | val bfields_more = bfields @ [(moreN,moreT)]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1769 | val fields_more = fields @ [(full_moreN,moreT)]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1770 | val vars_more = vars @ [more]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1771 | val named_vars_more = named_vars @[(full_moreN,more)]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1772 | val all_vars_more = all_vars @ [more]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1773 | val all_named_vars_more = all_named_vars @ [(full_moreN,more)]; | 
| 17261 | 1774 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1775 | (* 1st stage: extension_thy *) | 
| 15215 | 1776 | val (extension_thy,extT,ext_induct,ext_inject,ext_dest_convs,ext_split,u_convs) = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1777 | thy | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1778 | |> Theory.add_path bname | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1779 | |> extension_definition full extN fields names alphas_ext zeta moreT more vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1780 | |
| 17261 | 1781 | val _ = timing_msg "record preparing definitions"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1782 | val Type extension_scheme = extT; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1783 | val extension_name = unsuffix ext_typeN (fst extension_scheme); | 
| 17261 | 1784 | val extension = let val (n,Ts) = extension_scheme in (n,subst_last HOLogic.unitT Ts) end; | 
| 1785 | val extension_names = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1786 | (map ((unsuffix ext_typeN) o fst o #extension) parents) @ [extN]; | 
| 15570 | 1787 |     val extension_id = Library.foldl (op ^) ("",extension_names);
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1788 | |
| 17261 | 1789 | |
| 21078 | 1790 | fun rec_schemeT n = mk_recordT (map #extension (prune n parents)) extT; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1791 | val rec_schemeT0 = rec_schemeT 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1792 | |
| 17261 | 1793 | fun recT n = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1794 | let val (c,Ts) = extension | 
| 21078 | 1795 | in mk_recordT (map #extension (prune n parents)) (Type (c,subst_last HOLogic.unitT Ts)) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1796 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1797 | val recT0 = recT 0; | 
| 17261 | 1798 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1799 | fun mk_rec args n = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1800 | let val (args',more) = chop_last args; | 
| 17261 | 1801 | fun mk_ext' (((name,T),args),more) = mk_ext (name,T) (args@[more]); | 
| 1802 | fun build Ts = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1803 | foldr mk_ext' more (prune n (extension_names ~~ Ts ~~ (chunks parent_chunks args'))) | 
| 17261 | 1804 | in | 
| 1805 | if more = HOLogic.unit | |
| 1806 | then build (map recT (0 upto parent_len)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1807 | else build (map rec_schemeT (0 upto parent_len)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1808 | end; | 
| 17261 | 1809 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1810 | val r_rec0 = mk_rec all_vars_more 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1811 | val r_rec_unit0 = mk_rec (all_vars@[HOLogic.unit]) 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1812 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1813 | fun r n = Free (rN, rec_schemeT n) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1814 | val r0 = r 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1815 | fun r_unit n = Free (rN, recT n) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1816 | val r_unit0 = r_unit 0; | 
| 15215 | 1817 | val w = Free (wN, rec_schemeT 0) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1818 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1819 | (* prepare print translation functions *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1820 | val field_tr's = | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18964diff
changeset | 1821 | print_translation (distinct (op =) | 
| 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18964diff
changeset | 1822 | (List.concat (map NameSpace.accesses' (full_moreN :: names)))); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1823 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1824 | val adv_ext_tr's = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1825 | let | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1826 | val trnames = NameSpace.accesses' extN; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1827 | in map (gen_record_tr') trnames end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1828 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1829 | val adv_record_type_abbr_tr's = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1830 | let val trnames = NameSpace.accesses' (hd extension_names); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1831 | val lastExt = (unsuffix ext_typeN (fst extension)); | 
| 17190 | 1832 | in map (gen_record_type_abbr_tr' name alphas zeta lastExt rec_schemeT0) trnames | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1833 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1834 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1835 | val adv_record_type_tr's = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1836 | let val trnames = if parent_len > 0 then NameSpace.accesses' extN else []; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1837 | (* avoid conflict with adv_record_type_abbr_tr's *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1838 | in map (gen_record_type_tr') trnames | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1839 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1840 | |
| 17261 | 1841 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1842 | (* prepare declarations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1843 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1844 | val sel_decls = map (mk_selC rec_schemeT0) bfields_more; | 
| 15215 | 1845 | val upd_decls = map (mk_updC updateN rec_schemeT0) bfields_more; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1846 | val make_decl = (makeN, all_types ---> recT0); | 
| 17261 | 1847 | val fields_decl = (fields_selN, types ---> Type extension); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1848 | val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1849 | val truncate_decl = (truncateN, rec_schemeT0 --> recT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1850 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1851 | (* prepare definitions *) | 
| 17261 | 1852 | |
| 1853 | fun parent_more s = | |
| 1854 | if null parents then s | |
| 16124 | 1855 | else mk_sel s (NameSpace.qualified (#name (List.last parents)) moreN, extT); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1856 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1857 | fun parent_more_upd v s = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1858 | if null parents then v$s | 
| 16124 | 1859 | else let val mp = NameSpace.qualified (#name (List.last parents)) moreN; | 
| 15215 | 1860 | in mk_upd updateN mp v s end; | 
| 17261 | 1861 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1862 | (*record (scheme) type abbreviation*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1863 | val recordT_specs = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1864 | [(suffix schemeN bname, alphas @ [zeta], rec_schemeT0, Syntax.NoSyn), | 
| 17261 | 1865 | (bname, alphas, recT0, Syntax.NoSyn)]; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1866 | |
| 17261 | 1867 | (*selectors*) | 
| 1868 | fun mk_sel_spec (c,T) = | |
| 1869 | Const (mk_selC rec_schemeT0 (c,T)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1870 | :== (lambda r0 (Const (mk_selC extT (suffix ext_dest c,T))$parent_more r0)); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1871 | val sel_specs = map mk_sel_spec fields_more; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1872 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1873 | (*updates*) | 
| 15215 | 1874 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1875 | fun mk_upd_spec (c,T) = | 
| 17261 | 1876 | let | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1877 | val new = mk_upd' updN c (Free (base c,T-->T)) extT(*(parent_more r0)*); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1878 | in Const (mk_updC updateN rec_schemeT0 (c,T))$(Free (base c,T-->T))$r0 | 
| 15215 | 1879 | :== (parent_more_upd new r0) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1880 | end; | 
| 17261 | 1881 | val upd_specs = map mk_upd_spec fields_more; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1882 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1883 | (*derived operations*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1884 | val make_spec = Const (full makeN, all_types ---> recT0) $$ all_vars :== | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1885 | mk_rec (all_vars @ [HOLogic.unit]) 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1886 | val fields_spec = Const (full fields_selN, types ---> Type extension) $$ vars :== | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1887 | mk_rec (all_vars @ [HOLogic.unit]) parent_len; | 
| 17261 | 1888 | val extend_spec = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1889 | Const (full extendN, recT0-->moreT-->rec_schemeT0) $ r_unit0 $ more :== | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1890 | mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1891 | val truncate_spec = Const (full truncateN, rec_schemeT0 --> recT0) $ r0 :== | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1892 | mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1893 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1894 | (* 2st stage: defs_thy *) | 
| 17261 | 1895 | |
| 16379 | 1896 | fun mk_defs () = | 
| 1897 | extension_thy | |
| 17261 | 1898 | |> Theory.add_trfuns | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1899 | ([],[],field_tr's, []) | 
| 17261 | 1900 | |> Theory.add_advanced_trfuns | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1901 | ([],[],adv_ext_tr's @ adv_record_type_tr's @ adv_record_type_abbr_tr's,[]) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1902 | |> Theory.parent_path | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1903 | |> Theory.add_tyabbrs_i recordT_specs | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1904 | |> Theory.add_path bname | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1905 | |> Theory.add_consts_i | 
| 18330 | 1906 | (map2 (fn (x, T) => fn mx => (x, T, mx)) sel_decls (field_syntax @ [Syntax.NoSyn])) | 
| 17261 | 1907 | |> (Theory.add_consts_i o map Syntax.no_syn) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1908 | (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl]) | 
| 18377 | 1909 | |> ((PureThy.add_defs_i false o map Thm.no_attributes) sel_specs) | 
| 1910 | ||>> ((PureThy.add_defs_i false o map Thm.no_attributes) upd_specs) | |
| 1911 | ||>> ((PureThy.add_defs_i false o map Thm.no_attributes) | |
| 1912 | [make_spec, fields_spec, extend_spec, truncate_spec]) | |
| 1913 | |> swap | |
| 16379 | 1914 | val (defs_thy,((sel_defs,upd_defs),derived_defs)) = | 
| 1915 | timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:" | |
| 1916 | mk_defs; | |
| 17261 | 1917 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1918 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1919 | (* prepare propositions *) | 
| 17261 | 1920 | val _ = timing_msg "record preparing propositions"; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1921 | val P = Free (Name.variant all_variants "P", rec_schemeT0-->HOLogic.boolT); | 
| 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1922 | val C = Free (Name.variant all_variants "C", HOLogic.boolT); | 
| 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1923 | val P_unit = Free (Name.variant all_variants "P", recT0-->HOLogic.boolT); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1924 | |
| 17261 | 1925 | (*selectors*) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1926 | val sel_conv_props = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1927 | map (fn (c, x as Free (_,T)) => mk_sel r_rec0 (c,T) === x) named_vars_more; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1928 | |
| 17261 | 1929 | (*updates*) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1930 | fun mk_upd_prop (i,(c,T)) = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1931 | let val x' = Free (Name.variant all_variants (base c ^ "'"),T-->T); | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1932 | val n = parent_fields_len + i; | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1933 | val args' = nth_map n (K (x'$nth all_vars_more n)) all_vars_more | 
| 15215 | 1934 | in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1935 | val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1936 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1937 | (*induct*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1938 | val induct_scheme_prop = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1939 | All (map dest_Free all_vars_more) (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0); | 
| 17261 | 1940 | val induct_prop = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1941 | (All (map dest_Free all_vars) (Trueprop (P_unit $ r_rec_unit0)), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1942 | Trueprop (P_unit $ r_unit0)); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1943 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1944 | (*surjective*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1945 | val surjective_prop = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1946 | let val args = map (fn (c,Free (_,T)) => mk_sel r0 (c,T)) all_named_vars_more | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1947 | in r0 === mk_rec args 0 end; | 
| 17261 | 1948 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1949 | (*cases*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1950 | val cases_scheme_prop = | 
| 17261 | 1951 | (All (map dest_Free all_vars_more) | 
| 1952 | (Trueprop (HOLogic.mk_eq (r0,r_rec0)) ==> Trueprop C)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1953 | ==> Trueprop C; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1954 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1955 | val cases_prop = | 
| 17261 | 1956 | (All (map dest_Free all_vars) | 
| 1957 | (Trueprop (HOLogic.mk_eq (r_unit0,r_rec_unit0)) ==> Trueprop C)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1958 | ==> Trueprop C; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1959 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1960 | (*split*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1961 | val split_meta_prop = | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1962 | let val P = Free (Name.variant all_variants "P", rec_schemeT0-->Term.propT) in | 
| 17261 | 1963 | Logic.mk_equals | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1964 | (All [dest_Free r0] (P $ r0), All (map dest_Free all_vars_more) (P $ r_rec0)) | 
| 17261 | 1965 | end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1966 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1967 | val split_object_prop = | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1968 | let fun ALL vs t = foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) t vs | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1969 | in (ALL [dest_Free r0] (P $ r0)) === (ALL (map dest_Free all_vars_more) (P $ r_rec0)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1970 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1971 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1972 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1973 | val split_ex_prop = | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1974 | let fun EX vs t = foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) t vs | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1975 | in (EX [dest_Free r0] (P $ r0)) === (EX (map dest_Free all_vars_more) (P $ r_rec0)) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1976 | end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1977 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1978 | (*equality*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1979 | val equality_prop = | 
| 17261 | 1980 | let | 
| 1981 | val s' = Free (rN ^ "'", rec_schemeT0) | |
| 1982 | fun mk_sel_eq (c,Free (_,T)) = mk_sel r0 (c,T) === mk_sel s' (c,T) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1983 | val seleqs = map mk_sel_eq all_named_vars_more | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1984 | in All (map dest_Free [r0,s']) (Logic.list_implies (seleqs,r0 === s')) end; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1985 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1986 | (* 3rd stage: thms_thy *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1987 | |
| 17510 | 1988 | fun prove stndrd = quick_and_dirty_prove stndrd defs_thy; | 
| 1989 | val prove_standard = quick_and_dirty_prove true defs_thy; | |
| 17261 | 1990 | |
| 15215 | 1991 | fun prove_simp stndrd ss simps = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1992 | let val tac = simp_all_tac ss simps | 
| 15215 | 1993 | in fn prop => prove stndrd [] prop (K tac) end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1994 | |
| 17510 | 1995 | val ss = get_simpset defs_thy; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1996 | |
| 17261 | 1997 | fun sel_convs_prf () = map (prove_simp false ss | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1998 | (sel_defs@ext_dest_convs)) sel_conv_props; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1999 | val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf; | 
| 15215 | 2000 | fun sel_convs_standard_prf () = map standard sel_convs | 
| 17261 | 2001 | val sel_convs_standard = | 
| 2002 | timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2003 | |
| 17261 | 2004 | fun upd_convs_prf () = | 
| 2005 | map (prove_simp true ss (upd_defs@u_convs)) upd_conv_props; | |
| 2006 | ||
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2007 | val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf; | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2008 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2009 | val parent_induct = if null parents then [] else [#induct (hd (rev parents))]; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2010 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2011 | fun induct_scheme_prf () = prove_standard [] induct_scheme_prop (fn prems => | 
| 17261 | 2012 | (EVERY [if null parent_induct | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2013 | then all_tac else try_param_tac rN (hd parent_induct) 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2014 | try_param_tac rN ext_induct 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2015 | asm_simp_tac HOL_basic_ss 1])); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2016 | val induct_scheme = timeit_msg "record induct_scheme proof:" induct_scheme_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2017 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2018 | fun induct_prf () = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2019 | let val (assm, concl) = induct_prop; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2020 | in | 
| 20248 | 2021 |         prove_standard [assm] concl (fn {prems, ...} =>
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2022 | try_param_tac rN induct_scheme 1 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2023 | THEN try_param_tac "more" unit_induct 1 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2024 | THEN resolve_tac prems 1) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2025 | end; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2026 | val induct = timeit_msg "record induct proof:" induct_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2027 | |
| 17261 | 2028 | fun surjective_prf () = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2029 | prove_standard [] surjective_prop (fn prems => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2030 | (EVERY [try_param_tac rN induct_scheme 1, | 
| 15215 | 2031 | simp_tac (ss addsimps sel_convs_standard) 1])) | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2032 | val surjective = timeit_msg "record surjective proof:" surjective_prf; | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2033 | |
| 15215 | 2034 | fun cases_scheme_prf_opt () = | 
| 17261 | 2035 | let | 
| 15215 | 2036 | val (_$(Pvar$_)) = concl_of induct_scheme; | 
| 17261 | 2037 | val ind = cterm_instantiate | 
| 18858 | 2038 | [(cterm_of defs_thy Pvar, cterm_of defs_thy | 
| 15215 | 2039 | (lambda w (HOLogic.imp$HOLogic.mk_eq(r0,w)$C)))] | 
| 2040 | induct_scheme; | |
| 2041 | in standard (ObjectLogic.rulify (mp OF [ind, refl])) end; | |
| 2042 | ||
| 2043 | fun cases_scheme_prf_noopt () = | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2044 | prove_standard [] cases_scheme_prop (fn prems => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2045 | EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2046 | try_param_tac rN induct_scheme 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2047 | rtac impI 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2048 | REPEAT (etac allE 1), | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2049 | etac mp 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2050 | rtac refl 1]) | 
| 15215 | 2051 | val cases_scheme_prf = quick_and_dirty_prf cases_scheme_prf_noopt cases_scheme_prf_opt; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2052 | val cases_scheme = timeit_msg "record cases_scheme proof:" cases_scheme_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2053 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2054 | fun cases_prf () = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2055 | prove_standard [] cases_prop (fn _ => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2056 | try_param_tac rN cases_scheme 1 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2057 | THEN simp_all_tac HOL_basic_ss [unit_all_eq1]); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2058 | val cases = timeit_msg "record cases proof:" cases_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2059 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2060 | fun split_meta_prf () = | 
| 15215 | 2061 | prove false [] split_meta_prop (fn prems => | 
| 21687 | 2062 | EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1, | 
| 17960 | 2063 | etac meta_allE 1, atac 1, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2064 | rtac (prop_subst OF [surjective]) 1, | 
| 17960 | 2065 | REPEAT (etac meta_allE 1), atac 1]); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2066 | val split_meta = timeit_msg "record split_meta proof:" split_meta_prf; | 
| 15215 | 2067 | val split_meta_standard = standard split_meta; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2068 | |
| 15215 | 2069 | fun split_object_prf_opt () = | 
| 17261 | 2070 | let | 
| 18858 | 2071 | val cPI= cterm_of defs_thy (lambda r0 (Trueprop (P$r0))); | 
| 15215 | 2072 | val (_$Abs(_,_,P$_)) = fst (Logic.dest_equals (concl_of split_meta_standard)); | 
| 18858 | 2073 | val cP = cterm_of defs_thy P; | 
| 15215 | 2074 | val split_meta' = cterm_instantiate [(cP,cPI)] split_meta_standard; | 
| 2075 | val (l,r) = HOLogic.dest_eq (HOLogic.dest_Trueprop split_object_prop); | |
| 18858 | 2076 | val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l); | 
| 2077 | val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r); | |
| 15215 | 2078 | val thl = assume cl (*All r. P r*) (* 1 *) | 
| 2079 | |> obj_to_meta_all (*!!r. P r*) | |
| 17261 | 2080 | |> equal_elim split_meta' (*!!n m more. P (ext n m more)*) | 
| 2081 | |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*) | |
| 15215 | 2082 | |> implies_intr cl (* 1 ==> 2 *) | 
| 2083 | val thr = assume cr (*All n m more. P (ext n m more)*) | |
| 2084 | |> obj_to_meta_all (*!!n m more. P (ext n m more)*) | |
| 17261 | 2085 | |> equal_elim (symmetric split_meta') (*!!r. P r*) | 
| 15215 | 2086 | |> meta_to_obj_all (*All r. P r*) | 
| 2087 | |> implies_intr cr (* 2 ==> 1 *) | |
| 17261 | 2088 | in standard (thr COMP (thl COMP iffI)) end; | 
| 15215 | 2089 | |
| 2090 | fun split_object_prf_noopt () = | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2091 | prove_standard [] split_object_prop (fn prems => | 
| 17261 | 2092 | EVERY [rtac iffI 1, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2093 | REPEAT (rtac allI 1), etac allE 1, atac 1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2094 | rtac allI 1, rtac induct_scheme 1,REPEAT (etac allE 1),atac 1]); | 
| 15215 | 2095 | |
| 17261 | 2096 | val split_object_prf = quick_and_dirty_prf split_object_prf_noopt split_object_prf_opt; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2097 | val split_object = timeit_msg "record split_object proof:" split_object_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2098 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2099 | |
| 17261 | 2100 | fun split_ex_prf () = | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2101 | prove_standard [] split_ex_prop (fn prems => | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2102 | EVERY [rtac iffI 1, | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2103 | etac exE 1, | 
| 15215 | 2104 | simp_tac (HOL_basic_ss addsimps [split_meta_standard]) 1, | 
| 2105 | ex_inst_tac 1, | |
| 2106 | (*REPEAT (rtac exI 1),*) | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2107 | atac 1, | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2108 | REPEAT (etac exE 1), | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2109 | rtac exI 1, | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2110 | atac 1]); | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2111 | val split_ex = timeit_msg "record split_ex proof:" split_ex_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2112 | |
| 17261 | 2113 | fun equality_tac thms = | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2114 | let val (s'::s::eqs) = rev thms; | 
| 15215 | 2115 | val ss' = ss addsimps (s'::s::sel_convs_standard); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2116 | val eqs' = map (simplify ss') eqs; | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2117 | in simp_tac (HOL_basic_ss addsimps (s'::s::eqs')) 1 end; | 
| 17261 | 2118 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2119 | fun equality_prf () = prove_standard [] equality_prop (fn _ => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2120 | fn st => let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2121 | st |> (res_inst_tac [(rN, s)] cases_scheme 1 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2122 | THEN res_inst_tac [(rN, s')] cases_scheme 1 | 
| 17261 | 2123 | THEN (METAHYPS equality_tac 1)) | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2124 | (* simp_all_tac ss (sel_convs) would also work but is less efficient *) | 
| 17261 | 2125 | end); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2126 | val equality = timeit_msg "record equality proof:" equality_prf; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2127 | |
| 18377 | 2128 | val ((([sel_convs',upd_convs',sel_defs',upd_defs',[split_meta',split_object',split_ex'],derived_defs'], | 
| 2129 | [surjective',equality']),[induct_scheme',induct',cases_scheme',cases']), thms_thy) = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2130 | defs_thy | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2131 | |> (PureThy.add_thmss o map Thm.no_attributes) | 
| 15215 | 2132 |          [("select_convs", sel_convs_standard),
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2133 |           ("update_convs", upd_convs),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2134 |           ("select_defs", sel_defs),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2135 |           ("update_defs", upd_defs),
 | 
| 15215 | 2136 |           ("splits", [split_meta_standard,split_object,split_ex]),
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2137 |           ("defs", derived_defs)]
 | 
| 18377 | 2138 | ||>> (PureThy.add_thms o map Thm.no_attributes) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2139 |           [("surjective", surjective),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2140 |            ("equality", equality)]
 | 
| 18377 | 2141 | ||>> PureThy.add_thms | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2142 |         [(("induct_scheme", induct_scheme), induct_type_global (suffix schemeN name)),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2143 |          (("induct", induct), induct_type_global name),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2144 |          (("cases_scheme", cases_scheme), cases_type_global (suffix schemeN name)),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2145 |          (("cases", cases), cases_type_global name)];
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2146 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2147 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2148 | val sel_upd_simps = sel_convs' @ upd_convs'; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2149 | val iffs = [ext_inject] | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2150 | val final_thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2151 | thms_thy | 
| 18377 | 2152 | |> (snd oo PureThy.add_thmss) | 
| 18728 | 2153 |           [(("simps", sel_upd_simps), [Simplifier.simp_add]),
 | 
| 2154 |            (("iffs",iffs), [iff_add])]
 | |
| 17261 | 2155 | |> put_record name (make_record_info args parent fields extension induct_scheme') | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2156 | |> put_sel_upd (names @ [full_moreN]) sel_upd_simps | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2157 | |> add_record_equalities extension_id equality' | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2158 | |> add_extinjects ext_inject | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2159 | |> add_extsplit extension_name ext_split | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2160 | |> add_record_splits extension_id (split_meta',split_object',split_ex',induct_scheme') | 
| 17261 | 2161 | |> add_extfields extension_name (fields @ [(full_moreN,moreT)]) | 
| 2162 | |> add_fieldext (extension_name,snd extension) (names @ [full_moreN]) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2163 | |> Theory.parent_path; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2164 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2165 | in final_thy | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2166 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2167 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2168 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2169 | |
| 4895 | 2170 | (*we do all preparations and error checks here, deferring the real | 
| 2171 | work to record_definition*) | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2172 | 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 | 2173 | let | 
| 17261 | 2174 | val _ = Theory.requires thy "Record" "record definitions"; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2175 | val sign = Theory.sign_of thy; | 
| 5698 | 2176 |     val _ = message ("Defining record " ^ quote bname ^ " ...");
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2177 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2178 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2179 | (* parents *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2180 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2181 | fun prep_inst T = snd (cert_typ sign ([], T)); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2182 | |
| 15570 | 2183 | val parent = Option.map (apfst (map prep_inst) o prep_raw_parent sign) raw_parent | 
| 18678 | 2184 |       handle ERROR msg => cat_error msg ("The error(s) above in parent record specification");
 | 
| 12247 | 2185 | val parents = add_parents thy parent []; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2186 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2187 | val init_env = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2188 | (case parent of | 
| 15531 | 2189 | NONE => [] | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 2190 | | SOME (types, _) => foldr Term.add_typ_tfrees [] types); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2191 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2192 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2193 | (* fields *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2194 | |
| 12506 | 2195 | fun prep_field (env, (c, raw_T, mx)) = | 
| 18678 | 2196 | let val (env', T) = prep_typ sign (env, raw_T) handle ERROR msg => | 
| 2197 |         cat_error msg ("The error(s) above occured in field " ^ quote c)
 | |
| 12506 | 2198 | in (env', (c, T, mx)) end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2199 | |
| 4967 | 2200 | 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 | 2201 | val envir_names = map fst envir; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2202 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2203 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2204 | (* args *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2205 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2206 | val defaultS = Sign.defaultS sign; | 
| 17485 | 2207 | val args = map (fn x => (x, AList.lookup (op =) envir x |> the_default defaultS)) params; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2208 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2209 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2210 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2211 | |
| 4890 | 2212 | val name = Sign.full_name sign bname; | 
| 17261 | 2213 | val err_dup_record = | 
| 4890 | 2214 | if is_none (get_record thy name) then [] | 
| 2215 | else ["Duplicate definition of record " ^ quote name]; | |
| 2216 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2217 | val err_dup_parms = | 
| 18964 | 2218 | (case duplicates (op =) params of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2219 | [] => [] | 
| 4890 | 2220 | | dups => ["Duplicate parameter(s) " ^ commas dups]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2221 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2222 | val err_extra_frees = | 
| 20951 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 haftmann parents: 
20484diff
changeset | 2223 | (case subtract (op =) params envir_names of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2224 | [] => [] | 
| 4890 | 2225 | | extras => ["Extra free type variable(s) " ^ commas extras]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2226 | |
| 4890 | 2227 | 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 | 2228 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2229 | val err_dup_fields = | 
| 18964 | 2230 | (case duplicates (op =) (map #1 bfields) of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2231 | [] => [] | 
| 4890 | 2232 | | dups => ["Duplicate field(s) " ^ commas_quote dups]); | 
| 2233 | ||
| 2234 | val err_bad_fields = | |
| 12506 | 2235 | if forall (not_equal moreN o #1) bfields then [] | 
| 4890 | 2236 | else ["Illegal field name " ^ quote moreN]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2237 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2238 | val err_dup_sorts = | 
| 18964 | 2239 | (case duplicates (op =) envir_names of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2240 | [] => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2241 | | dups => ["Inconsistent sort constraints for " ^ commas dups]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2242 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2243 | val errs = | 
| 4890 | 2244 | err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @ | 
| 2245 | err_dup_fields @ err_bad_fields @ err_dup_sorts; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2246 | in | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2247 | if null errs then () else error (cat_lines errs) ; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2248 | thy |> record_definition (args, bname) parent parents bfields | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2249 | end | 
| 18678 | 2250 |   handle ERROR msg => cat_error msg ("Failed to define record " ^ quote bname);
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2251 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2252 | 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 | 2253 | 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 | 2254 | |
| 6358 | 2255 | (* setup theory *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2256 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2257 | val setup = | 
| 18708 | 2258 | RecordsData.init #> | 
| 2259 | Theory.add_trfuns ([], parse_translation, [], []) #> | |
| 2260 | Theory.add_advanced_trfuns ([], adv_parse_translation, [], []) #> | |
| 17875 | 2261 | (fn thy => (Simplifier.change_simpset_of thy | 
| 18708 | 2262 | (fn ss => ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]); thy)); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2263 | |
| 6358 | 2264 | (* outer syntax *) | 
| 2265 | ||
| 17057 | 2266 | local structure P = OuterParse and K = OuterKeyword in | 
| 6358 | 2267 | |
| 2268 | val record_decl = | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12590diff
changeset | 2269 | P.type_args -- P.name -- | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12590diff
changeset | 2270 | (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const); | 
| 6358 | 2271 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 2272 | val recordP = | 
| 17261 | 2273 | OuterSyntax.command "record" "define extensible record" K.thy_decl | 
| 2274 | (record_decl >> (fn (x, (y, z)) => Toplevel.theory (add_record x y z))); | |
| 6358 | 2275 | |
| 2276 | val _ = OuterSyntax.add_parsers [recordP]; | |
| 2277 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2278 | end; | 
| 5698 | 2279 | |
| 6384 | 2280 | end; | 
| 2281 | ||
| 15215 | 2282 | |
| 5698 | 2283 | structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage; | 
| 17261 | 2284 | open BasicRecordPackage; |