| author | bulwahn | 
| Mon, 29 Mar 2010 17:30:53 +0200 | |
| changeset 36034 | ee84eac07290 | 
| parent 35994 | 9cc3df9a606e | 
| child 36137 | 0be811a98d3a | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 1 | (* Title: HOL/Tools/record.ML | 
| 32763 | 2 | Author: Wolfgang Naraschewski, TU Muenchen | 
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | Author: Norbert Schirmer, TU Muenchen | |
| 5 | Author: Thomas Sewell, NICTA | |
| 6 | ||
| 7 | Extensible records with structural subtyping. | |
| 5698 | 8 | *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 9 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 10 | signature BASIC_RECORD = | 
| 5698 | 11 | sig | 
| 35138 | 12 | type record_info = | 
| 13 |    {args: (string * sort) list,
 | |
| 14 | parent: (typ list * string) option, | |
| 15 | fields: (string * typ) list, | |
| 16 | extension: (string * typ list), | |
| 17 | ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm, | |
| 18 | select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list, | |
| 19 | fold_congs: thm list, unfold_congs: thm list, splits: thm list, defs: thm list, | |
| 20 | surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm, | |
| 21 | cases: thm, simps: thm list, iffs: thm list} | |
| 22 | val get_record: theory -> string -> record_info option | |
| 23 | val the_record: theory -> string -> record_info | |
| 7178 | 24 | val record_simproc: simproc | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 25 | val record_eq_simproc: simproc | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 26 | val record_upd_simproc: simproc | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 27 | val record_split_simproc: (term -> int) -> simproc | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 28 | val record_ex_sel_eq_simproc: simproc | 
| 5698 | 29 | val record_split_tac: int -> tactic | 
| 15273 
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
 schirmer parents: 
15258diff
changeset | 30 | val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic | 
| 5713 | 31 | val record_split_name: string | 
| 5698 | 32 | val record_split_wrapper: string * wrapper | 
| 32740 | 33 | val print_record_type_abbr: bool Unsynchronized.ref | 
| 34 | val print_record_type_as_fields: bool Unsynchronized.ref | |
| 5698 | 35 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 36 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 37 | signature RECORD = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 38 | sig | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 39 | include BASIC_RECORD | 
| 32740 | 40 | val timing: bool Unsynchronized.ref | 
| 8574 | 41 | val updateN: string | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 42 | val ext_typeN: string | 
| 21363 | 43 | val extN: string | 
| 44 | val makeN: string | |
| 45 | val moreN: string | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 46 | val last_extT: typ -> (string * typ list) option | 
| 32972 | 47 | val dest_recTs: typ -> (string * typ list) list | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 48 | val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 49 | val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ) | 
| 26088 | 50 | val get_parent: theory -> string -> (typ list * string) option | 
| 51 | val get_extension: theory -> string -> (string * typ list) option | |
| 16458 | 52 | val get_extinjects: theory -> thm list | 
| 53 | val get_simpset: theory -> simpset | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 54 | val print_records: theory -> unit | 
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 55 | val read_typ: Proof.context -> string -> (string * sort) list -> typ * (string * sort) list | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 56 | val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list | 
| 35136 | 57 | val add_record: bool -> string list * binding -> (typ list * string) option -> | 
| 58 | (binding * typ * mixfix) list -> theory -> theory | |
| 59 | val add_record_cmd: bool -> string list * binding -> string option -> | |
| 60 | (binding * string * mixfix) list -> theory -> theory | |
| 18708 | 61 | val setup: theory -> theory | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 62 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 63 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 64 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 65 | signature ISO_TUPLE_SUPPORT = | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 66 | sig | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 67 | val add_iso_tuple_type: bstring * string list -> typ * typ -> theory -> (term * term) * theory | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 68 | val mk_cons_tuple: term * term -> term | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 69 | val dest_cons_tuple: term -> term * term | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 70 | val iso_tuple_intros_tac: int -> tactic | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 71 | val named_cterm_instantiate: (string * cterm) list -> thm -> thm | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 72 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 73 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 74 | structure Iso_Tuple_Support: ISO_TUPLE_SUPPORT = | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 75 | struct | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 76 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 77 | val isoN = "_Tuple_Iso"; | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 78 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 79 | val iso_tuple_intro = @{thm isomorphic_tuple_intro};
 | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 80 | val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros};
 | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 81 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 82 | val tuple_iso_tuple = (@{const_name tuple_iso_tuple}, @{thm tuple_iso_tuple});
 | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 83 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 84 | fun named_cterm_instantiate values thm = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 85 | let | 
| 35135 | 86 | fun match name ((name', _), _) = name = name'; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 87 | fun getvar name = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 88 | (case find_first (match name) (Term.add_vars (prop_of thm) []) of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 89 | SOME var => cterm_of (theory_of_thm thm) (Var var) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 90 |       | NONE => raise THM ("named_cterm_instantiate: " ^ name, 0, [thm]));
 | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 91 | in | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 92 | cterm_instantiate (map (apfst getvar) values) thm | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 93 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 94 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 95 | structure Iso_Tuple_Thms = Theory_Data | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 96 | ( | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 97 | type T = thm Symtab.table; | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 98 | val empty = Symtab.make [tuple_iso_tuple]; | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 99 | val extend = I; | 
| 33522 | 100 | fun merge data = Symtab.merge Thm.eq_thm_prop data; (* FIXME handle Symtab.DUP ?? *) | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 101 | ); | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 102 | |
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 103 | fun do_typedef name repT alphas thy = | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 104 | let | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 105 | fun get_thms thy name = | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 106 | let | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35989diff
changeset | 107 |         val [({Abs_name, abs_type = absT, ...}, {Rep_inject, Abs_inverse, ...})] =
 | 
| 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35989diff
changeset | 108 | Typedef.get_info_global thy name; | 
| 32972 | 109 | val rewrite_rule = | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 110 |           MetaSimplifier.rewrite_rule [@{thm iso_tuple_UNIV_I}, @{thm iso_tuple_True_simp}];
 | 
| 32972 | 111 | in | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35989diff
changeset | 112 | (map rewrite_rule [Rep_inject, Abs_inverse], Const (Abs_name, repT --> absT), absT) | 
| 32972 | 113 | end; | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 114 | in | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 115 | thy | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 116 | |> Typecopy.typecopy (Binding.name name, alphas) repT NONE | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 117 | |-> (fn (name, _) => `(fn thy => get_thms thy name)) | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 118 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 119 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 120 | fun mk_cons_tuple (left, right) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 121 | let | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 122 | val (leftT, rightT) = (fastype_of left, fastype_of right); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 123 | val prodT = HOLogic.mk_prodT (leftT, rightT); | 
| 32972 | 124 |     val isomT = Type (@{type_name tuple_isomorphism}, [prodT, leftT, rightT]);
 | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 125 | in | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 126 |     Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> prodT) $
 | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 127 | Const (fst tuple_iso_tuple, isomT) $ left $ right | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 128 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 129 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 130 | fun dest_cons_tuple (Const (@{const_name iso_tuple_cons}, _) $ Const _ $ t $ u) = (t, u)
 | 
| 32972 | 131 |   | dest_cons_tuple t = raise TERM ("dest_cons_tuple", [t]);
 | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 132 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 133 | fun add_iso_tuple_type (name, alphas) (leftT, rightT) thy = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 134 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 135 | val repT = HOLogic.mk_prodT (leftT, rightT); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 136 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 137 | val (([rep_inject, abs_inverse], absC, absT), typ_thy) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 138 | thy | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 139 | |> do_typedef name repT alphas | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 140 | ||> Sign.add_path name; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 141 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 142 | (*construct a type and body for the isomorphism constant by | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 143 | instantiating the theorem to which the definition will be applied*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 144 | val intro_inst = | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 145 |       rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] iso_tuple_intro;
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 146 | val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst)); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 147 | val isomT = fastype_of body; | 
| 35239 | 148 | val isom_binding = Binding.name (name ^ isoN); | 
| 149 | val isom_name = Sign.full_name typ_thy isom_binding; | |
| 32972 | 150 | val isom = Const (isom_name, isomT); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 151 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 152 | val ([isom_def], cdef_thy) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 153 | typ_thy | 
| 35239 | 154 | |> Sign.declare_const ((isom_binding, isomT), NoSyn) |> snd | 
| 35142 | 155 | |> PureThy.add_defs false | 
| 35239 | 156 | [((Binding.conceal (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])]; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 157 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 158 | val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro)); | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 159 |     val cons = Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 160 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 161 | val thm_thy = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 162 | cdef_thy | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 163 | |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple)) | 
| 35239 | 164 | |> Sign.restore_naming thy | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
33522diff
changeset | 165 | |> Code.add_default_eqn isom_def; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 166 | in | 
| 32972 | 167 | ((isom, cons $ isom), thm_thy) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 168 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 169 | |
| 35133 | 170 | val iso_tuple_intros_tac = | 
| 171 | resolve_from_net_tac iso_tuple_intros THEN' | |
| 32975 | 172 | CSUBGOAL (fn (cgoal, i) => | 
| 173 | let | |
| 174 | val thy = Thm.theory_of_cterm cgoal; | |
| 175 | val goal = Thm.term_of cgoal; | |
| 176 | ||
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 177 | val isthms = Iso_Tuple_Thms.get thy; | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 178 |       fun err s t = raise TERM ("iso_tuple_intros_tac: " ^ s, [t]);
 | 
| 32975 | 179 | |
| 180 | val goal' = Envir.beta_eta_contract goal; | |
| 181 | val is = | |
| 182 | (case goal' of | |
| 183 |           Const (@{const_name Trueprop}, _) $
 | |
| 184 |             (Const (@{const_name isomorphic_tuple}, _) $ Const is) => is
 | |
| 185 | | _ => err "unexpected goal format" goal'); | |
| 186 | val isthm = | |
| 187 | (case Symtab.lookup isthms (#1 is) of | |
| 188 | SOME isthm => isthm | |
| 189 | | NONE => err "no thm found for constant" (Const is)); | |
| 190 | in rtac isthm i end); | |
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 191 | |
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 192 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 193 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 194 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 195 | structure Record: RECORD = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 196 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 197 | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 198 | val eq_reflection = @{thm eq_reflection};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 199 | val atomize_all = @{thm HOL.atomize_all};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 200 | val atomize_imp = @{thm HOL.atomize_imp};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 201 | val meta_allE = @{thm Pure.meta_allE};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 202 | val prop_subst = @{thm prop_subst};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 203 | val K_record_comp = @{thm K_record_comp};
 | 
| 35136 | 204 | val K_comp_convs = [@{thm o_apply}, K_record_comp];
 | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 205 | val o_assoc = @{thm o_assoc};
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 206 | val id_apply = @{thm id_apply};
 | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 207 | val id_o_apps = [@{thm id_apply}, @{thm id_o}, @{thm o_id}];
 | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 208 | val Not_eq_iff = @{thm Not_eq_iff};
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 209 | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 210 | val refl_conj_eq = @{thm refl_conj_eq};
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 211 | |
| 35133 | 212 | val surject_assistI = @{thm iso_tuple_surjective_proof_assistI};
 | 
| 213 | val surject_assist_idE = @{thm iso_tuple_surjective_proof_assist_idE};
 | |
| 214 | ||
| 215 | val updacc_accessor_eqE = @{thm update_accessor_accessor_eqE};
 | |
| 216 | val updacc_updator_eqE = @{thm update_accessor_updator_eqE};
 | |
| 217 | val updacc_eq_idI = @{thm iso_tuple_update_accessor_eq_assist_idI};
 | |
| 218 | val updacc_eq_triv = @{thm iso_tuple_update_accessor_eq_assist_triv};
 | |
| 219 | ||
| 220 | val updacc_foldE = @{thm update_accessor_congruence_foldE};
 | |
| 221 | val updacc_unfoldE = @{thm update_accessor_congruence_unfoldE};
 | |
| 222 | val updacc_noopE = @{thm update_accessor_noopE};
 | |
| 223 | val updacc_noop_compE = @{thm update_accessor_noop_compE};
 | |
| 224 | val updacc_cong_idI = @{thm update_accessor_cong_assist_idI};
 | |
| 225 | val updacc_cong_triv = @{thm update_accessor_cong_assist_triv};
 | |
| 226 | val updacc_cong_from_eq = @{thm iso_tuple_update_accessor_cong_from_eq};
 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 227 | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 228 | val o_eq_dest = @{thm o_eq_dest};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 229 | val o_eq_id_dest = @{thm o_eq_id_dest};
 | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 230 | val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
 | 
| 11832 | 231 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 232 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 233 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 234 | (** name components **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 235 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 236 | val rN = "r"; | 
| 15215 | 237 | val wN = "w"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 238 | val moreN = "more"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 239 | val schemeN = "_scheme"; | 
| 17261 | 240 | val ext_typeN = "_ext_type"; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 241 | val inner_typeN = "_inner_type"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 242 | val extN ="_ext"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 243 | val updateN = "_update"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 244 | val makeN = "make"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 245 | val fields_selN = "fields"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 246 | val extendN = "extend"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 247 | val truncateN = "truncate"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 248 | |
| 32335 | 249 | |
| 250 | ||
| 4894 | 251 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 252 | |
| 14709 
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
 schirmer parents: 
14702diff
changeset | 253 | fun but_last xs = fst (split_last xs); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 254 | |
| 19748 | 255 | fun varifyT midx = | 
| 256 | let fun varify (a, S) = TVar ((a, midx + 1), S); | |
| 257 | in map_type_tfree varify end; | |
| 258 | ||
| 32335 | 259 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 260 | (* timing *) | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 261 | |
| 32740 | 262 | val timing = Unsynchronized.ref false; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 263 | fun timeit_msg s x = if ! timing then (warning s; timeit x) else x (); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 264 | fun timing_msg s = if ! timing then warning s else (); | 
| 17261 | 265 | |
| 32335 | 266 | |
| 12255 | 267 | (* syntax *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 268 | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 269 | val Trueprop = HOLogic.mk_Trueprop; | 
| 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 270 | fun All xs t = Term.list_all_free (xs, t); | 
| 4894 | 271 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 272 | infix 0 :== ===; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 273 | infixr 0 ==>; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 274 | |
| 35989 
3418cdf1855e
moved Primitive_Defs.mk_defpair to OldGoals.mk_defpair;
 wenzelm parents: 
35742diff
changeset | 275 | val op :== = OldGoals.mk_defpair; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 276 | val op === = Trueprop o HOLogic.mk_eq; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 277 | val op ==> = Logic.mk_implies; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 278 | |
| 32335 | 279 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 280 | (* constructor *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 281 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 282 | fun mk_ext (name, T) ts = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 283 | let val Ts = map fastype_of ts | 
| 35239 | 284 | in list_comb (Const (suffix extN name, Ts ---> T), ts) end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 285 | |
| 32335 | 286 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 287 | (* selector *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 288 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 289 | fun mk_selC sT (c, T) = (c, sT --> T); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 290 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 291 | fun mk_sel s (c, T) = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 292 | let val sT = fastype_of s | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 293 | in Const (mk_selC sT (c, T)) $ s end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 294 | |
| 32335 | 295 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 296 | (* updates *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 297 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 298 | 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 | 299 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 300 | 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 | 301 | let val vT = domain_type (fastype_of v); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 302 | in Const (mk_updC sfx sT (c, vT)) $ v end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 303 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 304 | 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 | 305 | |
| 32335 | 306 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 307 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 308 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 309 | fun dest_recT (typ as Type (c_ext_type, Ts as (_ :: _))) = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 310 | (case try (unsuffix ext_typeN) c_ext_type of | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 311 |         NONE => raise TYPE ("Record.dest_recT", [typ], [])
 | 
| 15570 | 312 | | SOME c => ((c, Ts), List.last Ts)) | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 313 |   | dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []);
 | 
| 5197 | 314 | |
| 32975 | 315 | val is_recT = can dest_recT; | 
| 11833 | 316 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 317 | fun dest_recTs T = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 318 | let val ((c, Ts), U) = dest_recT T | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 319 | in (c, Ts) :: dest_recTs U | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 320 | end handle TYPE _ => []; | 
| 14255 | 321 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 322 | fun last_extT T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 323 | let val ((c, Ts), U) = dest_recT T in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 324 | (case last_extT U of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 325 | NONE => SOME (c, Ts) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 326 | | SOME l => SOME l) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 327 | end handle TYPE _ => NONE; | 
| 14255 | 328 | |
| 17261 | 329 | fun rec_id i T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 330 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 331 | val rTs = dest_recTs T; | 
| 33957 | 332 | val rTs' = if i < 0 then rTs else take i rTs; | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 333 | in implode (map #1 rTs') end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 334 | |
| 32335 | 335 | |
| 336 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 337 | (*** extend theory by record definition ***) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 338 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 339 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 340 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 341 | (* type record_info and parent_info *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 342 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 343 | type record_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 344 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 345 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 346 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 347 | extension: (string * typ list), | 
| 35138 | 348 | |
| 349 | ext_induct: thm, | |
| 350 | ext_inject: thm, | |
| 351 | ext_surjective: thm, | |
| 352 | ext_split: thm, | |
| 353 | ext_def: thm, | |
| 354 | ||
| 355 | select_convs: thm list, | |
| 356 | update_convs: thm list, | |
| 357 | select_defs: thm list, | |
| 358 | update_defs: thm list, | |
| 359 | fold_congs: thm list, | |
| 360 | unfold_congs: thm list, | |
| 361 | splits: thm list, | |
| 362 | defs: thm list, | |
| 363 | ||
| 364 | surjective: thm, | |
| 365 | equality: thm, | |
| 366 | induct_scheme: thm, | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 367 | induct: thm, | 
| 35138 | 368 | cases_scheme: thm, | 
| 369 | cases: thm, | |
| 370 | ||
| 371 | simps: thm list, | |
| 372 | iffs: thm list}; | |
| 373 | ||
| 374 | fun make_record_info args parent fields extension | |
| 375 | ext_induct ext_inject ext_surjective ext_split ext_def | |
| 376 | select_convs update_convs select_defs update_defs fold_congs unfold_congs splits defs | |
| 377 | surjective equality induct_scheme induct cases_scheme cases | |
| 378 | simps iffs : record_info = | |
| 17261 | 379 |  {args = args, parent = parent, fields = fields, extension = extension,
 | 
| 35138 | 380 | ext_induct = ext_induct, ext_inject = ext_inject, ext_surjective = ext_surjective, | 
| 381 | ext_split = ext_split, ext_def = ext_def, select_convs = select_convs, | |
| 382 | update_convs = update_convs, select_defs = select_defs, update_defs = update_defs, | |
| 383 | fold_congs = fold_congs, unfold_congs = unfold_congs, splits = splits, defs = defs, | |
| 384 | surjective = surjective, equality = equality, induct_scheme = induct_scheme, | |
| 385 | induct = induct, cases_scheme = cases_scheme, cases = cases, simps = simps, iffs = iffs}; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 386 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 387 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 388 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 389 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 390 | extension: (string * typ list), | 
| 35138 | 391 | induct_scheme: thm, | 
| 392 | ext_def: thm}; | |
| 393 | ||
| 394 | fun make_parent_info name fields extension ext_def induct_scheme : parent_info = | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 395 |  {name = name, fields = fields, extension = extension,
 | 
| 35138 | 396 | ext_def = ext_def, induct_scheme = induct_scheme}; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 397 | |
| 22846 | 398 | |
| 399 | (* theory data *) | |
| 5001 | 400 | |
| 7178 | 401 | type record_data = | 
| 402 |  {records: record_info Symtab.table,
 | |
| 403 | sel_upd: | |
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 404 |    {selectors: (int * bool) Symtab.table,
 | 
| 7178 | 405 | updates: string Symtab.table, | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 406 | simpset: Simplifier.simpset, | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 407 | defset: Simplifier.simpset, | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 408 | foldcong: Simplifier.simpset, | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 409 | unfoldcong: Simplifier.simpset}, | 
| 14255 | 410 | equalities: thm Symtab.table, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 411 | extinjects: thm list, | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 412 | extsplit: thm Symtab.table, (*maps extension name to split rule*) | 
| 35135 | 413 | splits: (thm * thm * thm * thm) Symtab.table, (*!!, ALL, EX - split-equalities, induct rule*) | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 414 | extfields: (string * typ) list Symtab.table, (*maps extension to its fields*) | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 415 | fieldext: (string * typ list) Symtab.table}; (*maps field to its extension*) | 
| 7178 | 416 | |
| 17261 | 417 | fun make_record_data | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 418 | records sel_upd equalities extinjects extsplit splits extfields fieldext = | 
| 17261 | 419 |  {records = records, sel_upd = sel_upd,
 | 
| 420 | equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits, | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 421 | extfields = extfields, fieldext = fieldext }: record_data; | 
| 7178 | 422 | |
| 35137 | 423 | structure Records_Data = Theory_Data | 
| 22846 | 424 | ( | 
| 7178 | 425 | type T = record_data; | 
| 426 | val empty = | |
| 427 | make_record_data Symtab.empty | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 428 |       {selectors = Symtab.empty, updates = Symtab.empty,
 | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 429 | simpset = HOL_basic_ss, defset = HOL_basic_ss, | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 430 | foldcong = HOL_basic_ss, unfoldcong = HOL_basic_ss} | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 431 | Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; | 
| 16458 | 432 | val extend = I; | 
| 33522 | 433 | fun merge | 
| 7178 | 434 |    ({records = recs1,
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 435 | sel_upd = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 436 |       {selectors = sels1, updates = upds1,
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 437 | simpset = ss1, defset = ds1, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 438 | foldcong = fc1, unfoldcong = uc1}, | 
| 14255 | 439 | equalities = equalities1, | 
| 17261 | 440 | extinjects = extinjects1, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 441 | extsplit = extsplit1, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 442 | splits = splits1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 443 | extfields = extfields1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 444 | fieldext = fieldext1}, | 
| 7178 | 445 |     {records = recs2,
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 446 | sel_upd = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 447 |       {selectors = sels2, updates = upds2,
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 448 | simpset = ss2, defset = ds2, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 449 | foldcong = fc2, unfoldcong = uc2}, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 450 | equalities = equalities2, | 
| 17261 | 451 | extinjects = extinjects2, | 
| 452 | extsplit = extsplit2, | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 453 | splits = splits2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 454 | extfields = extfields2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 455 | fieldext = fieldext2}) = | 
| 17261 | 456 | make_record_data | 
| 7178 | 457 | (Symtab.merge (K true) (recs1, recs2)) | 
| 458 |       {selectors = Symtab.merge (K true) (sels1, sels2),
 | |
| 459 | updates = Symtab.merge (K true) (upds1, upds2), | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 460 | simpset = Simplifier.merge_ss (ss1, ss2), | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 461 | defset = Simplifier.merge_ss (ds1, ds2), | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 462 | foldcong = Simplifier.merge_ss (fc1, fc2), | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 463 | unfoldcong = Simplifier.merge_ss (uc1, uc2)} | 
| 22634 | 464 | (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2)) | 
| 33522 | 465 | (Thm.merge_thms (extinjects1, extinjects2)) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 466 | (Symtab.merge Thm.eq_thm_prop (extsplit1, extsplit2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 467 | (Symtab.merge (fn ((a, b, c, d), (w, x, y, z)) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 468 | Thm.eq_thm (a, w) andalso Thm.eq_thm (b, x) andalso | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 469 | Thm.eq_thm (c, y) andalso Thm.eq_thm (d, z)) (splits1, splits2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 470 | (Symtab.merge (K true) (extfields1, extfields2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 471 | (Symtab.merge (K true) (fieldext1, fieldext2)); | 
| 22846 | 472 | ); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 473 | |
| 22846 | 474 | fun print_records thy = | 
| 475 | let | |
| 35137 | 476 |     val {records = recs, ...} = Records_Data.get thy;
 | 
| 26943 
aec0d97a01c4
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26626diff
changeset | 477 | val prt_typ = Syntax.pretty_typ_global thy; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 478 | |
| 22846 | 479 | fun pretty_parent NONE = [] | 
| 480 | | pretty_parent (SOME (Ts, name)) = | |
| 481 | [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 482 | |
| 22846 | 483 | fun pretty_field (c, T) = Pretty.block | 
| 484 | [Pretty.str (Sign.extern_const thy c), Pretty.str " ::", | |
| 485 | Pretty.brk 1, Pretty.quote (prt_typ T)]; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 486 | |
| 22846 | 487 |     fun pretty_record (name, {args, parent, fields, ...}: record_info) =
 | 
| 488 | Pretty.block (Pretty.fbreaks (Pretty.block | |
| 489 | [prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: | |
| 490 | pretty_parent parent @ map pretty_field fields)); | |
| 491 | in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end; | |
| 5006 | 492 | |
| 16458 | 493 | |
| 7178 | 494 | (* access 'records' *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 495 | |
| 35137 | 496 | val get_record = Symtab.lookup o #records o Records_Data.get; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 497 | |
| 35138 | 498 | fun the_record thy name = | 
| 499 | (case get_record thy name of | |
| 500 | SOME info => info | |
| 501 |   | NONE => error ("Unknown record type " ^ quote name));
 | |
| 502 | ||
| 4890 | 503 | fun put_record name info thy = | 
| 7178 | 504 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 505 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 506 | Records_Data.get thy; | 
| 17412 | 507 | val data = make_record_data (Symtab.update (name, info) records) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 508 | sel_upd equalities extinjects extsplit splits extfields fieldext; | 
| 35137 | 509 | in Records_Data.put data thy end; | 
| 7178 | 510 | |
| 22846 | 511 | |
| 7178 | 512 | (* access 'sel_upd' *) | 
| 513 | ||
| 35137 | 514 | val get_sel_upd = #sel_upd o Records_Data.get; | 
| 7178 | 515 | |
| 17510 | 516 | val is_selector = Symtab.defined o #selectors o get_sel_upd; | 
| 17412 | 517 | val get_updates = Symtab.lookup o #updates o get_sel_upd; | 
| 35232 
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
 wenzelm parents: 
35149diff
changeset | 518 | fun get_ss_with_context getss thy = Simplifier.global_context thy (getss (get_sel_upd thy)); | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 519 | |
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 520 | val get_simpset = get_ss_with_context #simpset; | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 521 | val get_sel_upd_defs = get_ss_with_context #defset; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 522 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 523 | fun get_update_details u thy = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 524 | let val sel_upd = get_sel_upd thy in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 525 | (case Symtab.lookup (#updates sel_upd) u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 526 | SOME s => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 527 | let val SOME (dep, ismore) = Symtab.lookup (#selectors sel_upd) s | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 528 | in SOME (s, dep, ismore) end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 529 | | NONE => NONE) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 530 | end; | 
| 7178 | 531 | |
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 532 | fun put_sel_upd names more depth simps defs (folds, unfolds) thy = | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 533 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 534 | val all = names @ [more]; | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 535 | val sels = map (rpair (depth, false)) names @ [(more, (depth, true))]; | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 536 | val upds = map (suffix updateN) all ~~ all; | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 537 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 538 |     val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong},
 | 
| 35137 | 539 | equalities, extinjects, extsplit, splits, extfields, fieldext} = Records_Data.get thy; | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 540 | val data = make_record_data records | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 541 |       {selectors = fold Symtab.update_new sels selectors,
 | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 542 | updates = fold Symtab.update_new upds updates, | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 543 | simpset = Simplifier.addsimps (simpset, simps), | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 544 | defset = Simplifier.addsimps (defset, defs), | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 545 | foldcong = foldcong addcongs folds, | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 546 | unfoldcong = unfoldcong addcongs unfolds} | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 547 | equalities extinjects extsplit splits extfields fieldext; | 
| 35137 | 548 | in Records_Data.put data thy end; | 
| 22846 | 549 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 550 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 551 | (* access 'equalities' *) | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 552 | |
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 553 | fun add_record_equalities name thm thy = | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 554 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 555 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 556 | Records_Data.get thy; | 
| 17261 | 557 | val data = make_record_data records sel_upd | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 558 | (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext; | 
| 35137 | 559 | in Records_Data.put data thy end; | 
| 560 | ||
| 561 | val get_equalities = Symtab.lookup o #equalities o Records_Data.get; | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 562 | |
| 22846 | 563 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 564 | (* access 'extinjects' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 565 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 566 | fun add_extinjects thm thy = | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 567 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 568 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 569 | Records_Data.get thy; | 
| 22846 | 570 | val data = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 571 | make_record_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 572 | extsplit splits extfields fieldext; | 
| 35137 | 573 | in Records_Data.put data thy end; | 
| 574 | ||
| 575 | val get_extinjects = rev o #extinjects o Records_Data.get; | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 576 | |
| 22846 | 577 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 578 | (* access 'extsplit' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 579 | |
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 580 | fun add_extsplit name thm thy = | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 581 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 582 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 583 | Records_Data.get thy; | 
| 35144 | 584 | val data = | 
| 585 | make_record_data records sel_upd equalities extinjects | |
| 586 | (Symtab.update_new (name, thm) extsplit) splits extfields fieldext; | |
| 35137 | 587 | in Records_Data.put data thy end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 588 | |
| 26088 | 589 | |
| 14255 | 590 | (* access 'splits' *) | 
| 591 | ||
| 592 | fun add_record_splits name thmP thy = | |
| 593 | let | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 594 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 595 | Records_Data.get thy; | 
| 35144 | 596 | val data = | 
| 597 | make_record_data records sel_upd equalities extinjects extsplit | |
| 598 | (Symtab.update_new (name, thmP) splits) extfields fieldext; | |
| 35137 | 599 | in Records_Data.put data thy end; | 
| 600 | ||
| 601 | val get_splits = Symtab.lookup o #splits o Records_Data.get; | |
| 14255 | 602 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 603 | |
| 26088 | 604 | (* parent/extension of named record *) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 605 | |
| 35137 | 606 | val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Records_Data.get); | 
| 607 | val get_extension = Option.map #extension oo (Symtab.lookup o #records o Records_Data.get); | |
| 17261 | 608 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 609 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 610 | (* access 'extfields' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 611 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 612 | fun add_extfields name fields thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 613 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 614 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 615 | Records_Data.get thy; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 616 | val data = | 
| 35144 | 617 | make_record_data records sel_upd equalities extinjects extsplit splits | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 618 | (Symtab.update_new (name, fields) extfields) fieldext; | 
| 35137 | 619 | in Records_Data.put data thy end; | 
| 620 | ||
| 621 | val get_extfields = Symtab.lookup o #extfields o Records_Data.get; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 622 | |
| 18858 | 623 | fun get_extT_fields thy T = | 
| 15059 | 624 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 625 | val ((name, Ts), moreT) = dest_recT T; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 626 | val recname = | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 627 | let val (nm :: _ :: rst) = rev (Long_Name.explode name) (* FIXME !? *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 628 | in Long_Name.implode (rev (nm :: rst)) end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 629 | val midx = maxidx_of_typs (moreT :: Ts); | 
| 19748 | 630 | val varifyT = varifyT midx; | 
| 35137 | 631 |     val {records, extfields, ...} = Records_Data.get thy;
 | 
| 35149 | 632 | val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name); | 
| 17412 | 633 | val args = map varifyT (snd (#extension (the (Symtab.lookup records recname)))); | 
| 15058 | 634 | |
| 35149 | 635 | val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) Vartab.empty; | 
| 636 | val fields' = map (apsnd (Envir.norm_type subst o varifyT)) fields; | |
| 637 | in (fields', (more, moreT)) end; | |
| 15058 | 638 | |
| 18858 | 639 | fun get_recT_fields thy T = | 
| 17261 | 640 | let | 
| 35149 | 641 | val (root_fields, (root_more, root_moreT)) = get_extT_fields thy T; | 
| 642 | val (rest_fields, rest_more) = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 643 | if is_recT root_moreT then get_recT_fields thy root_moreT | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 644 | else ([], (root_more, root_moreT)); | 
| 35149 | 645 | in (root_fields @ rest_fields, rest_more) end; | 
| 15059 | 646 | |
| 15058 | 647 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 648 | (* access 'fieldext' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 649 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 650 | fun add_fieldext extname_types fields thy = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 651 | let | 
| 17261 | 652 |     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
 | 
| 35137 | 653 | Records_Data.get thy; | 
| 17261 | 654 | val fieldext' = | 
| 17412 | 655 | fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 656 | val data = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 657 | make_record_data records sel_upd equalities extinjects | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 658 | extsplit splits extfields fieldext'; | 
| 35137 | 659 | in Records_Data.put data thy end; | 
| 660 | ||
| 661 | val get_fieldext = Symtab.lookup o #fieldext o Records_Data.get; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 662 | |
| 21962 | 663 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 664 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 665 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 666 | fun add_parents _ NONE parents = parents | 
| 15531 | 667 | | add_parents thy (SOME (types, name)) parents = | 
| 12247 | 668 | let | 
| 669 | fun err msg = error (msg ^ " parent record " ^ quote name); | |
| 12255 | 670 | |
| 35138 | 671 |         val {args, parent, fields, extension, induct_scheme, ext_def, ...} =
 | 
| 15531 | 672 | (case get_record thy name of SOME info => info | NONE => err "Unknown"); | 
| 12247 | 673 | val _ = if length types <> length args then err "Bad number of arguments for" else (); | 
| 12255 | 674 | |
| 12247 | 675 | fun bad_inst ((x, S), T) = | 
| 22578 | 676 | if Sign.of_sort thy (T, S) then NONE else SOME x | 
| 32952 | 677 | val bads = map_filter bad_inst (args ~~ types); | 
| 21962 | 678 |         val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
 | 
| 12255 | 679 | |
| 12247 | 680 | val inst = map fst args ~~ types; | 
| 17377 | 681 | val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); | 
| 15570 | 682 | val parent' = Option.map (apfst (map subst)) parent; | 
| 12247 | 683 | val fields' = map (apsnd subst) fields; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 684 | val extension' = apsnd (map subst) extension; | 
| 12247 | 685 | in | 
| 12255 | 686 | add_parents thy parent' | 
| 35138 | 687 | (make_parent_info name fields' extension' ext_def induct_scheme :: parents) | 
| 12247 | 688 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 689 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 690 | |
| 21962 | 691 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 692 | (** concrete syntax for records **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 693 | |
| 22693 | 694 | (* decode type *) | 
| 695 | ||
| 696 | fun decode_type thy t = | |
| 697 | let | |
| 35144 | 698 | fun get_sort env xi = | 
| 699 | the_default (Sign.defaultS thy) (AList.lookup (op =) env (xi: indexname)); | |
| 22693 | 700 | in | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 701 | Syntax.typ_of_term (get_sort (Syntax.term_sorts t)) t | 
| 22693 | 702 | end; | 
| 703 | ||
| 704 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 705 | (* parse translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 706 | |
| 35144 | 707 | local | 
| 708 | ||
| 35146 | 709 | fun field_type_tr ((Const (@{syntax_const "_field_type"}, _) $ Const (name, _) $ arg)) =
 | 
| 710 | (name, arg) | |
| 711 |   | field_type_tr t = raise TERM ("field_type_tr", [t]);
 | |
| 712 | ||
| 713 | fun field_types_tr (Const (@{syntax_const "_field_types"}, _) $ t $ u) =
 | |
| 714 | field_type_tr t :: field_types_tr u | |
| 715 | | field_types_tr t = [field_type_tr t]; | |
| 716 | ||
| 717 | fun record_field_types_tr more ctxt t = | |
| 17261 | 718 | let | 
| 21772 | 719 | val thy = ProofContext.theory_of ctxt; | 
| 35146 | 720 |     fun err msg = raise TERM ("Error in record-type input: " ^ msg, [t]);
 | 
| 721 | ||
| 722 | fun split_args (field :: fields) ((name, arg) :: fargs) = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 723 | if can (unsuffix name) field then | 
| 35146 | 724 | let val (args, rest) = split_args fields fargs | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 725 | in (arg :: args, rest) end | 
| 35135 | 726 |           else err ("expecting field " ^ field ^ " but got " ^ name)
 | 
| 35146 | 727 | | split_args [] (fargs as (_ :: _)) = ([], fargs) | 
| 728 | | split_args (_ :: _) [] = err "expecting more fields" | |
| 729 | | split_args _ _ = ([], []); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 730 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 731 | fun mk_ext (fargs as (name, _) :: _) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 732 | (case get_fieldext thy (Sign.intern_const thy name) of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 733 | SOME (ext, alphas) => | 
| 18858 | 734 | (case get_extfields thy ext of | 
| 35146 | 735 | SOME fields => | 
| 736 | let | |
| 737 | val fields' = but_last fields; | |
| 738 | val types = map snd fields'; | |
| 739 | val (args, rest) = split_args (map fst fields') fargs; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 740 | val argtypes = map (Sign.certify_typ thy o decode_type thy) args; | 
| 33029 | 741 | val midx = fold Term.maxidx_typ argtypes 0; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 742 | val varifyT = varifyT midx; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 743 | val vartypes = map varifyT types; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 744 | |
| 35146 | 745 | val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes) Vartab.empty | 
| 746 | handle Type.TYPE_MATCH => err "type is no proper record (extension)"; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 747 | val alphas' = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 748 | map (Syntax.term_of_typ (! Syntax.show_sorts) o Envir.norm_type subst o varifyT) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 749 | (but_last alphas); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 750 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 751 | val more' = mk_ext rest; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 752 | in | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 753 | list_comb | 
| 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 754 | (Syntax.const (Syntax.mark_type (suffix ext_typeN ext)), alphas' @ [more']) | 
| 35146 | 755 | end | 
| 35135 | 756 |               | NONE => err ("no fields defined for " ^ ext))
 | 
| 757 | | NONE => err (name ^ " is no proper field")) | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 758 | | mk_ext [] = more; | 
| 35146 | 759 | in | 
| 760 | mk_ext (field_types_tr t) | |
| 761 | end; | |
| 762 | ||
| 35363 | 763 | fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const @{type_syntax unit}) ctxt t
 | 
| 35146 | 764 |   | record_type_tr _ ts = raise TERM ("record_type_tr", ts);
 | 
| 765 | ||
| 766 | fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t | |
| 767 |   | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts);
 | |
| 768 | ||
| 35147 | 769 | |
| 770 | fun field_tr ((Const (@{syntax_const "_field"}, _) $ Const (name, _) $ arg)) = (name, arg)
 | |
| 771 |   | field_tr t = raise TERM ("field_tr", [t]);
 | |
| 772 | ||
| 773 | fun fields_tr (Const (@{syntax_const "_fields"}, _) $ t $ u) = field_tr t :: fields_tr u
 | |
| 774 | | fields_tr t = [field_tr t]; | |
| 775 | ||
| 776 | fun record_fields_tr more ctxt t = | |
| 777 | let | |
| 778 | val thy = ProofContext.theory_of ctxt; | |
| 779 |     fun err msg = raise TERM ("Error in record input: " ^ msg, [t]);
 | |
| 780 | ||
| 781 | fun split_args (field :: fields) ((name, arg) :: fargs) = | |
| 782 | if can (unsuffix name) field | |
| 783 | then | |
| 784 | let val (args, rest) = split_args fields fargs | |
| 785 | in (arg :: args, rest) end | |
| 786 |           else err ("expecting field " ^ field ^ " but got " ^ name)
 | |
| 787 | | split_args [] (fargs as (_ :: _)) = ([], fargs) | |
| 788 | | split_args (_ :: _) [] = err "expecting more fields" | |
| 789 | | split_args _ _ = ([], []); | |
| 790 | ||
| 791 | fun mk_ext (fargs as (name, _) :: _) = | |
| 792 | (case get_fieldext thy (Sign.intern_const thy name) of | |
| 793 | SOME (ext, _) => | |
| 794 | (case get_extfields thy ext of | |
| 795 | SOME fields => | |
| 796 | let | |
| 797 | val (args, rest) = split_args (map fst (but_last fields)) fargs; | |
| 798 | val more' = mk_ext rest; | |
| 35262 
9ea4445d2ccf
slightly more abstract syntax mark/unmark operations;
 wenzelm parents: 
35240diff
changeset | 799 | in list_comb (Syntax.const (Syntax.mark_const (ext ^ extN)), args @ [more']) end | 
| 35147 | 800 |               | NONE => err ("no fields defined for " ^ ext))
 | 
| 801 | | NONE => err (name ^ " is no proper field")) | |
| 802 | | mk_ext [] = more; | |
| 803 | in mk_ext (fields_tr t) end; | |
| 804 | ||
| 805 | fun record_tr ctxt [t] = record_fields_tr (Syntax.const @{const_syntax Unity}) ctxt t
 | |
| 806 |   | record_tr _ ts = raise TERM ("record_tr", ts);
 | |
| 807 | ||
| 808 | fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t | |
| 809 |   | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts);
 | |
| 810 | ||
| 811 | ||
| 812 | fun field_update_tr (Const (@{syntax_const "_field_update"}, _) $ Const (name, _) $ arg) =
 | |
| 813 |       Syntax.const (suffix updateN name) $ Abs ("_", dummyT, arg)
 | |
| 814 |   | field_update_tr t = raise TERM ("field_update_tr", [t]);
 | |
| 815 | ||
| 816 | fun field_updates_tr (Const (@{syntax_const "_field_updates"}, _) $ t $ u) =
 | |
| 817 | field_update_tr t :: field_updates_tr u | |
| 818 | | field_updates_tr t = [field_update_tr t]; | |
| 819 | ||
| 820 | fun record_update_tr [t, u] = fold (curry op $) (field_updates_tr u) t | |
| 821 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | |
| 822 | ||
| 35146 | 823 | in | 
| 15215 | 824 | |
| 24867 | 825 | val parse_translation = | 
| 35145 
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
 wenzelm parents: 
35144diff
changeset | 826 |  [(@{syntax_const "_record_update"}, record_update_tr)];
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 827 | |
| 35146 | 828 | val advanced_parse_translation = | 
| 829 |  [(@{syntax_const "_record"}, record_tr),
 | |
| 830 |   (@{syntax_const "_record_scheme"}, record_scheme_tr),
 | |
| 831 |   (@{syntax_const "_record_type"}, record_type_tr),
 | |
| 832 |   (@{syntax_const "_record_type_scheme"}, record_type_scheme_tr)];
 | |
| 833 | ||
| 834 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 835 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 836 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 837 | (* print translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 838 | |
| 32740 | 839 | val print_record_type_abbr = Unsynchronized.ref true; | 
| 840 | val print_record_type_as_fields = Unsynchronized.ref true; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 841 | |
| 35149 | 842 | |
| 843 | local | |
| 844 | ||
| 845 | (* FIXME early extern (!??) *) | |
| 846 | (* FIXME Syntax.free (??) *) | |
| 847 | fun field_type_tr' (c, t) = Syntax.const @{syntax_const "_field_type"} $ Syntax.const c $ t;
 | |
| 848 | ||
| 849 | fun field_types_tr' (t, u) = Syntax.const @{syntax_const "_field_types"} $ t $ u;
 | |
| 850 | ||
| 851 | fun record_type_tr' ctxt t = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 852 | let | 
| 35149 | 853 | val thy = ProofContext.theory_of ctxt; | 
| 854 | ||
| 855 | val T = decode_type thy t; | |
| 856 | val varifyT = varifyT (Term.maxidx_of_typ T); | |
| 857 | ||
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 858 | val term_of_type = Syntax.term_of_typ (! Syntax.show_sorts); | 
| 35149 | 859 | |
| 860 | fun strip_fields T = | |
| 861 | (case T of | |
| 35615 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 862 | Type (ext, args as _ :: _) => | 
| 35149 | 863 | (case try (unsuffix ext_typeN) ext of | 
| 864 | SOME ext' => | |
| 865 | (case get_extfields thy ext' of | |
| 35615 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 866 | SOME (fields as (x, _) :: _) => | 
| 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 867 | (case get_fieldext thy x of | 
| 35149 | 868 | SOME (_, alphas) => | 
| 869 | (let | |
| 870 | val f :: fs = but_last fields; | |
| 871 | val fields' = | |
| 872 | apfst (Sign.extern_const thy) f :: map (apfst Long_Name.base_name) fs; | |
| 873 | val (args', more) = split_last args; | |
| 874 | val alphavars = map varifyT (but_last alphas); | |
| 875 | val subst = fold (Sign.typ_match thy) (alphavars ~~ args') Vartab.empty; | |
| 876 | val fields'' = (map o apsnd) (Envir.norm_type subst o varifyT) fields'; | |
| 877 | in fields'' @ strip_fields more end | |
| 878 |                       handle Type.TYPE_MATCH => [("", T)]
 | |
| 879 |                         | Library.UnequalLengths => [("", T)])
 | |
| 35615 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 880 |                   | _ => [("", T)])
 | 
| 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 881 |               | _ => [("", T)])
 | 
| 
61bb9f8af129
record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
 wenzelm parents: 
35614diff
changeset | 882 |           | _ => [("", T)])
 | 
| 35149 | 883 |       | _ => [("", T)]);
 | 
| 884 | ||
| 885 | val (fields, (_, moreT)) = split_last (strip_fields T); | |
| 886 | val _ = null fields andalso raise Match; | |
| 887 | val u = foldr1 field_types_tr' (map (field_type_tr' o apsnd term_of_type) fields); | |
| 888 | in | |
| 889 | if not (! print_record_type_as_fields) orelse null fields then raise Match | |
| 890 |     else if moreT = HOLogic.unitT then Syntax.const @{syntax_const "_record_type"} $ u
 | |
| 891 |     else Syntax.const @{syntax_const "_record_type_scheme"} $ u $ term_of_type moreT
 | |
| 892 | end; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 893 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 894 | (*try to reconstruct the record name type abbreviation from | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 895 | the (nested) extension types*) | 
| 35149 | 896 | fun record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt tm = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 897 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 898 | val thy = ProofContext.theory_of ctxt; | 
| 35614 
d7afa8700622
record_type_abbr_tr': removed obsolete workaround for decode_type, which now retains syntactic categories of variables vs. constructors (authentic syntax);
 wenzelm parents: 
35430diff
changeset | 899 | val T = decode_type thy tm; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 900 | val midx = maxidx_of_typ T; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 901 | val varifyT = varifyT midx; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 902 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 903 | fun mk_type_abbr subst name alphas = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 904 | let val abbrT = Type (name, map (fn a => varifyT (TFree (a, Sign.defaultS thy))) alphas) in | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 905 | Syntax.term_of_typ (! Syntax.show_sorts) (Envir.norm_type subst abbrT) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 906 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 907 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 908 | fun match rT T = Sign.typ_match thy (varifyT rT, T) Vartab.empty; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 909 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 910 | if ! print_record_type_abbr then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 911 | (case last_extT T of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 912 | SOME (name, _) => | 
| 35148 | 913 | if name = last_ext then | 
| 35149 | 914 | let val subst = match schemeT T in | 
| 32335 | 915 | if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree (zeta, Sign.defaultS thy)))) | 
| 916 | then mk_type_abbr subst abbr alphas | |
| 917 | else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta]) | |
| 35149 | 918 | end handle Type.TYPE_MATCH => record_type_tr' ctxt tm | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 919 | else raise Match (*give print translation of specialised record a chance*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 920 | | _ => raise Match) | 
| 35149 | 921 | else record_type_tr' ctxt tm | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 922 | end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 923 | |
| 35149 | 924 | in | 
| 925 | ||
| 926 | fun record_ext_type_tr' name = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 927 | let | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 928 | val ext_type_name = Syntax.mark_type (suffix ext_typeN name); | 
| 35149 | 929 | fun tr' ctxt ts = | 
| 930 | record_type_tr' ctxt (list_comb (Syntax.const ext_type_name, ts)); | |
| 931 | in (ext_type_name, tr') end; | |
| 932 | ||
| 933 | fun record_ext_type_abbr_tr' abbr alphas zeta last_ext schemeT name = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 934 | let | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 935 | val ext_type_name = Syntax.mark_type (suffix ext_typeN name); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 936 | fun tr' ctxt ts = | 
| 35149 | 937 | record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt | 
| 938 | (list_comb (Syntax.const ext_type_name, ts)); | |
| 939 | in (ext_type_name, tr') end; | |
| 940 | ||
| 941 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 942 | |
| 32335 | 943 | |
| 35240 | 944 | local | 
| 945 | ||
| 946 | (* FIXME Syntax.free (??) *) | |
| 947 | fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t;
 | |
| 948 | fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u;
 | |
| 949 | ||
| 950 | fun record_tr' ctxt t = | |
| 951 | let | |
| 952 | val thy = ProofContext.theory_of ctxt; | |
| 953 | val extern = Consts.extern (ProofContext.consts_of ctxt); | |
| 954 | ||
| 955 | fun strip_fields t = | |
| 956 | (case strip_comb t of | |
| 957 | (Const (ext, _), args as (_ :: _)) => | |
| 35262 
9ea4445d2ccf
slightly more abstract syntax mark/unmark operations;
 wenzelm parents: 
35240diff
changeset | 958 | (case try (Syntax.unmark_const o unsuffix extN) ext of | 
| 35240 | 959 | SOME ext' => | 
| 960 | (case get_extfields thy ext' of | |
| 961 | SOME fields => | |
| 962 | (let | |
| 963 | val f :: fs = but_last (map fst fields); | |
| 964 | val fields' = extern f :: map Long_Name.base_name fs; | |
| 965 | val (args', more) = split_last args; | |
| 966 | in (fields' ~~ args') @ strip_fields more end | |
| 967 |                   handle Library.UnequalLengths => [("", t)])
 | |
| 968 |               | NONE => [("", t)])
 | |
| 969 |           | NONE => [("", t)])
 | |
| 970 |        | _ => [("", t)]);
 | |
| 971 | ||
| 972 | val (fields, (_, more)) = split_last (strip_fields t); | |
| 973 | val _ = null fields andalso raise Match; | |
| 974 | val u = foldr1 fields_tr' (map field_tr' fields); | |
| 975 | in | |
| 976 | case more of | |
| 977 |       Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u
 | |
| 978 |     | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more
 | |
| 979 | end; | |
| 980 | ||
| 981 | in | |
| 982 | ||
| 983 | fun record_ext_tr' name = | |
| 984 | let | |
| 35262 
9ea4445d2ccf
slightly more abstract syntax mark/unmark operations;
 wenzelm parents: 
35240diff
changeset | 985 | val ext_name = Syntax.mark_const (name ^ extN); | 
| 35240 | 986 | fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts)); | 
| 987 | in (ext_name, tr') end; | |
| 988 | ||
| 989 | end; | |
| 990 | ||
| 991 | ||
| 992 | local | |
| 993 | ||
| 994 | fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) = | |
| 995 | let | |
| 996 | val extern = Consts.extern (ProofContext.consts_of ctxt); | |
| 997 | val t = | |
| 998 | (case k of | |
| 999 | Abs (_, _, Abs (_, _, t) $ Bound 0) => | |
| 1000 | if null (loose_bnos t) then t else raise Match | |
| 1001 | | Abs (_, _, t) => | |
| 1002 | if null (loose_bnos t) then t else raise Match | |
| 1003 | | _ => raise Match); | |
| 1004 | in | |
| 35262 
9ea4445d2ccf
slightly more abstract syntax mark/unmark operations;
 wenzelm parents: 
35240diff
changeset | 1005 | (case Option.map extern (try Syntax.unmark_const c) of | 
| 35240 | 1006 | SOME update_name => | 
| 1007 | (case try (unsuffix updateN) update_name of | |
| 1008 | SOME name => | |
| 1009 |                 apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t))
 | |
| 1010 | (field_updates_tr' ctxt u) | |
| 1011 | | NONE => ([], tm)) | |
| 1012 | | NONE => ([], tm)) | |
| 1013 | end | |
| 1014 | | field_updates_tr' _ tm = ([], tm); | |
| 1015 | ||
| 1016 | fun record_update_tr' ctxt tm = | |
| 1017 | (case field_updates_tr' ctxt tm of | |
| 1018 | ([], _) => raise Match | |
| 1019 | | (ts, u) => | |
| 1020 |       Syntax.const @{syntax_const "_record_update"} $ u $
 | |
| 1021 |         foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts));
 | |
| 1022 | ||
| 1023 | in | |
| 1024 | ||
| 1025 | fun field_update_tr' name = | |
| 1026 | let | |
| 35262 
9ea4445d2ccf
slightly more abstract syntax mark/unmark operations;
 wenzelm parents: 
35240diff
changeset | 1027 | val update_name = Syntax.mark_const (name ^ updateN); | 
| 35240 | 1028 | fun tr' ctxt [t, u] = record_update_tr' ctxt (Syntax.const update_name $ t $ u) | 
| 1029 | | tr' _ _ = raise Match; | |
| 1030 | in (update_name, tr') end; | |
| 1031 | ||
| 1032 | end; | |
| 1033 | ||
| 1034 | ||
| 32335 | 1035 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1036 | (** record simprocs **) | 
| 14358 | 1037 | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1038 | fun future_forward_prf_standard thy prf prop () = | 
| 33711 | 1039 | let val thm = | 
| 1040 | if ! quick_and_dirty then Skip_Proof.make_thm thy prop | |
| 1041 | else if Goal.future_enabled () then | |
| 1042 | Goal.future_result (ProofContext.init thy) (Future.fork_pri ~1 prf) prop | |
| 1043 | else prf () | |
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1044 | in Drule.export_without_context thm end; | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1045 | |
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1046 | fun prove_common immediate stndrd thy asms prop tac = | 
| 33711 | 1047 | let | 
| 1048 | val prv = | |
| 1049 | if ! quick_and_dirty then Skip_Proof.prove | |
| 1050 | else if immediate orelse not (Goal.future_enabled ()) then Goal.prove | |
| 1051 | else Goal.prove_future; | |
| 1052 | val prf = prv (ProofContext.init thy) [] asms prop tac; | |
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1053 | in if stndrd then Drule.export_without_context prf else prf end; | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1054 | |
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1055 | val prove_future_global = prove_common false; | 
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1056 | val prove_global = prove_common true; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1057 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1058 | fun is_sel_upd_pair thy (Const (s, _)) (Const (u, t')) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1059 | (case get_updates thy u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1060 | SOME u_name => u_name = s | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1061 |   | NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')]));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1062 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1063 | fun mk_comp f g = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1064 | let | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1065 | val X = fastype_of g; | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1066 | val A = domain_type X; | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1067 | val B = range_type X; | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1068 | val C = range_type (fastype_of f); | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1069 | val T = (B --> C) --> (A --> B) --> A --> C; | 
| 35133 | 1070 |   in Const (@{const_name Fun.comp}, T) $ f $ g end;
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1071 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1072 | fun mk_comp_id f = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1073 | let val T = range_type (fastype_of f) | 
| 35133 | 1074 |   in mk_comp (Const (@{const_name Fun.id}, T --> T)) f end;
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1075 | |
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 1076 | fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1077 | | get_upd_funs _ = []; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1078 | |
| 32975 | 1079 | fun get_accupd_simps thy term defset = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1080 | let | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1081 | val (acc, [body]) = strip_comb term; | 
| 35408 | 1082 | val upd_funs = sort_distinct Term_Ord.fast_term_ord (get_upd_funs body); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1083 | fun get_simp upd = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1084 | let | 
| 35133 | 1085 | (* FIXME fresh "f" (!?) *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1086 | val T = domain_type (fastype_of upd); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1087 |         val lhs = mk_comp acc (upd $ Free ("f", T));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1088 | val rhs = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1089 | if is_sel_upd_pair thy acc upd | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1090 |           then mk_comp (Free ("f", T)) acc
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1091 | else mk_comp_id acc; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1092 | val prop = lhs === rhs; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1093 | val othm = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1094 | Goal.prove (ProofContext.init thy) [] [] prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1095 | (fn _ => | 
| 32975 | 1096 | simp_tac defset 1 THEN | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1097 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN | 
| 32975 | 1098 | TRY (simp_tac (HOL_ss addsimps id_o_apps) 1)); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1099 | val dest = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1100 | if is_sel_upd_pair thy acc upd | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1101 | then o_eq_dest | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1102 | else o_eq_id_dest; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1103 | in Drule.export_without_context (othm RS dest) end; | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 1104 | in map get_simp upd_funs end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1105 | |
| 32975 | 1106 | fun get_updupd_simp thy defset u u' comp = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1107 | let | 
| 35133 | 1108 | (* FIXME fresh "f" (!?) *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1109 |     val f = Free ("f", domain_type (fastype_of u));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1110 |     val f' = Free ("f'", domain_type (fastype_of u'));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1111 | val lhs = mk_comp (u $ f) (u' $ f'); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1112 | val rhs = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1113 | if comp | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1114 | then u $ mk_comp f f' | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1115 | else mk_comp (u' $ f') (u $ f); | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1116 | val prop = lhs === rhs; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1117 | val othm = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1118 | Goal.prove (ProofContext.init thy) [] [] prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1119 | (fn _ => | 
| 32975 | 1120 | simp_tac defset 1 THEN | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1121 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN | 
| 32975 | 1122 | TRY (simp_tac (HOL_ss addsimps [id_apply]) 1)); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1123 | val dest = if comp then o_eq_dest_lhs else o_eq_dest; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1124 | in Drule.export_without_context (othm RS dest) end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1125 | |
| 32975 | 1126 | fun get_updupd_simps thy term defset = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1127 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1128 | val upd_funs = get_upd_funs term; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1129 | val cname = fst o dest_Const; | 
| 32975 | 1130 | fun getswap u u' = get_updupd_simp thy defset u u' (cname u = cname u'); | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1131 | fun build_swaps_to_eq _ [] swaps = swaps | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1132 | | build_swaps_to_eq upd (u :: us) swaps = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1133 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1134 | val key = (cname u, cname upd); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1135 | val newswaps = | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1136 | if Symreltab.defined swaps key then swaps | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1137 | else Symreltab.insert (K true) (key, getswap u upd) swaps; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1138 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1139 | if cname u = cname upd then newswaps | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1140 | else build_swaps_to_eq upd us newswaps | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1141 | end; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1142 | fun swaps_needed [] _ _ swaps = map snd (Symreltab.dest swaps) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1143 | | swaps_needed (u :: us) prev seen swaps = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1144 | if Symtab.defined seen (cname u) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1145 | then swaps_needed us prev seen (build_swaps_to_eq u prev swaps) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1146 | else swaps_needed us (u :: prev) (Symtab.insert (K true) (cname u, ()) seen) swaps; | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1147 | in swaps_needed upd_funs [] Symtab.empty Symreltab.empty end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1148 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1149 | val named_cterm_instantiate = Iso_Tuple_Support.named_cterm_instantiate; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1150 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1151 | fun prove_unfold_defs thy ex_simps ex_simprs prop = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1152 | let | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1153 | val defset = get_sel_upd_defs thy; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1154 | val prop' = Envir.beta_eta_contract prop; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1155 | val (lhs, _) = Logic.dest_equals (Logic.strip_assums_concl prop'); | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1156 | val (_, args) = strip_comb lhs; | 
| 32975 | 1157 | val simps = (if length args = 1 then get_accupd_simps else get_updupd_simps) thy lhs defset; | 
| 16973 | 1158 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1159 | Goal.prove (ProofContext.init thy) [] [] prop' | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1160 | (fn _ => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1161 | simp_tac (HOL_basic_ss addsimps (simps @ [K_record_comp])) 1 THEN | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1162 | TRY (simp_tac (HOL_basic_ss addsimps ex_simps addsimprocs ex_simprs) 1)) | 
| 15203 | 1163 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1164 | |
| 15215 | 1165 | |
| 15059 | 1166 | local | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1167 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1168 | fun eq (s1: string) (s2: string) = (s1 = s2); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1169 | |
| 16822 | 1170 | fun has_field extfields f T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1171 | exists (fn (eN, _) => exists (eq f o fst) (Symtab.lookup_list extfields eN)) (dest_recTs T); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1172 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1173 | fun K_skeleton n (T as Type (_, [_, kT])) (b as Bound i) (Abs (x, xT, t)) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1174 | if null (loose_bnos t) then ((n, kT), (Abs (x, xT, Bound (i + 1)))) else ((n, T), b) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1175 | | K_skeleton n T b _ = ((n, T), b); | 
| 25705 | 1176 | |
| 15059 | 1177 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1178 | |
| 14255 | 1179 | (* record_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1180 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1181 | (* | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1182 | Simplify selections of an record update: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1183 | (1) S (S_update k r) = k (S r) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1184 | (2) S (X_update k r) = S r | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1185 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1186 | The simproc skips multiple updates at once, eg: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1187 | S (X_update x (Y_update y (S_update k r))) = k (S r) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1188 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1189 | But be careful in (2) because of the extensibility of records. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1190 | - If S is a more-selector we have to make sure that the update on component | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1191 | X does not affect the selected subrecord. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1192 | - If X is a more-selector we have to make sure that S is not in the updated | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1193 | subrecord. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1194 | *) | 
| 13462 | 1195 | val record_simproc = | 
| 29064 | 1196 |   Simplifier.simproc @{theory HOL} "record_simp" ["x"]
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1197 | (fn thy => fn _ => fn t => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1198 | (case t of | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1199 | (sel as Const (s, Type (_, [_, rangeS]))) $ | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1200 | ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) => | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1201 | if is_selector thy s andalso is_some (get_updates thy u) then | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1202 | let | 
| 35137 | 1203 |               val {sel_upd = {updates, ...}, extfields, ...} = Records_Data.get thy;
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1204 | |
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1205 | fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) = | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1206 | (case Symtab.lookup updates u of | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1207 | NONE => NONE | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1208 | | SOME u_name => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1209 | if u_name = s then | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1210 | (case mk_eq_terms r of | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1211 | NONE => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1212 | let | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1213 |                                 val rv = ("r", rT);
 | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1214 | val rb = Bound 0; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1215 | val (kv, kb) = K_skeleton "k" kT (Bound 1) k; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1216 | in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1217 | | SOME (trm, trm', vars) => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1218 | let | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1219 | val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1220 | in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end) | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1221 | else if has_field extfields u_name rangeS orelse | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1222 | has_field extfields s (domain_type kT) then NONE | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1223 | else | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1224 | (case mk_eq_terms r of | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1225 | SOME (trm, trm', vars) => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1226 | let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1227 | in SOME (upd $ kb $ trm, trm', kv :: vars) end | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1228 | | NONE => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1229 | let | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1230 |                                 val rv = ("r", rT);
 | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1231 | val rb = Bound 0; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1232 | val (kv, kb) = K_skeleton "k" kT (Bound 1) k; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1233 | in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end)) | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1234 | | mk_eq_terms _ = NONE; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1235 | in | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1236 | (case mk_eq_terms (upd $ k $ r) of | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1237 | SOME (trm, trm', vars) => | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1238 | SOME | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1239 | (prove_unfold_defs thy [] [] | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1240 | (list_all (vars, Logic.mk_equals (sel $ trm, trm')))) | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1241 | | NONE => NONE) | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1242 | end | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1243 | else NONE | 
| 15531 | 1244 | | _ => NONE)); | 
| 7178 | 1245 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1246 | fun get_upd_acc_cong_thm upd acc thy simpset = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1247 | let | 
| 32975 | 1248 |     val insts = [("upd", cterm_of thy upd), ("acc", cterm_of thy acc)];
 | 
| 1249 | val prop = Thm.concl_of (named_cterm_instantiate insts updacc_cong_triv); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1250 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1251 | Goal.prove (ProofContext.init thy) [] [] prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1252 | (fn _ => | 
| 32975 | 1253 | simp_tac simpset 1 THEN | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1254 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN | 
| 32975 | 1255 | TRY (resolve_tac [updacc_cong_idI] 1)) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1256 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1257 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1258 | |
| 17261 | 1259 | (* record_upd_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1260 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1261 | (*Simplify multiple updates: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1262 | (1) "N_update y (M_update g (N_update x (M_update f r))) = | 
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 1263 | (N_update (y o x) (M_update (g o f) r))" | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1264 | (2) "r(|M:= M r|) = r" | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1265 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1266 | In both cases "more" updates complicate matters: for this reason | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1267 | we omit considering further updates if doing so would introduce | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1268 | both a more update and an update to a field within it.*) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1269 | val record_upd_simproc = | 
| 29064 | 1270 |   Simplifier.simproc @{theory HOL} "record_upd_simp" ["x"]
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1271 | (fn thy => fn _ => fn t => | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1272 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1273 | (*We can use more-updators with other updators as long | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1274 | as none of the other updators go deeper than any more | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1275 | updator. min here is the depth of the deepest other | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1276 | updator, max the depth of the shallowest more updator.*) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1277 | fun include_depth (dep, true) (min, max) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1278 | if min <= dep | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1279 | then SOME (min, if dep <= max orelse max = ~1 then dep else max) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1280 | else NONE | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1281 | | include_depth (dep, false) (min, max) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1282 | if dep <= max orelse max = ~1 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1283 | then SOME (if min <= dep then dep else min, max) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1284 | else NONE; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1285 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1286 | fun getupdseq (term as (upd as Const (u, _)) $ f $ tm) min max = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1287 | (case get_update_details u thy of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1288 | SOME (s, dep, ismore) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1289 | (case include_depth (dep, ismore) (min, max) of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1290 | SOME (min', max') => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1291 | let val (us, bs, _) = getupdseq tm min' max' | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1292 | in ((upd, s, f) :: us, bs, fastype_of term) end | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1293 | | NONE => ([], term, HOLogic.unitT)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1294 | | NONE => ([], term, HOLogic.unitT)) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1295 | | getupdseq term _ _ = ([], term, HOLogic.unitT); | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1296 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1297 | val (upds, base, baseT) = getupdseq t 0 ~1; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1298 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1299 | fun is_upd_noop s (Abs (n, T, Const (s', T') $ tm')) tm = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1300 | if s = s' andalso null (loose_bnos tm') | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1301 | andalso subst_bound (HOLogic.unit, tm') = tm | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1302 | then (true, Abs (n, T, Const (s', T') $ Bound 1)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1303 | else (false, HOLogic.unit) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1304 | | is_upd_noop _ _ _ = (false, HOLogic.unit); | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1305 | |
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1306 | fun get_noop_simps (upd as Const _) (Abs (_, _, (acc as Const _) $ _)) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1307 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1308 | val ss = get_sel_upd_defs thy; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1309 | val uathm = get_upd_acc_cong_thm upd acc thy ss; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1310 | in | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1311 | [Drule.export_without_context (uathm RS updacc_noopE), | 
| 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1312 | Drule.export_without_context (uathm RS updacc_noop_compE)] | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1313 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1314 | |
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1315 | (*If f is constant then (f o g) = f. We know that K_skeleton | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1316 | only returns constant abstractions thus when we see an | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1317 | abstraction we can discard inner updates.*) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1318 | fun add_upd (f as Abs _) fs = [f] | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1319 | | add_upd f fs = (f :: fs); | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1320 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1321 | (*mk_updterm returns | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1322 | (orig-term-skeleton, simplified-skeleton, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1323 | variables, duplicate-updates, simp-flag, noop-simps) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1324 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1325 | where duplicate-updates is a table used to pass upward | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1326 | the list of update functions which can be composed | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1327 | into an update above them, simp-flag indicates whether | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1328 | any simplification was achieved, and noop-simps are | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1329 | used for eliminating case (2) defined above*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1330 | fun mk_updterm ((upd as Const (u, T), s, f) :: upds) above term = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1331 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1332 | val (lhs, rhs, vars, dups, simp, noops) = | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1333 | mk_updterm upds (Symtab.update (u, ()) above) term; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1334 | val (fvar, skelf) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1335 | K_skeleton (Long_Name.base_name s) (domain_type T) (Bound (length vars)) f; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1336 | val (isnoop, skelf') = is_upd_noop s f term; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1337 | val funT = domain_type T; | 
| 35133 | 1338 | fun mk_comp_local (f, f') = | 
| 1339 |                   Const (@{const_name Fun.comp}, funT --> funT --> funT) $ f $ f';
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1340 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1341 | if isnoop then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1342 | (upd $ skelf' $ lhs, rhs, vars, | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1343 | Symtab.update (u, []) dups, true, | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1344 | if Symtab.defined noops u then noops | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1345 | else Symtab.update (u, get_noop_simps upd skelf') noops) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1346 | else if Symtab.defined above u then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1347 | (upd $ skelf $ lhs, rhs, fvar :: vars, | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1348 | Symtab.map_default (u, []) (add_upd skelf) dups, | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1349 | true, noops) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1350 | else | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1351 | (case Symtab.lookup dups u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1352 | SOME fs => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1353 | (upd $ skelf $ lhs, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1354 | upd $ foldr1 mk_comp_local (add_upd skelf fs) $ rhs, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1355 | fvar :: vars, dups, true, noops) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1356 | | NONE => (upd $ skelf $ lhs, upd $ skelf $ rhs, fvar :: vars, dups, simp, noops)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1357 | end | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1358 | | mk_updterm [] _ _ = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1359 |               (Bound 0, Bound 0, [("r", baseT)], Symtab.empty, false, Symtab.empty)
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1360 |           | mk_updterm us _ _ = raise TERM ("mk_updterm match", map (fn (x, _, _) => x) us);
 | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1361 | |
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1362 | val (lhs, rhs, vars, _, simp, noops) = mk_updterm upds Symtab.empty base; | 
| 32952 | 1363 | val noops' = maps snd (Symtab.dest noops); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1364 | in | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1365 | if simp then | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1366 | SOME | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1367 | (prove_unfold_defs thy noops' [record_simproc] | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1368 | (list_all (vars, Logic.mk_equals (lhs, rhs)))) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1369 | else NONE | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1370 | end); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1371 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1372 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1373 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1374 | |
| 14255 | 1375 | (* record_eq_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1376 | |
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1377 | (*Look up the most specific record-equality. | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1378 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1379 | Note on efficiency: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1380 | Testing equality of records boils down to the test of equality of all components. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1381 | Therefore the complexity is: #components * complexity for single component. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1382 | Especially if a record has a lot of components it may be better to split up | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1383 | the record first and do simplification on that (record_split_simp_tac). | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1384 | e.g. r(|lots of updates|) = x | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1385 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1386 | record_eq_simproc record_split_simp_tac | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1387 | Complexity: #components * #updates #updates | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1388 | *) | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1389 | val record_eq_simproc = | 
| 29064 | 1390 |   Simplifier.simproc @{theory HOL} "record_eq_simp" ["r = s"]
 | 
| 18858 | 1391 | (fn thy => fn _ => fn t => | 
| 35133 | 1392 |       (case t of Const (@{const_name "op ="}, Type (_, [T, _])) $ _ $ _ =>
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1393 | (case rec_id ~1 T of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1394 | "" => NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1395 | | name => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1396 | (case get_equalities thy name of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1397 | NONE => NONE | 
| 35410 | 1398 |             | SOME thm => SOME (thm RS @{thm Eq_TrueI})))
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1399 | | _ => NONE)); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1400 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1401 | |
| 14255 | 1402 | (* record_split_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1403 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1404 | (*Split quantified occurrences of records, for which P holds. P can peek on the | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1405 | subterm starting at the quantified occurrence of the record (including the quantifier): | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1406 | P t = 0: do not split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1407 | P t = ~1: completely split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1408 | P t > 0: split up to given bound of record extensions.*) | 
| 14255 | 1409 | fun record_split_simproc P = | 
| 29064 | 1410 |   Simplifier.simproc @{theory HOL} "record_split_simp" ["x"]
 | 
| 18858 | 1411 | (fn thy => fn _ => fn t => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1412 | (case t of | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1413 | Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ => | 
| 35147 | 1414 |           if quantifier = @{const_name all} orelse
 | 
| 1415 |             quantifier = @{const_name All} orelse
 | |
| 35133 | 1416 |             quantifier = @{const_name Ex}
 | 
| 1417 | then | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1418 | (case rec_id ~1 T of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1419 | "" => NONE | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1420 | | _ => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1421 | let val split = P t in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1422 | if split <> 0 then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1423 | (case get_splits thy (rec_id split T) of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1424 | NONE => NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1425 | | SOME (all_thm, All_thm, Ex_thm, _) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1426 | SOME | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1427 | (case quantifier of | 
| 35133 | 1428 |                             @{const_name all} => all_thm
 | 
| 1429 |                           | @{const_name All} => All_thm RS eq_reflection
 | |
| 1430 |                           | @{const_name Ex} => Ex_thm RS eq_reflection
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1431 | | _ => error "record_split_simproc")) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1432 | else NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1433 | end) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1434 | else NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1435 | | _ => NONE)); | 
| 7178 | 1436 | |
| 14427 | 1437 | val record_ex_sel_eq_simproc = | 
| 29064 | 1438 |   Simplifier.simproc @{theory HOL} "record_ex_sel_eq_simproc" ["Ex t"]
 | 
| 18858 | 1439 | (fn thy => fn ss => fn t => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1440 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1441 | fun prove prop = | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1442 | prove_global true thy [] prop | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1443 | (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy) | 
| 35410 | 1444 |                 addsimps @{thms simp_thms} addsimprocs [record_split_simproc (K ~1)]) 1);
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1445 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1446 | fun mkeq (lr, Teq, (sel, Tsel), x) i = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1447 | if is_selector thy sel then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1448 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1449 | val x' = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1450 | if not (loose_bvar1 (x, 0)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1451 |                 then Free ("x" ^ string_of_int i, range_type Tsel)
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1452 |                 else raise TERM ("", [x]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1453 | val sel' = Const (sel, Tsel) $ Bound 0; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1454 | val (l, r) = if lr then (sel', x') else (x', sel'); | 
| 35133 | 1455 |             in Const (@{const_name "op ="}, Teq) $ l $ r end
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1456 |           else raise TERM ("", [Const (sel, Tsel)]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1457 | |
| 35133 | 1458 |         fun dest_sel_eq (Const (@{const_name "op ="}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1459 | (true, Teq, (sel, Tsel), X) | 
| 35133 | 1460 |           | dest_sel_eq (Const (@{const_name "op ="}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1461 | (false, Teq, (sel, Tsel), X) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1462 |           | dest_sel_eq _ = raise TERM ("", []);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1463 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1464 | (case t of | 
| 35133 | 1465 |           Const (@{const_name Ex}, Tex) $ Abs (s, T, t) =>
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1466 | (let | 
| 35133 | 1467 | val eq = mkeq (dest_sel_eq t) 0; | 
| 1468 | val prop = | |
| 1469 |                list_all ([("r", T)],
 | |
| 1470 | Logic.mk_equals | |
| 1471 |                   (Const (@{const_name Ex}, Tex) $ Abs (s, T, eq), HOLogic.true_const));
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1472 | in SOME (prove prop) end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1473 | handle TERM _ => NONE) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1474 | | _ => NONE) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1475 | end); | 
| 14427 | 1476 | |
| 5698 | 1477 | |
| 14255 | 1478 | (* record_split_simp_tac *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1479 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1480 | (*Split (and simplify) all records in the goal for which P holds. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1481 | For quantified occurrences of a record | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1482 | P can peek on the whole subterm (including the quantifier); for free variables P | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1483 | can only peek on the variable itself. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1484 | P t = 0: do not split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1485 | P t = ~1: completely split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1486 | P t > 0: split up to given bound of record extensions.*) | 
| 32975 | 1487 | fun record_split_simp_tac thms P = CSUBGOAL (fn (cgoal, i) => | 
| 14255 | 1488 | let | 
| 32975 | 1489 | val thy = Thm.theory_of_cterm cgoal; | 
| 1490 | ||
| 1491 | val goal = term_of cgoal; | |
| 1492 | val frees = filter (is_recT o #2) (Term.add_frees goal []); | |
| 14255 | 1493 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1494 | val has_rec = exists_Const | 
| 14255 | 1495 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 35133 | 1496 |           (s = @{const_name all} orelse s = @{const_name All} orelse s = @{const_name Ex}) andalso
 | 
| 1497 | is_recT T | |
| 14255 | 1498 | | _ => false); | 
| 1499 | ||
| 17261 | 1500 | fun mk_split_free_tac free induct_thm i = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1501 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1502 | val cfree = cterm_of thy free; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1503 | val _$ (_ $ r) = concl_of induct_thm; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1504 | val crec = cterm_of thy r; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1505 | val thm = cterm_instantiate [(crec, cfree)] induct_thm; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1506 | in | 
| 32975 | 1507 |         simp_tac (HOL_basic_ss addsimps @{thms induct_atomize}) i THEN
 | 
| 1508 | rtac thm i THEN | |
| 1509 |         simp_tac (HOL_basic_ss addsimps @{thms induct_rulify}) i
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1510 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1511 | |
| 32975 | 1512 | val split_frees_tacs = | 
| 1513 | frees |> map_filter (fn (x, T) => | |
| 1514 | (case rec_id ~1 T of | |
| 1515 | "" => NONE | |
| 1516 | | _ => | |
| 1517 | let | |
| 1518 | val free = Free (x, T); | |
| 1519 | val split = P free; | |
| 1520 | in | |
| 1521 | if split <> 0 then | |
| 1522 | (case get_splits thy (rec_id split T) of | |
| 1523 | NONE => NONE | |
| 1524 | | SOME (_, _, _, induct_thm) => | |
| 1525 | SOME (mk_split_free_tac free induct_thm i)) | |
| 1526 | else NONE | |
| 1527 | end)); | |
| 17261 | 1528 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1529 | val simprocs = if has_rec goal then [record_split_simproc P] else []; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1530 | val thms' = K_comp_convs @ thms; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1531 | in | 
| 32975 | 1532 | EVERY split_frees_tacs THEN | 
| 1533 | Simplifier.full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i | |
| 1534 | end); | |
| 14255 | 1535 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1536 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1537 | (* record_split_tac *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1538 | |
| 35147 | 1539 | (*Split all records in the goal, which are quantified by !! or ALL.*) | 
| 32975 | 1540 | val record_split_tac = CSUBGOAL (fn (cgoal, i) => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1541 | let | 
| 32975 | 1542 | val goal = term_of cgoal; | 
| 1543 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1544 | val has_rec = exists_Const | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1545 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 35147 | 1546 |           (s = @{const_name all} orelse s = @{const_name All}) andalso is_recT T
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1547 | | _ => false); | 
| 17261 | 1548 | |
| 35240 | 1549 |     fun is_all (Const (@{const_name all}, _) $ _) = ~1
 | 
| 1550 |       | is_all (Const (@{const_name All}, _) $ _) = ~1
 | |
| 1551 | | is_all _ = 0; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1552 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1553 | if has_rec goal then | 
| 32975 | 1554 | Simplifier.full_simp_tac (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i | 
| 1555 | else no_tac | |
| 1556 | end); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1557 | |
| 32335 | 1558 | |
| 6358 | 1559 | (* wrapper *) | 
| 1560 | ||
| 5707 | 1561 | val record_split_name = "record_split_tac"; | 
| 1562 | val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac); | |
| 5698 | 1563 | |
| 16330 | 1564 | |
| 1565 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1566 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1567 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1568 | (* prepare arguments *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1569 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1570 | fun read_raw_parent ctxt raw_T = | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1571 | (case ProofContext.read_typ_abbrev ctxt raw_T of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1572 | Type (name, Ts) => (Ts, name) | 
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1573 |   | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1574 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1575 | fun read_typ ctxt raw_T env = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1576 | let | 
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1577 | val ctxt' = fold (Variable.declare_typ o TFree) env ctxt; | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1578 | val T = Syntax.read_typ ctxt' raw_T; | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29265diff
changeset | 1579 | val env' = OldTerm.add_typ_tfrees (T, env); | 
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1580 | in (T, env') end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1581 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1582 | fun cert_typ ctxt raw_T env = | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1583 | let | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1584 | val thy = ProofContext.theory_of ctxt; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1585 | val T = Type.no_tvars (Sign.certify_typ thy raw_T) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1586 | handle TYPE (msg, _, _) => error msg; | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29265diff
changeset | 1587 | val env' = OldTerm.add_typ_tfrees (T, env); | 
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1588 | in (T, env') end; | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 1589 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1590 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1591 | (* attributes *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1592 | |
| 33368 | 1593 | fun case_names_fields x = Rule_Cases.case_names ["fields"] x; | 
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24712diff
changeset | 1594 | fun induct_type_global name = [case_names_fields, Induct.induct_type name]; | 
| 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24712diff
changeset | 1595 | fun cases_type_global name = [case_names_fields, Induct.cases_type name]; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1596 | |
| 32335 | 1597 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1598 | (* tactics *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1599 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1600 | 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 | 1601 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1602 | (*Do case analysis / induction according to rule on last parameter of ith subgoal | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1603 | (or on s if there are no parameters). | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1604 | Instatiation of record variable (and predicate) in rule is calculated to | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1605 | avoid problems with higher order unification.*) | 
| 32975 | 1606 | fun try_param_tac s rule = CSUBGOAL (fn (cgoal, i) => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1607 | let | 
| 32975 | 1608 | val cert = Thm.cterm_of (Thm.theory_of_cterm cgoal); | 
| 1609 | ||
| 1610 | val g = Thm.term_of cgoal; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1611 | val params = Logic.strip_params g; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1612 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g); | 
| 32975 | 1613 | val rule' = Thm.lift_rule cgoal rule; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1614 | val (P, ys) = strip_comb (HOLogic.dest_Trueprop | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1615 | (Logic.strip_assums_concl (prop_of rule'))); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1616 | (*ca indicates if rule is a case analysis or induction rule*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1617 | val (x, ca) = | 
| 33957 | 1618 | (case rev (drop (length params) ys) of | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1619 | [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1620 | (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 | 1621 | | [x] => (head_of x, false)); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1622 | val rule'' = cterm_instantiate (map (pairself cert) | 
| 32975 | 1623 | (case rev params of | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1624 | [] => | 
| 32975 | 1625 | (case AList.lookup (op =) (Term.add_frees g []) s of | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1626 | NONE => sys_error "try_param_tac: no such variable" | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1627 | | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))]) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1628 | | (_, T) :: _ => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1629 | [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))), | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1630 | (x, list_abs (params, Bound 0))])) rule'; | 
| 32975 | 1631 | in compose_tac (false, rule'', nprems_of rule) i end); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1632 | |
| 15215 | 1633 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1634 | fun extension_definition name fields alphas zeta moreT more vars thy = | 
| 17261 | 1635 | let | 
| 35239 | 1636 | val base_name = Long_Name.base_name name; | 
| 1637 | ||
| 32977 | 1638 | val fieldTs = map snd fields; | 
| 35239 | 1639 | val fields_moreTs = fieldTs @ [moreT]; | 
| 1640 | ||
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1641 | val alphas_zeta = alphas @ [zeta]; | 
| 32977 | 1642 | val alphas_zetaTs = map (fn a => TFree (a, HOLogic.typeS)) alphas_zeta; | 
| 35239 | 1643 | |
| 1644 | val ext_binding = Binding.name (suffix extN base_name); | |
| 1645 | val ext_name = suffix extN name; | |
| 1646 | val extT = Type (suffix ext_typeN name, alphas_zetaTs); | |
| 1647 | val ext_type = fields_moreTs ---> extT; | |
| 1648 | ||
| 1649 | ||
| 1650 | (* the tree of new types that will back the record extension *) | |
| 32767 | 1651 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1652 | val mktreeV = Balanced_Tree.make Iso_Tuple_Support.mk_cons_tuple; | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1653 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1654 | fun mk_iso_tuple (left, right) (thy, i) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1655 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1656 | val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i; | 
| 35239 | 1657 | val ((_, cons), thy') = thy | 
| 1658 | |> Iso_Tuple_Support.add_iso_tuple_type | |
| 1659 | (suffix suff base_name, alphas_zeta) (fastype_of left, fastype_of right); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1660 | in | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1661 | (cons $ left $ right, (thy', i + 1)) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1662 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1663 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1664 | (*trying to create a 1-element iso_tuple will fail, and is pointless anyway*) | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1665 | fun mk_even_iso_tuple [arg] = pair arg | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1666 | | mk_even_iso_tuple args = mk_iso_tuple (Iso_Tuple_Support.dest_cons_tuple (mktreeV args)); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1667 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1668 | fun build_meta_tree_type i thy vars more = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1669 | let val len = length vars in | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1670 |         if len < 1 then raise TYPE ("meta_tree_type args too short", [], vars)
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1671 | else if len > 16 then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1672 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1673 | fun group16 [] = [] | 
| 33957 | 1674 | | group16 xs = take 16 xs :: group16 (drop 16 xs); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1675 | val vars' = group16 vars; | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1676 | val (composites, (thy', i')) = fold_map mk_even_iso_tuple vars' (thy, i); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1677 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1678 | build_meta_tree_type i' thy' composites more | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1679 | end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1680 | else | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1681 | let val (term, (thy', _)) = mk_iso_tuple (mktreeV vars, more) (thy, 0) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1682 | in (term, thy') end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1683 | end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1684 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1685 | val _ = timing_msg "record extension preparing definitions"; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1686 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1687 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1688 | (* 1st stage part 1: introduce the tree of new types *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1689 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1690 | fun get_meta_tree () = build_meta_tree_type 1 thy vars more; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1691 | val (ext_body, typ_thy) = | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1692 | timeit_msg "record extension nested type def:" get_meta_tree; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1693 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1694 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1695 | (* prepare declarations and definitions *) | 
| 17261 | 1696 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1697 | (* 1st stage part 2: define the ext constant *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1698 | |
| 35239 | 1699 | fun mk_ext args = list_comb (Const (ext_name, ext_type), args); | 
| 1700 | val ext_spec = Logic.mk_equals (mk_ext (vars @ [more]), ext_body); | |
| 1701 | ||
| 16379 | 1702 | fun mk_defs () = | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1703 | typ_thy | 
| 35239 | 1704 | |> Sign.declare_const ((ext_binding, ext_type), NoSyn) |> snd | 
| 1705 | |> PureThy.add_defs false [((Thm.def_binding ext_binding, ext_spec), [])] | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1706 | ||> Theory.checkpoint | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1707 | val ([ext_def], defs_thy) = | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1708 | timeit_msg "record extension constructor def:" mk_defs; | 
| 17261 | 1709 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1710 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1711 | (* prepare propositions *) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1712 | |
| 16379 | 1713 | val _ = timing_msg "record extension preparing propositions"; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1714 | val vars_more = vars @ [more]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1715 | val variants = map (fn Free (x, _) => x) vars_more; | 
| 15215 | 1716 | val ext = mk_ext vars_more; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1717 | val s = Free (rN, extT); | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1718 | val P = Free (Name.variant variants "P", extT --> HOLogic.boolT); | 
| 17261 | 1719 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1720 | val inject_prop = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1721 | let val vars_more' = map (fn (Free (x, T)) => Free (x ^ "'", T)) vars_more in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1722 | HOLogic.mk_conj (HOLogic.eq_const extT $ | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1723 | mk_ext vars_more $ mk_ext vars_more', HOLogic.true_const) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1724 | === | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1725 | foldr1 HOLogic.mk_conj | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1726 | (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [HOLogic.true_const]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1727 | end; | 
| 17261 | 1728 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1729 | val induct_prop = | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1730 | (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 | 1731 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1732 | val split_meta_prop = | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1733 | let val P = Free (Name.variant variants "P", extT --> Term.propT) in | 
| 17261 | 1734 | Logic.mk_equals | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1735 | (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext)) | 
| 17261 | 1736 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1737 | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 1738 | val prove_standard = prove_future_global true defs_thy; | 
| 17261 | 1739 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1740 | fun inject_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1741 | simplify HOL_ss | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1742 | (prove_standard [] inject_prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1743 | (fn _ => | 
| 32975 | 1744 | simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN | 
| 1745 | REPEAT_DETERM | |
| 1746 | (rtac refl_conj_eq 1 ORELSE | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1747 | Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE | 
| 32975 | 1748 | rtac refl 1))); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1749 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1750 | 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 | 1751 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1752 | (*We need a surjection property r = (| f = f r, g = g r ... |) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1753 | to prove other theorems. We haven't given names to the accessors | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1754 | f, g etc yet however, so we generate an ext structure with | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1755 | free variables as all arguments and allow the introduction tactic to | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1756 | operate on it as far as it can. We then use Drule.export_without_context | 
| 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1757 | to convert the free variables into unifiable variables and unify them with | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1758 | (roughly) the definition of the accessor.*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1759 | fun surject_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1760 | let | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1761 | val cterm_ext = cterm_of defs_thy ext; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1762 |         val start = named_cterm_instantiate [("y", cterm_ext)] surject_assist_idE;
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1763 | val tactic1 = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1764 | simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1765 | REPEAT_ALL_NEW Iso_Tuple_Support.iso_tuple_intros_tac 1; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1766 | val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1); | 
| 32972 | 1767 | val [halfway] = Seq.list_of (tactic1 start); | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 1768 | val [surject] = Seq.list_of (tactic2 (Drule.export_without_context halfway)); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1769 | in | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1770 | surject | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1771 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1772 | val surject = timeit_msg "record extension surjective proof:" surject_prf; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1773 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1774 | fun split_meta_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1775 | prove_standard [] split_meta_prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1776 | (fn _ => | 
| 32975 | 1777 | EVERY1 | 
| 1778 | [rtac equal_intr_rule, Goal.norm_hhf_tac, | |
| 1779 | etac meta_allE, atac, | |
| 1780 | rtac (prop_subst OF [surject]), | |
| 1781 | REPEAT o etac meta_allE, atac]); | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1782 | val split_meta = timeit_msg "record extension split_meta proof:" split_meta_prf; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1783 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1784 | fun induct_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1785 | let val (assm, concl) = induct_prop in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1786 | prove_standard [assm] concl | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1787 |           (fn {prems, ...} =>
 | 
| 32975 | 1788 | cut_rules_tac [split_meta RS Drule.equal_elim_rule2] 1 THEN | 
| 1789 | resolve_tac prems 2 THEN | |
| 1790 | asm_simp_tac HOL_ss 1) | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1791 | end; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1792 | 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 | 1793 | |
| 35138 | 1794 | val ([induct', inject', surjective', split_meta'], thm_thy) = | 
| 17261 | 1795 | defs_thy | 
| 33612 | 1796 | |> PureThy.add_thms (map (Thm.no_attributes o apfst Binding.name) | 
| 35138 | 1797 |            [("ext_induct", induct),
 | 
| 1798 |             ("ext_inject", inject),
 | |
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 1799 |             ("ext_surjective", surject),
 | 
| 33612 | 1800 |             ("ext_split", split_meta)])
 | 
| 1801 | ||> Code.add_default_eqn ext_def; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1802 | |
| 35138 | 1803 | in ((extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1804 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1805 | fun chunks [] [] = [] | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1806 | | chunks [] xs = [xs] | 
| 33957 | 1807 | | chunks (l :: ls) xs = take l xs :: chunks ls (drop l xs); | 
| 17261 | 1808 | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1809 | fun chop_last [] = error "chop_last: list should not be empty" | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1810 | | chop_last [x] = ([], x) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1811 | | chop_last (x :: xs) = let val (tl, l) = chop_last xs in (x :: tl, l) end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1812 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1813 | fun subst_last _ [] = error "subst_last: list should not be empty" | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1814 | | subst_last s [_] = [s] | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1815 | | subst_last s (x :: xs) = x :: subst_last s xs; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1816 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1817 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1818 | (* mk_recordT *) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1819 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1820 | (*builds up the record type from the current extension tpye extT and a list | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1821 | of parent extensions, starting with the root of the record hierarchy*) | 
| 21078 | 1822 | fun mk_recordT extT = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1823 | fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT; | 
| 15215 | 1824 | |
| 1825 | ||
| 1826 | fun obj_to_meta_all thm = | |
| 1827 | let | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1828 | fun E thm = (* FIXME proper name *) | 
| 35239 | 1829 | (case SOME (spec OF [thm]) handle THM _ => NONE of | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1830 | SOME thm' => E thm' | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1831 | | NONE => thm); | 
| 15215 | 1832 | val th1 = E thm; | 
| 1833 | val th2 = Drule.forall_intr_vars th1; | |
| 1834 | in th2 end; | |
| 1835 | ||
| 1836 | fun meta_to_obj_all thm = | |
| 1837 | let | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26496diff
changeset | 1838 | val thy = Thm.theory_of_thm thm; | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26496diff
changeset | 1839 | val prop = Thm.prop_of thm; | 
| 15215 | 1840 | val params = Logic.strip_params prop; | 
| 1841 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl prop); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1842 | val ct = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.list_all (params, concl))); | 
| 15215 | 1843 | val thm' = Seq.hd (REPEAT (rtac allI 1) (Thm.trivial ct)); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1844 | in Thm.implies_elim thm' thm end; | 
| 15215 | 1845 | |
| 1846 | ||
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1847 | (* record_definition *) | 
| 25070 
e2a39b6526b0
replaced NameSpace.accesses' by NameSpace.external_names (depening on naming);
 wenzelm parents: 
24867diff
changeset | 1848 | |
| 35239 | 1849 | fun record_definition (args, binding) parent (parents: parent_info list) raw_fields thy = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1850 | let | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1851 | val alphas = map fst args; | 
| 35136 | 1852 | |
| 35239 | 1853 | val name = Sign.full_name thy binding; | 
| 1854 | val full = Sign.full_name_path thy (Binding.name_of binding); (* FIXME Binding.qualified (!?) *) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1855 | |
| 35136 | 1856 | val bfields = map (fn (x, T, _) => (x, T)) raw_fields; | 
| 1857 | val field_syntax = map #3 raw_fields; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1858 | |
| 32952 | 1859 | val parent_fields = maps #fields parents; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1860 | val parent_chunks = map (length o #fields) parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1861 | val parent_names = map fst parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1862 | val parent_types = map snd parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1863 | val parent_fields_len = length parent_fields; | 
| 35239 | 1864 | val parent_variants = | 
| 1865 | Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1866 | 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 | 1867 | val parent_len = length parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1868 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1869 | val fields = map (apfst full) bfields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1870 | val names = map fst fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1871 | val types = map snd fields; | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1872 | val alphas_fields = fold Term.add_tfree_namesT types []; | 
| 33049 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 haftmann parents: 
33039diff
changeset | 1873 | val alphas_ext = inter (op =) alphas_fields alphas; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1874 | val len = length fields; | 
| 30715 
e23e15f52d42
avoid mixing of left/right associative infixes, to make it work with experimental Poly/ML 5.3 branch;
 wenzelm parents: 
30364diff
changeset | 1875 | val variants = | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1876 | Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) | 
| 35136 | 1877 | (map (Binding.name_of o fst) bfields); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1878 | val vars = ListPair.map Free (variants, types); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1879 | val named_vars = names ~~ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1880 | val idxms = 0 upto len; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1881 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1882 | val all_fields = parent_fields @ fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1883 | val all_types = parent_types @ types; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1884 | val all_variants = parent_variants @ variants; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1885 | val all_vars = parent_vars @ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1886 | 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 | 1887 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1888 | |
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 1889 | val zeta = Name.variant alphas "'z"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1890 | val moreT = TFree (zeta, HOLogic.typeS); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1891 | val more = Free (moreN, moreT); | 
| 35136 | 1892 | val full_moreN = full (Binding.name moreN); | 
| 1893 | val bfields_more = bfields @ [(Binding.name moreN, moreT)]; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1894 | val fields_more = fields @ [(full_moreN, moreT)]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1895 | val named_vars_more = named_vars @ [(full_moreN, more)]; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1896 | val all_vars_more = all_vars @ [more]; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1897 | val all_named_vars_more = all_named_vars @ [(full_moreN, more)]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1898 | |
| 17261 | 1899 | |
| 35239 | 1900 | (* 1st stage: ext_thy *) | 
| 1901 | ||
| 1902 | val extension_name = full binding; | |
| 1903 | ||
| 1904 | val ((extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1905 | thy | 
| 35239 | 1906 | |> Sign.qualified_path false binding | 
| 1907 | |> extension_definition extension_name fields alphas_ext zeta moreT more vars; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1908 | |
| 17261 | 1909 | val _ = timing_msg "record preparing definitions"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1910 | val Type extension_scheme = extT; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1911 | val extension_name = unsuffix ext_typeN (fst extension_scheme); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1912 | val extension = let val (n, Ts) = extension_scheme in (n, subst_last HOLogic.unitT Ts) end; | 
| 35239 | 1913 | val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extension_name]; | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1914 | val extension_id = implode extension_names; | 
| 17261 | 1915 | |
| 33957 | 1916 | fun rec_schemeT n = mk_recordT (map #extension (drop n parents)) extT; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1917 | val rec_schemeT0 = rec_schemeT 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1918 | |
| 17261 | 1919 | fun recT n = | 
| 32972 | 1920 | let val (c, Ts) = extension in | 
| 33957 | 1921 | mk_recordT (map #extension (drop n parents)) | 
| 32972 | 1922 | (Type (c, subst_last HOLogic.unitT Ts)) | 
| 1923 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1924 | val recT0 = recT 0; | 
| 17261 | 1925 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1926 | fun mk_rec args n = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1927 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1928 | val (args', more) = chop_last args; | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1929 | fun mk_ext' ((name, T), args) more = mk_ext (name, T) (args @ [more]); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1930 | fun build Ts = | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1931 | fold_rev mk_ext' (drop n ((extension_names ~~ Ts) ~~ chunks parent_chunks args')) more; | 
| 17261 | 1932 | in | 
| 1933 | if more = HOLogic.unit | |
| 33063 | 1934 | then build (map_range recT (parent_len + 1)) | 
| 1935 | else build (map_range rec_schemeT (parent_len + 1)) | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1936 | end; | 
| 17261 | 1937 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1938 | val r_rec0 = mk_rec all_vars_more 0; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1939 | val r_rec_unit0 = mk_rec (all_vars @ [HOLogic.unit]) 0; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1940 | |
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1941 | fun r n = Free (rN, rec_schemeT n); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1942 | val r0 = r 0; | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1943 | fun r_unit n = Free (rN, recT n); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1944 | val r_unit0 = r_unit 0; | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1945 | val w = Free (wN, rec_schemeT 0); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1946 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1947 | |
| 35239 | 1948 | (* print translations *) | 
| 1949 | ||
| 35149 | 1950 | val record_ext_type_abbr_tr's = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1951 | let | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1952 | val trname = hd extension_names; | 
| 35148 | 1953 | val last_ext = unsuffix ext_typeN (fst extension); | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1954 | in [record_ext_type_abbr_tr' name alphas zeta last_ext rec_schemeT0 trname] end; | 
| 35149 | 1955 | |
| 1956 | val record_ext_type_tr's = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1957 | let | 
| 35149 | 1958 | (*avoid conflict with record_type_abbr_tr's*) | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1959 | val trnames = if parent_len > 0 then [extension_name] else []; | 
| 35149 | 1960 | in map record_ext_type_tr' trnames end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1961 | |
| 35239 | 1962 | val advanced_print_translation = | 
| 1963 | map field_update_tr' (full_moreN :: names) @ [record_ext_tr' extension_name] @ | |
| 1964 | record_ext_type_tr's @ record_ext_type_abbr_tr's; | |
| 1965 | ||
| 17261 | 1966 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1967 | (* prepare declarations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1968 | |
| 35136 | 1969 | val sel_decls = map (mk_selC rec_schemeT0 o apfst Binding.name_of) bfields_more; | 
| 1970 | val upd_decls = map (mk_updC updateN rec_schemeT0 o apfst Binding.name_of) bfields_more; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1971 | val make_decl = (makeN, all_types ---> recT0); | 
| 17261 | 1972 | val fields_decl = (fields_selN, types ---> Type extension); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1973 | val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1974 | val truncate_decl = (truncateN, rec_schemeT0 --> recT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1975 | |
| 35133 | 1976 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1977 | (* prepare definitions *) | 
| 17261 | 1978 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1979 | (*record (scheme) type abbreviation*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1980 | val recordT_specs = | 
| 35239 | 1981 | [(Binding.suffix_name schemeN binding, alphas @ [zeta], rec_schemeT0, NoSyn), | 
| 1982 | (binding, alphas, recT0, NoSyn)]; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1983 | |
| 35138 | 1984 | val ext_defs = ext_def :: map #ext_def parents; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1985 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1986 | (*Theorems from the iso_tuple intros. | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1987 | By unfolding ext_defs from r_rec0 we create a tree of constructor | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1988 | calls (many of them Pair, but others as well). The introduction | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1989 | rules for update_accessor_eq_assist can unify two different ways | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1990 | on these constructors. If we take the complete result sequence of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1991 | running a the introduction tactic, we get one theorem for each upd/acc | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1992 | pair, from which we can derive the bodies of our selector and | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1993 | updator and their convs.*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1994 | fun get_access_update_thms () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1995 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1996 | val r_rec0_Vars = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1997 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1998 | (*pick variable indices of 1 to avoid possible variable | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1999 | collisions with existing variables in updacc_eq_triv*) | 
| 32757 
4e97fc468a53
Avoid a possible variable name conflict in instantiating a theorem.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32752diff
changeset | 2000 | fun to_Var (Free (c, T)) = Var ((c, 1), T); | 
| 
4e97fc468a53
Avoid a possible variable name conflict in instantiating a theorem.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32752diff
changeset | 2001 | in mk_rec (map to_Var all_vars_more) 0 end; | 
| 
4e97fc468a53
Avoid a possible variable name conflict in instantiating a theorem.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32752diff
changeset | 2002 | |
| 35239 | 2003 | val cterm_rec = cterm_of ext_thy r_rec0; | 
| 2004 | val cterm_vrs = cterm_of ext_thy r_rec0_Vars; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2005 |         val insts = [("v", cterm_rec), ("v'", cterm_vrs)];
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2006 | val init_thm = named_cterm_instantiate insts updacc_eq_triv; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2007 | val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2008 | val tactic = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2009 | simp_tac (HOL_basic_ss addsimps ext_defs) 1 THEN | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 2010 | REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE terminal); | 
| 32972 | 2011 | val updaccs = Seq.list_of (tactic init_thm); | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2012 | in | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2013 | (updaccs RL [updacc_accessor_eqE], | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2014 | updaccs RL [updacc_updator_eqE], | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2015 | updaccs RL [updacc_cong_from_eq]) | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2016 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2017 | val (accessor_thms, updator_thms, upd_acc_cong_assists) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2018 | timeit_msg "record getting tree access/updates:" get_access_update_thms; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2019 | |
| 33957 | 2020 | fun lastN xs = drop parent_fields_len xs; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2021 | |
| 17261 | 2022 | (*selectors*) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2023 | fun mk_sel_spec ((c, T), thm) = | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2024 | let | 
| 35239 | 2025 | val (acc $ arg, _) = | 
| 2026 | HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm))); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2027 | val _ = | 
| 35239 | 2028 | if arg aconv r_rec0 then () | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2029 |           else raise TERM ("mk_sel_spec: different arg", [arg]);
 | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2030 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2031 | Const (mk_selC rec_schemeT0 (c, T)) :== acc | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2032 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2033 | val sel_specs = map mk_sel_spec (fields_more ~~ lastN accessor_thms); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2034 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2035 | (*updates*) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2036 | fun mk_upd_spec ((c, T), thm) = | 
| 17261 | 2037 | let | 
| 35239 | 2038 | val (upd $ _ $ arg, _) = | 
| 2039 | HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm))); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2040 | val _ = | 
| 35135 | 2041 | if arg aconv r_rec0 then () | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2042 |           else raise TERM ("mk_sel_spec: different arg", [arg]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2043 | in Const (mk_updC updateN rec_schemeT0 (c, T)) :== upd end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2044 | val upd_specs = map mk_upd_spec (fields_more ~~ lastN updator_thms); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2045 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2046 | (*derived operations*) | 
| 35144 | 2047 | val make_spec = | 
| 2048 | list_comb (Const (full (Binding.name makeN), all_types ---> recT0), all_vars) :== | |
| 2049 | mk_rec (all_vars @ [HOLogic.unit]) 0; | |
| 2050 | val fields_spec = | |
| 2051 | list_comb (Const (full (Binding.name fields_selN), types ---> Type extension), vars) :== | |
| 2052 | mk_rec (all_vars @ [HOLogic.unit]) parent_len; | |
| 17261 | 2053 | val extend_spec = | 
| 35136 | 2054 | Const (full (Binding.name extendN), recT0 --> moreT --> rec_schemeT0) $ r_unit0 $ more :== | 
| 35144 | 2055 | mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0; | 
| 2056 | val truncate_spec = | |
| 2057 | Const (full (Binding.name truncateN), rec_schemeT0 --> recT0) $ r0 :== | |
| 2058 | mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2059 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2060 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2061 | (* 2st stage: defs_thy *) | 
| 17261 | 2062 | |
| 16379 | 2063 | fun mk_defs () = | 
| 35239 | 2064 | ext_thy | 
| 2065 | |> Sign.add_advanced_trfuns ([], [], advanced_print_translation, []) | |
| 2066 | |> Sign.restore_naming thy | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 2067 | |> Sign.add_tyabbrs_i recordT_specs | 
| 35239 | 2068 | |> Sign.qualified_path false binding | 
| 2069 | |> fold (fn ((x, T), mx) => snd o Sign.declare_const ((Binding.name x, T), mx)) | |
| 2070 | (sel_decls ~~ (field_syntax @ [NoSyn])) | |
| 2071 | |> fold (fn (x, T) => snd o Sign.declare_const ((Binding.name x, T), NoSyn)) | |
| 2072 | (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl]) | |
| 35142 | 2073 | |> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name))) | 
| 2074 | sel_specs | |
| 2075 | ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name))) | |
| 2076 | upd_specs | |
| 2077 | ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name))) | |
| 2078 | [make_spec, fields_spec, extend_spec, truncate_spec] | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2079 | |-> | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2080 | (fn defs as ((sel_defs, upd_defs), derived_defs) => | 
| 28370 | 2081 | fold Code.add_default_eqn sel_defs | 
| 2082 | #> fold Code.add_default_eqn upd_defs | |
| 2083 | #> fold Code.add_default_eqn derived_defs | |
| 22747 
0c9c413b4678
add definitions explicitly to code generator table
 haftmann parents: 
22693diff
changeset | 2084 | #> pair defs) | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2085 | ||> Theory.checkpoint | 
| 22747 
0c9c413b4678
add definitions explicitly to code generator table
 haftmann parents: 
22693diff
changeset | 2086 | val (((sel_defs, upd_defs), derived_defs), defs_thy) = | 
| 
0c9c413b4678
add definitions explicitly to code generator table
 haftmann parents: 
22693diff
changeset | 2087 | timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:" | 
| 
0c9c413b4678
add definitions explicitly to code generator table
 haftmann parents: 
22693diff
changeset | 2088 | mk_defs; | 
| 17261 | 2089 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2090 | (* prepare propositions *) | 
| 17261 | 2091 | val _ = timing_msg "record preparing propositions"; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2092 | val P = Free (Name.variant all_variants "P", rec_schemeT0 --> HOLogic.boolT); | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 2093 | val C = Free (Name.variant all_variants "C", HOLogic.boolT); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2094 | 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 | 2095 | |
| 17261 | 2096 | (*selectors*) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2097 | val sel_conv_props = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2098 | map (fn (c, x as Free (_, T)) => mk_sel r_rec0 (c, T) === x) named_vars_more; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2099 | |
| 17261 | 2100 | (*updates*) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2101 | fun mk_upd_prop (i, (c, T)) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2102 | let | 
| 35239 | 2103 | val x' = Free (Name.variant all_variants (Long_Name.base_name c ^ "'"), T --> T); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2104 | val n = parent_fields_len + i; | 
| 35239 | 2105 | val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2106 | 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 | 2107 | 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 | 2108 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2109 | (*induct*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2110 | val induct_scheme_prop = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2111 | All (map dest_Free all_vars_more) (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0); | 
| 17261 | 2112 | val induct_prop = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2113 | (All (map dest_Free all_vars) (Trueprop (P_unit $ r_rec_unit0)), | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2114 | Trueprop (P_unit $ r_unit0)); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2115 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2116 | (*surjective*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2117 | val surjective_prop = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2118 | let val args = map (fn (c, Free (_, T)) => mk_sel r0 (c, T)) all_named_vars_more | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2119 | in r0 === mk_rec args 0 end; | 
| 17261 | 2120 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2121 | (*cases*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2122 | val cases_scheme_prop = | 
| 32809 
e72347dd3e64
avoid mixed l/r infixes, which do not work in some versions of SML;
 wenzelm parents: 
32808diff
changeset | 2123 | (All (map dest_Free all_vars_more) ((r0 === r_rec0) ==> Trueprop C)) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 2124 | ==> Trueprop C; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2125 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2126 | val cases_prop = | 
| 32809 
e72347dd3e64
avoid mixed l/r infixes, which do not work in some versions of SML;
 wenzelm parents: 
32808diff
changeset | 2127 | (All (map dest_Free all_vars) ((r_unit0 === r_rec_unit0) ==> Trueprop C)) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 2128 | ==> Trueprop C; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2129 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2130 | (*split*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2131 | val split_meta_prop = | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20049diff
changeset | 2132 | let val P = Free (Name.variant all_variants "P", rec_schemeT0-->Term.propT) in | 
| 17261 | 2133 | Logic.mk_equals | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2134 | (All [dest_Free r0] (P $ r0), All (map dest_Free all_vars_more) (P $ r_rec0)) | 
| 17261 | 2135 | end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2136 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2137 | val split_object_prop = | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2138 | let val ALL = fold_rev (fn (v, T) => fn t => HOLogic.mk_all (v, T, t)) | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2139 | in ALL [dest_Free r0] (P $ r0) === ALL (map dest_Free all_vars_more) (P $ r_rec0) end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2140 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2141 | val split_ex_prop = | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2142 | let val EX = fold_rev (fn (v, T) => fn t => HOLogic.mk_exists (v, T, t)) | 
| 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2143 | in EX [dest_Free r0] (P $ r0) === EX (map dest_Free all_vars_more) (P $ r_rec0) end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2144 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2145 | (*equality*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2146 | val equality_prop = | 
| 17261 | 2147 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2148 | val s' = Free (rN ^ "'", rec_schemeT0); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2149 | fun mk_sel_eq (c, Free (_, T)) = mk_sel r0 (c, T) === mk_sel s' (c, T); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2150 | val seleqs = map mk_sel_eq all_named_vars_more; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2151 | in All (map dest_Free [r0, s']) (Logic.list_implies (seleqs, r0 === s')) end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2152 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2153 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2154 | (* 3rd stage: thms_thy *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2155 | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2156 | fun prove stndrd = prove_future_global stndrd defs_thy; | 
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2157 | val prove_standard = prove_future_global true defs_thy; | 
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2158 | val future_forward_prf = future_forward_prf_standard defs_thy; | 
| 17261 | 2159 | |
| 15215 | 2160 | fun prove_simp stndrd ss simps = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2161 | let val tac = simp_all_tac ss simps | 
| 15215 | 2162 | 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 | 2163 | |
| 17510 | 2164 | val ss = get_simpset defs_thy; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2165 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2166 | fun sel_convs_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2167 | map (prove_simp false ss (sel_defs @ accessor_thms)) sel_conv_props; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2168 | val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 2169 | fun sel_convs_standard_prf () = map Drule.export_without_context sel_convs; | 
| 17261 | 2170 | val sel_convs_standard = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2171 | timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2172 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2173 | fun upd_convs_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2174 | map (prove_simp false ss (upd_defs @ updator_thms)) upd_conv_props; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2175 | val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 2176 | fun upd_convs_standard_prf () = map Drule.export_without_context upd_convs; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2177 | val upd_convs_standard = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2178 | timeit_msg "record upd_convs_standard proof:" upd_convs_standard_prf; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2179 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2180 | fun get_upd_acc_congs () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2181 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2182 | val symdefs = map symmetric (sel_defs @ upd_defs); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2183 | val fold_ss = HOL_basic_ss addsimps symdefs; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 2184 | val ua_congs = map (Drule.export_without_context o simplify fold_ss) upd_acc_cong_assists; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2185 | in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2186 | val (fold_congs, unfold_congs) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2187 | timeit_msg "record upd fold/unfold congs:" get_upd_acc_congs; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2188 | |
| 35138 | 2189 | val parent_induct = Option.map #induct_scheme (try List.last parents); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2190 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2191 | fun induct_scheme_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2192 | prove_standard [] induct_scheme_prop | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2193 | (fn _ => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2194 | EVERY | 
| 35138 | 2195 | [case parent_induct of NONE => all_tac | SOME ind => try_param_tac rN ind 1, | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2196 | try_param_tac rN ext_induct 1, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2197 | asm_simp_tac HOL_basic_ss 1]); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2198 | 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 | 2199 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2200 | fun induct_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2201 | let val (assm, concl) = induct_prop in | 
| 20248 | 2202 |         prove_standard [assm] concl (fn {prems, ...} =>
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2203 | try_param_tac rN induct_scheme 1 | 
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26957diff
changeset | 2204 |           THEN try_param_tac "more" @{thm unit.induct} 1
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2205 | THEN resolve_tac prems 1) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2206 | end; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2207 | 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 | 2208 | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2209 | fun cases_scheme_prf () = | 
| 17261 | 2210 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2211 | val _ $ (Pvar $ _) = concl_of induct_scheme; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2212 | val ind = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2213 | cterm_instantiate | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2214 | [(cterm_of defs_thy Pvar, cterm_of defs_thy | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2215 | (lambda w (HOLogic.imp $ HOLogic.mk_eq (r0, w) $ C)))] | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2216 | induct_scheme; | 
| 35625 | 2217 | in Object_Logic.rulify (mp OF [ind, refl]) end; | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2218 | |
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2219 | val cases_scheme_prf = future_forward_prf cases_scheme_prf cases_scheme_prop; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2220 | 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 | 2221 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2222 | fun cases_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2223 | prove_standard [] cases_prop | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2224 | (fn _ => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2225 | try_param_tac rN cases_scheme 1 THEN | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2226 | simp_all_tac HOL_basic_ss [unit_all_eq1]); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2227 | 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 | 2228 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2229 | fun surjective_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2230 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2231 | val leaf_ss = get_sel_upd_defs defs_thy addsimps (sel_defs @ (o_assoc :: id_o_apps)); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2232 | val init_ss = HOL_basic_ss addsimps ext_defs; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2233 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2234 | prove_standard [] surjective_prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 2235 | (fn _ => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2236 | EVERY | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2237 | [rtac surject_assist_idE 1, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2238 | simp_tac init_ss 1, | 
| 32975 | 2239 | REPEAT | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 2240 | (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE | 
| 32975 | 2241 | (rtac surject_assistI 1 THEN simp_tac leaf_ss 1))]) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2242 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2243 | val surjective = timeit_msg "record surjective proof:" surjective_prf; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2244 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2245 | fun split_meta_prf () = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2246 | prove false [] split_meta_prop | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 2247 | (fn _ => | 
| 32975 | 2248 | EVERY1 | 
| 2249 | [rtac equal_intr_rule, Goal.norm_hhf_tac, | |
| 2250 | etac meta_allE, atac, | |
| 2251 | rtac (prop_subst OF [surjective]), | |
| 2252 | REPEAT o etac meta_allE, atac]); | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2253 | val split_meta = timeit_msg "record split_meta proof:" split_meta_prf; | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34151diff
changeset | 2254 | fun split_meta_standardise () = Drule.export_without_context split_meta; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2255 | val split_meta_standard = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2256 | timeit_msg "record split_meta standard:" split_meta_standardise; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2257 | |
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2258 | fun split_object_prf () = | 
| 17261 | 2259 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2260 | val cPI= cterm_of defs_thy (lambda r0 (Trueprop (P $ r0))); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2261 | val _ $ Abs (_, _, P $ _) = fst (Logic.dest_equals (concl_of split_meta_standard)); | 
| 18858 | 2262 | val cP = cterm_of defs_thy P; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2263 | val split_meta' = cterm_instantiate [(cP, cPI)] split_meta_standard; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2264 | val (l, r) = HOLogic.dest_eq (HOLogic.dest_Trueprop split_object_prop); | 
| 18858 | 2265 | val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l); | 
| 2266 | val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2267 | val thl = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2268 | assume cl (*All r. P r*) (* 1 *) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2269 | |> obj_to_meta_all (*!!r. P r*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2270 | |> equal_elim split_meta' (*!!n m more. P (ext n m more)*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2271 | |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2272 | |> implies_intr cl (* 1 ==> 2 *) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2273 | val thr = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2274 | assume cr (*All n m more. P (ext n m more)*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2275 | |> obj_to_meta_all (*!!n m more. P (ext n m more)*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2276 | |> equal_elim (symmetric split_meta') (*!!r. P r*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2277 | |> meta_to_obj_all (*All r. P r*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2278 | |> implies_intr cr (* 2 ==> 1 *) | 
| 33691 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2279 | in thr COMP (thl COMP iffI) end; | 
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2280 | |
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2281 | |
| 
3db22a5707f3
clarified quick-and-dirty usage in record package;
 schirmer <schirmer@in.tum.de> parents: 
33612diff
changeset | 2282 | val split_object_prf = future_forward_prf split_object_prf split_object_prop; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2283 | 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 | 2284 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2285 | |
| 17261 | 2286 | fun split_ex_prf () = | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2287 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2288 | val ss = HOL_basic_ss addsimps [not_ex RS sym, Not_eq_iff]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2289 | val P_nm = fst (dest_Free P); | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 2290 | val not_P = cterm_of defs_thy (lambda r0 (HOLogic.mk_not (P $ r0))); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2291 | val so' = named_cterm_instantiate ([(P_nm, not_P)]) split_object; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2292 | val so'' = simplify ss so'; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2293 | in | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 2294 | prove_standard [] split_ex_prop (fn _ => resolve_tac [so''] 1) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2295 | end; | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2296 | 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 | 2297 | |
| 17261 | 2298 | fun equality_tac thms = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2299 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2300 | val s' :: s :: eqs = rev thms; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2301 | val ss' = ss addsimps (s' :: s :: sel_convs_standard); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2302 | val eqs' = map (simplify ss') eqs; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2303 | in simp_tac (HOL_basic_ss addsimps (s' :: s :: eqs')) 1 end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2304 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2305 | fun equality_prf () = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2306 |       prove_standard [] equality_prop (fn {context, ...} =>
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2307 | fn st => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2308 | let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2309 | st |> (res_inst_tac context [((rN, 0), s)] cases_scheme 1 THEN | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2310 | res_inst_tac context [((rN, 0), s')] cases_scheme 1 THEN | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2311 |               Subgoal.FOCUS (fn {prems, ...} => equality_tac prems) context 1)
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2312 | (*simp_all_tac ss (sel_convs) would also work but is less efficient*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2313 | end); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2314 | 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 | 2315 | |
| 32335 | 2316 | val ((([sel_convs', upd_convs', sel_defs', upd_defs', | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 2317 | fold_congs', unfold_congs', | 
| 35138 | 2318 | splits' as [split_meta', split_object', split_ex'], derived_defs'], | 
| 32335 | 2319 | [surjective', equality']), | 
| 2320 | [induct_scheme', induct', cases_scheme', cases']), thms_thy) = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2321 | defs_thy | 
| 29579 | 2322 | |> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name)) | 
| 15215 | 2323 |          [("select_convs", sel_convs_standard),
 | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 2324 |           ("update_convs", upd_convs_standard),
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2325 |           ("select_defs", sel_defs),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2326 |           ("update_defs", upd_defs),
 | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 2327 |           ("fold_congs", fold_congs),
 | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 2328 |           ("unfold_congs", unfold_congs),
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2329 |           ("splits", [split_meta_standard, split_object, split_ex]),
 | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2330 |           ("defs", derived_defs)]
 | 
| 29579 | 2331 | ||>> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name)) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2332 |           [("surjective", surjective),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2333 |            ("equality", equality)]
 | 
| 29579 | 2334 | ||>> (PureThy.add_thms o (map o apfst o apfst) Binding.name) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2335 |         [(("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 | 2336 |          (("induct", induct), induct_type_global name),
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2337 |          (("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 | 2338 |          (("cases", cases), cases_type_global name)];
 | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2339 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2340 | val sel_upd_simps = sel_convs' @ upd_convs'; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2341 | val sel_upd_defs = sel_defs' @ upd_defs'; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2342 | val iffs = [ext_inject] | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2343 | val depth = parent_len + 1; | 
| 35138 | 2344 | |
| 2345 | val ([simps', iffs'], thms_thy') = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2346 | thms_thy | 
| 35138 | 2347 | |> PureThy.add_thmss | 
| 33053 
dabf9d1bb552
removed "nitpick_const_simp" attribute from Record's "simps";
 blanchet parents: 
32952diff
changeset | 2348 | [((Binding.name "simps", sel_upd_simps), [Simplifier.simp_add]), | 
| 35138 | 2349 | ((Binding.name "iffs", iffs), [iff_add])]; | 
| 2350 | ||
| 2351 | val info = | |
| 2352 | make_record_info args parent fields extension | |
| 2353 | ext_induct ext_inject ext_surjective ext_split ext_def | |
| 2354 | sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs' | |
| 2355 | surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs'; | |
| 2356 | ||
| 2357 | val final_thy = | |
| 2358 | thms_thy' | |
| 2359 | |> put_record name info | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 2360 | |> put_sel_upd names full_moreN depth sel_upd_simps sel_upd_defs (fold_congs', unfold_congs') | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2361 | |> add_record_equalities extension_id equality' | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2362 | |> add_extinjects ext_inject | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2363 | |> add_extsplit extension_name ext_split | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2364 | |> add_record_splits extension_id (split_meta', split_object', split_ex', induct_scheme') | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2365 | |> add_extfields extension_name (fields @ [(full_moreN, moreT)]) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2366 | |> add_fieldext (extension_name, snd extension) (names @ [full_moreN]) | 
| 35239 | 2367 | |> Sign.restore_naming thy; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2368 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2369 | in final_thy end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2370 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2371 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2372 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2373 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2374 | (*We do all preparations and error checks here, deferring the real | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2375 | work to record_definition.*) | 
| 35136 | 2376 | fun gen_add_record prep_typ prep_raw_parent quiet_mode | 
| 35239 | 2377 | (params, binding) raw_parent raw_fields thy = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2378 | let | 
| 17261 | 2379 | val _ = Theory.requires thy "Record" "record definitions"; | 
| 35136 | 2380 | val _ = | 
| 2381 | if quiet_mode then () | |
| 35239 | 2382 |       else writeln ("Defining record " ^ quote (Binding.str_of binding) ^ " ...");
 | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2383 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2384 | val ctxt = ProofContext.init thy; | 
| 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2385 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2386 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2387 | (* parents *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2388 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2389 | fun prep_inst T = fst (cert_typ ctxt T []); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2390 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2391 | val parent = Option.map (apfst (map prep_inst) o prep_raw_parent ctxt) raw_parent | 
| 18678 | 2392 |       handle ERROR msg => cat_error msg ("The error(s) above in parent record specification");
 | 
| 12247 | 2393 | val parents = add_parents thy parent []; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2394 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2395 | val init_env = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2396 | (case parent of | 
| 15531 | 2397 | NONE => [] | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2398 | | SOME (types, _) => fold Term.add_tfreesT types []); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2399 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2400 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2401 | (* fields *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2402 | |
| 35136 | 2403 | fun prep_field (x, raw_T, mx) env = | 
| 2404 | let | |
| 2405 | val (T, env') = | |
| 2406 | prep_typ ctxt raw_T env handle ERROR msg => | |
| 2407 |             cat_error msg ("The error(s) above occured in record field " ^ quote (Binding.str_of x));
 | |
| 2408 | in ((x, T, mx), env') end; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2409 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2410 | val (bfields, envir) = fold_map prep_field raw_fields init_env; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2411 | val envir_names = map fst envir; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2412 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2413 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2414 | (* args *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2415 | |
| 22578 | 2416 | val defaultS = Sign.defaultS thy; | 
| 17485 | 2417 | 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 | 2418 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2419 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2420 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2421 | |
| 35239 | 2422 | val name = Sign.full_name thy binding; | 
| 17261 | 2423 | val err_dup_record = | 
| 4890 | 2424 | if is_none (get_record thy name) then [] | 
| 2425 | else ["Duplicate definition of record " ^ quote name]; | |
| 2426 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2427 | val err_dup_parms = | 
| 18964 | 2428 | (case duplicates (op =) params of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2429 | [] => [] | 
| 4890 | 2430 | | dups => ["Duplicate parameter(s) " ^ commas dups]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2431 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2432 | val err_extra_frees = | 
| 20951 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 haftmann parents: 
20484diff
changeset | 2433 | (case subtract (op =) params envir_names of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2434 | [] => [] | 
| 4890 | 2435 | | extras => ["Extra free type variable(s) " ^ commas extras]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2436 | |
| 4890 | 2437 | 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 | 2438 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2439 | val err_dup_fields = | 
| 35136 | 2440 | (case duplicates Binding.eq_name (map #1 bfields) of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2441 | [] => [] | 
| 35136 | 2442 | | dups => ["Duplicate field(s) " ^ commas_quote (map Binding.str_of dups)]); | 
| 4890 | 2443 | |
| 2444 | val err_bad_fields = | |
| 35136 | 2445 | if forall (not_equal moreN o Binding.name_of o #1) bfields then [] | 
| 4890 | 2446 | else ["Illegal field name " ^ quote moreN]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2447 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2448 | val err_dup_sorts = | 
| 18964 | 2449 | (case duplicates (op =) envir_names of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2450 | [] => [] | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2451 | | dups => ["Inconsistent sort constraints for " ^ commas dups]); | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2452 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2453 | val errs = | 
| 4890 | 2454 | err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @ | 
| 2455 | err_dup_fields @ err_bad_fields @ err_dup_sorts; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2456 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2457 | val _ = if null errs then () else error (cat_lines errs); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2458 | in | 
| 35239 | 2459 | thy |> record_definition (args, binding) parent parents bfields | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2460 | end | 
| 35239 | 2461 |   handle ERROR msg => cat_error msg ("Failed to define record " ^ quote (Binding.str_of binding));
 | 
| 35136 | 2462 | |
| 2463 | val add_record = gen_add_record cert_typ (K I); | |
| 2464 | val add_record_cmd = gen_add_record read_typ read_raw_parent; | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2465 | |
| 32335 | 2466 | |
| 6358 | 2467 | (* setup theory *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2468 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2469 | val setup = | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 2470 | Sign.add_trfuns ([], parse_translation, [], []) #> | 
| 35146 | 2471 | Sign.add_advanced_trfuns ([], advanced_parse_translation, [], []) #> | 
| 26496 
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26477diff
changeset | 2472 | Simplifier.map_simpset (fn ss => | 
| 
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26477diff
changeset | 2473 | ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2474 | |
| 32335 | 2475 | |
| 6358 | 2476 | (* outer syntax *) | 
| 2477 | ||
| 17057 | 2478 | local structure P = OuterParse and K = OuterKeyword in | 
| 6358 | 2479 | |
| 24867 | 2480 | val _ = | 
| 17261 | 2481 | OuterSyntax.command "record" "define extensible record" K.thy_decl | 
| 35136 | 2482 | (P.type_args -- P.binding -- | 
| 2483 | (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const_binding) | |
| 2484 | >> (fn (x, (y, z)) => Toplevel.theory (add_record_cmd false x y z))); | |
| 6358 | 2485 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2486 | end; | 
| 5698 | 2487 | |
| 6384 | 2488 | end; | 
| 2489 | ||
| 32335 | 2490 | structure Basic_Record: BASIC_RECORD = Record; | 
| 2491 | open Basic_Record; |