| author | desharna | 
| Tue, 28 Sep 2021 10:47:18 +0200 | |
| changeset 74370 | d8dc8fdc46fc | 
| parent 74292 | 39c98371606f | 
| child 74383 | 107941e8fa01 | 
| 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 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 10 | signature RECORD = | 
| 5698 | 11 | sig | 
| 43682 | 12 | val type_abbr: bool Config.T | 
| 13 | val type_as_fields: bool Config.T | |
| 14 | val timing: bool Config.T | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 15 | |
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 16 | type info = | 
| 35138 | 17 |    {args: (string * sort) list,
 | 
| 18 | parent: (typ list * string) option, | |
| 19 | fields: (string * typ) list, | |
| 20 | extension: (string * typ list), | |
| 21 | ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm, | |
| 22 | select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list, | |
| 23 | fold_congs: thm list, unfold_congs: thm list, splits: thm list, defs: thm list, | |
| 24 | surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm, | |
| 25 | cases: thm, simps: thm list, iffs: thm list} | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 26 | val get_info: theory -> string -> info option | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 27 | val the_info: theory -> string -> info | 
| 41591 | 28 | val get_hierarchy: theory -> (string * typ list) -> (string * ((string * sort) * typ) list) list | 
| 61260 | 29 |   val add_record: {overloaded: bool} -> (string * sort) list * binding ->
 | 
| 30 | (typ list * string) option -> (binding * typ * mixfix) list -> theory -> theory | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 31 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 32 | val last_extT: typ -> (string * typ list) option | 
| 32972 | 33 | val dest_recTs: typ -> (string * typ list) list | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 34 | val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 35 | val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ) | 
| 26088 | 36 | val get_parent: theory -> string -> (typ list * string) option | 
| 37 | val get_extension: theory -> string -> (string * typ list) option | |
| 16458 | 38 | val get_extinjects: theory -> thm list | 
| 39 | val get_simpset: theory -> simpset | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 40 | val simproc: simproc | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 41 | val eq_simproc: simproc | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 42 | val upd_simproc: simproc | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 43 | val split_simproc: (term -> int) -> simproc | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 44 | val ex_sel_eq_simproc: simproc | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 45 | val split_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 46 | val split_simp_tac: Proof.context -> thm list -> (term -> int) -> int -> tactic | 
| 42793 | 47 | val split_wrapper: string * (Proof.context -> wrapper) | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 48 | |
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 49 | val pretty_recT: Proof.context -> typ -> Pretty.T | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 50 | val string_of_record: Proof.context -> string -> string | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 51 | |
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 52 | val codegen: bool Config.T | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 53 | val updateN: string | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 54 | val ext_typeN: string | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 55 | val extN: string | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 56 | end; | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 57 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 58 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 59 | 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 | 60 | sig | 
| 61260 | 61 |   val add_iso_tuple_type: {overloaded: bool} -> binding * (string * sort) list ->
 | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 62 | typ * typ -> theory -> (term * term) * theory | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 63 | val mk_cons_tuple: term * term -> term | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 64 | val dest_cons_tuple: term -> term * term | 
| 59164 | 65 | val iso_tuple_intros_tac: Proof.context -> int -> tactic | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 66 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 67 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 68 | 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 | 69 | struct | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 70 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 71 | val isoN = "_Tuple_Iso"; | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 72 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 73 | val iso_tuple_intro = @{thm isomorphic_tuple_intro};
 | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 74 | val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros};
 | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 75 | |
| 67149 | 76 | val tuple_iso_tuple = (\<^const_name>\<open>Record.tuple_iso_tuple\<close>, @{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 | 77 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 78 | 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 | 79 | ( | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 80 | type T = thm Symtab.table; | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 81 | 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 | 82 | val extend = I; | 
| 33522 | 83 | 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 | 84 | ); | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 85 | |
| 41575 | 86 | fun get_typedef_info tyco vs | 
| 87 |     (({rep_type, Abs_name, ...}, {Rep_inject, Abs_inverse, ... }) : Typedef.info) thy =
 | |
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 88 | let | 
| 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 89 | val exists_thm = | 
| 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 90 | UNIV_I | 
| 60801 | 91 | |> Thm.instantiate' [SOME (Thm.global_ctyp_of thy (Logic.varifyT_global rep_type))] []; | 
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 92 | val proj_constr = Abs_inverse OF [exists_thm]; | 
| 38534 | 93 | val absT = Type (tyco, map TFree vs); | 
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 94 | in | 
| 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 95 | thy | 
| 38534 | 96 | |> pair (tyco, ((Rep_inject, proj_constr), Const (Abs_name, rep_type --> absT), absT)) | 
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 97 | end | 
| 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 98 | |
| 61260 | 99 | fun do_typedef overloaded raw_tyco repT raw_vs thy = | 
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 100 | let | 
| 42361 | 101 | val ctxt = Proof_Context.init_global thy |> Variable.declare_typ repT; | 
| 102 | val vs = map (Proof_Context.check_tfree ctxt) raw_vs; | |
| 38529 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 103 | in | 
| 
4cc2ca4d6237
formally integrated typecopy layer into record package
 haftmann parents: 
38401diff
changeset | 104 | thy | 
| 69829 | 105 | |> Named_Target.theory_map_result (apsnd o Typedef.transform_info) | 
| 106 | (Typedef.add_typedef overloaded (raw_tyco, vs, NoSyn) | |
| 107 | (HOLogic.mk_UNIV repT) NONE (fn ctxt' => resolve_tac ctxt' [UNIV_witness] 1)) | |
| 38534 | 108 | |-> (fn (tyco, info) => get_typedef_info tyco vs info) | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 109 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 110 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 111 | fun mk_cons_tuple (left, right) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 112 | let | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 113 | val (leftT, rightT) = (fastype_of left, fastype_of right); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 114 | val prodT = HOLogic.mk_prodT (leftT, rightT); | 
| 67149 | 115 | val isomT = Type (\<^type_name>\<open>tuple_isomorphism\<close>, [prodT, leftT, rightT]); | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 116 | in | 
| 67149 | 117 | Const (\<^const_name>\<open>Record.iso_tuple_cons\<close>, isomT --> leftT --> rightT --> prodT) $ | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 118 | 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 | 119 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 120 | |
| 67149 | 121 | fun dest_cons_tuple (Const (\<^const_name>\<open>Record.iso_tuple_cons\<close>, _) $ Const _ $ t $ u) = (t, u) | 
| 32972 | 122 |   | 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 | 123 | |
| 61260 | 124 | fun add_iso_tuple_type overloaded (b, alphas) (leftT, rightT) thy = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 125 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 126 | val repT = HOLogic.mk_prodT (leftT, rightT); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 127 | |
| 38534 | 128 | val ((_, ((rep_inject, abs_inverse), absC, absT)), typ_thy) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 129 | thy | 
| 61260 | 130 | |> do_typedef overloaded b repT alphas | 
| 37470 | 131 | ||> Sign.add_path (Binding.name_of b); (*FIXME proper prefixing instead*) | 
| 60796 | 132 | val typ_ctxt = Proof_Context.init_global typ_thy; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 133 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 134 | (*construct a type and body for the isomorphism constant by | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 135 | instantiating the theorem to which the definition will be applied*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 136 | val intro_inst = | 
| 60796 | 137 | rep_inject RS | 
| 138 | infer_instantiate typ_ctxt | |
| 139 |           [(("abst", 0), Thm.cterm_of typ_ctxt absC)] iso_tuple_intro;
 | |
| 59582 | 140 | val (_, body) = Logic.dest_equals (List.last (Thm.prems_of intro_inst)); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 141 | val isomT = fastype_of body; | 
| 37470 | 142 | val isom_binding = Binding.suffix_name isoN b; | 
| 35239 | 143 | val isom_name = Sign.full_name typ_thy isom_binding; | 
| 32972 | 144 | val isom = Const (isom_name, isomT); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 145 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 146 | val ([isom_def], cdef_thy) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 147 | typ_thy | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 148 | |> Sign.declare_const_global ((isom_binding, isomT), NoSyn) |> snd | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39134diff
changeset | 149 | |> Global_Theory.add_defs false | 
| 59859 | 150 | [((Binding.concealed (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])]; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 151 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 152 | val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro)); | 
| 67149 | 153 | val cons = Const (\<^const_name>\<open>Record.iso_tuple_cons\<close>, isomT --> leftT --> rightT --> absT); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 154 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 155 | val thm_thy = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 156 | cdef_thy | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 157 | |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple)) | 
| 35239 | 158 | |> Sign.restore_naming thy | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 159 | in | 
| 32972 | 160 | ((isom, cons $ isom), thm_thy) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 161 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 162 | |
| 59164 | 163 | fun iso_tuple_intros_tac ctxt = | 
| 164 | resolve_from_net_tac ctxt iso_tuple_intros THEN' | |
| 32975 | 165 | CSUBGOAL (fn (cgoal, i) => | 
| 166 | let | |
| 167 | val goal = Thm.term_of cgoal; | |
| 168 | ||
| 60326 | 169 | val isthms = Iso_Tuple_Thms.get (Proof_Context.theory_of ctxt); | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 170 |       fun err s t = raise TERM ("iso_tuple_intros_tac: " ^ s, [t]);
 | 
| 32975 | 171 | |
| 172 | val goal' = Envir.beta_eta_contract goal; | |
| 173 | val is = | |
| 174 | (case goal' of | |
| 67149 | 175 | Const (\<^const_name>\<open>Trueprop\<close>, _) $ | 
| 176 | (Const (\<^const_name>\<open>isomorphic_tuple\<close>, _) $ Const is) => is | |
| 32975 | 177 | | _ => err "unexpected goal format" goal'); | 
| 178 | val isthm = | |
| 179 | (case Symtab.lookup isthms (#1 is) of | |
| 180 | SOME isthm => isthm | |
| 181 | | NONE => err "no thm found for constant" (Const is)); | |
| 60752 | 182 | in resolve_tac ctxt [isthm] i end); | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 183 | |
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 184 | end; | 
| 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 185 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 186 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 187 | structure Record: RECORD = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 188 | struct | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 189 | |
| 35133 | 190 | val surject_assistI = @{thm iso_tuple_surjective_proof_assistI};
 | 
| 191 | val surject_assist_idE = @{thm iso_tuple_surjective_proof_assist_idE};
 | |
| 192 | ||
| 193 | val updacc_accessor_eqE = @{thm update_accessor_accessor_eqE};
 | |
| 194 | val updacc_updator_eqE = @{thm update_accessor_updator_eqE};
 | |
| 195 | val updacc_eq_idI = @{thm iso_tuple_update_accessor_eq_assist_idI};
 | |
| 196 | val updacc_eq_triv = @{thm iso_tuple_update_accessor_eq_assist_triv};
 | |
| 197 | ||
| 198 | val updacc_foldE = @{thm update_accessor_congruence_foldE};
 | |
| 199 | val updacc_unfoldE = @{thm update_accessor_congruence_unfoldE};
 | |
| 200 | val updacc_noopE = @{thm update_accessor_noopE};
 | |
| 201 | val updacc_noop_compE = @{thm update_accessor_noop_compE};
 | |
| 202 | val updacc_cong_idI = @{thm update_accessor_cong_assist_idI};
 | |
| 203 | val updacc_cong_triv = @{thm update_accessor_cong_assist_triv};
 | |
| 204 | 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 | 205 | |
| 67149 | 206 | val codegen = Attrib.setup_config_bool \<^binding>\<open>record_codegen\<close> (K true); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 207 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 208 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 209 | (** name components **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 210 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 211 | val rN = "r"; | 
| 15215 | 212 | val wN = "w"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 213 | val moreN = "more"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 214 | val schemeN = "_scheme"; | 
| 38531 
a11a1e4e0403
authentic syntax allows simplification of type names
 haftmann parents: 
38530diff
changeset | 215 | val ext_typeN = "_ext"; | 
| 
a11a1e4e0403
authentic syntax allows simplification of type names
 haftmann parents: 
38530diff
changeset | 216 | val inner_typeN = "_inner"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 217 | val extN ="_ext"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 218 | val updateN = "_update"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 219 | val makeN = "make"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 220 | val fields_selN = "fields"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 221 | val extendN = "extend"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 222 | val truncateN = "truncate"; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 223 | |
| 32335 | 224 | |
| 225 | ||
| 4894 | 226 | (*** utilities ***) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 227 | |
| 45424 
01d75cf04497
localized Record.decode_type: use standard Proof_Context.get_sort;
 wenzelm parents: 
44653diff
changeset | 228 | fun varifyT idx = map_type_tfree (fn (a, S) => TVar ((a, idx), S)); | 
| 19748 | 229 | |
| 32335 | 230 | |
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 231 | (* timing *) | 
| 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 232 | |
| 67149 | 233 | val timing = Attrib.setup_config_bool \<^binding>\<open>record_timing\<close> (K false); | 
| 43682 | 234 | fun timeit_msg ctxt s x = if Config.get ctxt timing then (warning s; timeit x) else x (); | 
| 235 | fun timing_msg ctxt s = if Config.get ctxt timing then warning s else (); | |
| 17261 | 236 | |
| 32335 | 237 | |
| 12255 | 238 | (* syntax *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 239 | |
| 11927 
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
 wenzelm parents: 
11923diff
changeset | 240 | val Trueprop = HOLogic.mk_Trueprop; | 
| 4894 | 241 | |
| 11934 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 242 | infix 0 :== ===; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 243 | infixr 0 ==>; | 
| 
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
 wenzelm parents: 
11927diff
changeset | 244 | |
| 37781 
2fbbf0a48cef
moved misc legacy stuff from OldGoals to Misc_Legacy;
 wenzelm parents: 
37470diff
changeset | 245 | val op :== = Misc_Legacy.mk_defpair; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 246 | val op === = Trueprop o HOLogic.mk_eq; | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 247 | val op ==> = Logic.mk_implies; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 248 | |
| 32335 | 249 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 250 | (* constructor *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 251 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 252 | fun mk_ext (name, T) ts = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 253 | let val Ts = map fastype_of ts | 
| 35239 | 254 | 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 | 255 | |
| 32335 | 256 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 257 | (* selector *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 258 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 259 | fun mk_selC sT (c, T) = (c, sT --> T); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 260 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 261 | fun mk_sel s (c, T) = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 262 | let val sT = fastype_of s | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 263 | in Const (mk_selC sT (c, T)) $ s end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 264 | |
| 32335 | 265 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 266 | (* updates *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 267 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 268 | 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 | 269 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 270 | 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 | 271 | let val vT = domain_type (fastype_of v); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 272 | in Const (mk_updC sfx sT (c, vT)) $ v end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 273 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 274 | 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 | 275 | |
| 32335 | 276 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 277 | (* types *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 278 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 279 | 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 | 280 | (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 | 281 |         NONE => raise TYPE ("Record.dest_recT", [typ], [])
 | 
| 15570 | 282 | | SOME c => ((c, Ts), List.last Ts)) | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 283 |   | dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []);
 | 
| 5197 | 284 | |
| 32975 | 285 | val is_recT = can dest_recT; | 
| 11833 | 286 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 287 | fun dest_recTs T = | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 288 | let val ((c, Ts), U) = dest_recT T | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 289 | in (c, Ts) :: dest_recTs U | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 290 | end handle TYPE _ => []; | 
| 14255 | 291 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 292 | fun last_extT T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 293 | let val ((c, Ts), U) = dest_recT T in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 294 | (case last_extT U of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 295 | NONE => SOME (c, Ts) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 296 | | SOME l => SOME l) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 297 | end handle TYPE _ => NONE; | 
| 14255 | 298 | |
| 17261 | 299 | fun rec_id i T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 300 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 301 | val rTs = dest_recTs T; | 
| 33957 | 302 | 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 | 303 | in implode (map #1 rTs') end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 304 | |
| 32335 | 305 | |
| 306 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 307 | (*** extend theory by record definition ***) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 308 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 309 | (** record info **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 310 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 311 | (* type info and parent_info *) | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 312 | |
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 313 | type info = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 314 |  {args: (string * sort) list,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 315 | parent: (typ list * string) option, | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 316 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 317 | extension: (string * typ list), | 
| 35138 | 318 | |
| 319 | ext_induct: thm, | |
| 320 | ext_inject: thm, | |
| 321 | ext_surjective: thm, | |
| 322 | ext_split: thm, | |
| 323 | ext_def: thm, | |
| 324 | ||
| 325 | select_convs: thm list, | |
| 326 | update_convs: thm list, | |
| 327 | select_defs: thm list, | |
| 328 | update_defs: thm list, | |
| 47783 | 329 | fold_congs: thm list, (* potentially used in L4.verified *) | 
| 330 | unfold_congs: thm list, (* potentially used in L4.verified *) | |
| 35138 | 331 | splits: thm list, | 
| 332 | defs: thm list, | |
| 333 | ||
| 334 | surjective: thm, | |
| 335 | equality: thm, | |
| 336 | induct_scheme: thm, | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 337 | induct: thm, | 
| 35138 | 338 | cases_scheme: thm, | 
| 339 | cases: thm, | |
| 340 | ||
| 341 | simps: thm list, | |
| 342 | iffs: thm list}; | |
| 343 | ||
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 344 | fun make_info args parent fields extension | 
| 35138 | 345 | ext_induct ext_inject ext_surjective ext_split ext_def | 
| 346 | select_convs update_convs select_defs update_defs fold_congs unfold_congs splits defs | |
| 347 | surjective equality induct_scheme induct cases_scheme cases | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 348 | simps iffs : info = | 
| 17261 | 349 |  {args = args, parent = parent, fields = fields, extension = extension,
 | 
| 35138 | 350 | ext_induct = ext_induct, ext_inject = ext_inject, ext_surjective = ext_surjective, | 
| 351 | ext_split = ext_split, ext_def = ext_def, select_convs = select_convs, | |
| 352 | update_convs = update_convs, select_defs = select_defs, update_defs = update_defs, | |
| 353 | fold_congs = fold_congs, unfold_congs = unfold_congs, splits = splits, defs = defs, | |
| 354 | surjective = surjective, equality = equality, induct_scheme = induct_scheme, | |
| 355 | 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 | 356 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 357 | type parent_info = | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 358 |  {name: string,
 | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 359 | fields: (string * typ) list, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 360 | extension: (string * typ list), | 
| 35138 | 361 | induct_scheme: thm, | 
| 362 | ext_def: thm}; | |
| 363 | ||
| 364 | 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 | 365 |  {name = name, fields = fields, extension = extension,
 | 
| 35138 | 366 | ext_def = ext_def, induct_scheme = induct_scheme}; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 367 | |
| 22846 | 368 | |
| 369 | (* theory data *) | |
| 5001 | 370 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 371 | type data = | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 372 |  {records: info Symtab.table,
 | 
| 7178 | 373 | sel_upd: | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 374 |    {selectors: (int * bool) Symtab.table,
 | 
| 7178 | 375 | updates: string Symtab.table, | 
| 46047 
6170af176fbb
tuned -- afford slightly larger simpset in simp_defs_tac;
 wenzelm parents: 
46046diff
changeset | 376 | simpset: simpset, | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 377 | defset: simpset}, | 
| 14255 | 378 | equalities: thm Symtab.table, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 379 | extinjects: thm list, | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 380 | extsplit: thm Symtab.table, (*maps extension name to split rule*) | 
| 35135 | 381 | 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 | 382 | 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 | 383 | fieldext: (string * typ list) Symtab.table}; (*maps field to its extension*) | 
| 7178 | 384 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 385 | fun make_data | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 386 | records sel_upd equalities extinjects extsplit splits extfields fieldext = | 
| 17261 | 387 |  {records = records, sel_upd = sel_upd,
 | 
| 388 | equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits, | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 389 | extfields = extfields, fieldext = fieldext }: data; | 
| 7178 | 390 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 391 | structure Data = Theory_Data | 
| 22846 | 392 | ( | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 393 | type T = data; | 
| 7178 | 394 | val empty = | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 395 | make_data Symtab.empty | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 396 |       {selectors = Symtab.empty, updates = Symtab.empty,
 | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 397 | simpset = HOL_basic_ss, defset = HOL_basic_ss} | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 398 | Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; | 
| 16458 | 399 | val extend = I; | 
| 33522 | 400 | fun merge | 
| 7178 | 401 |    ({records = recs1,
 | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 402 |      sel_upd = {selectors = sels1, updates = upds1, simpset = ss1, defset = ds1},
 | 
| 14255 | 403 | equalities = equalities1, | 
| 17261 | 404 | extinjects = extinjects1, | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 405 | extsplit = extsplit1, | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 406 | splits = splits1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 407 | extfields = extfields1, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 408 | fieldext = fieldext1}, | 
| 7178 | 409 |     {records = recs2,
 | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 410 |      sel_upd = {selectors = sels2, updates = upds2, simpset = ss2, defset = ds2},
 | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 411 | equalities = equalities2, | 
| 17261 | 412 | extinjects = extinjects2, | 
| 413 | extsplit = extsplit2, | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 414 | splits = splits2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 415 | extfields = extfields2, | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 416 | fieldext = fieldext2}) = | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 417 | make_data | 
| 7178 | 418 | (Symtab.merge (K true) (recs1, recs2)) | 
| 419 |       {selectors = Symtab.merge (K true) (sels1, sels2),
 | |
| 420 | 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 | 421 | simpset = Simplifier.merge_ss (ss1, ss2), | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 422 | defset = Simplifier.merge_ss (ds1, ds2)} | 
| 22634 | 423 | (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2)) | 
| 33522 | 424 | (Thm.merge_thms (extinjects1, extinjects2)) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 425 | (Symtab.merge Thm.eq_thm_prop (extsplit1, extsplit2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 426 | (Symtab.merge (fn ((a, b, c, d), (w, x, y, z)) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 427 | Thm.eq_thm (a, w) andalso Thm.eq_thm (b, x) andalso | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 428 | Thm.eq_thm (c, y) andalso Thm.eq_thm (d, z)) (splits1, splits2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 429 | (Symtab.merge (K true) (extfields1, extfields2)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 430 | (Symtab.merge (K true) (fieldext1, fieldext2)); | 
| 22846 | 431 | ); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 432 | |
| 16458 | 433 | |
| 7178 | 434 | (* access 'records' *) | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 435 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 436 | val get_info = Symtab.lookup o #records o Data.get; | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 437 | |
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 438 | fun the_info thy name = | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 439 | (case get_info thy name of | 
| 35138 | 440 | SOME info => info | 
| 441 |   | NONE => error ("Unknown record type " ^ quote name));
 | |
| 442 | ||
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 443 | fun put_record name info = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 444 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 445 | make_data (Symtab.update (name, info) records) | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 446 | sel_upd equalities extinjects extsplit splits extfields fieldext); | 
| 7178 | 447 | |
| 22846 | 448 | |
| 7178 | 449 | (* access 'sel_upd' *) | 
| 450 | ||
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 451 | val get_sel_upd = #sel_upd o Data.get; | 
| 7178 | 452 | |
| 17510 | 453 | val is_selector = Symtab.defined o #selectors o get_sel_upd; | 
| 17412 | 454 | val get_updates = Symtab.lookup o #updates o get_sel_upd; | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 455 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 456 | val get_simpset = #simpset o get_sel_upd; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 457 | val get_sel_upd_defs = #defset o get_sel_upd; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 458 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 459 | fun get_update_details u thy = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 460 | let val sel_upd = get_sel_upd thy in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 461 | (case Symtab.lookup (#updates sel_upd) u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 462 | SOME s => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 463 | let val SOME (dep, ismore) = Symtab.lookup (#selectors sel_upd) s | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 464 | in SOME (s, dep, ismore) end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 465 | | NONE => NONE) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 466 | end; | 
| 7178 | 467 | |
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 468 | fun put_sel_upd names more depth simps defs thy = | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 469 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 470 | val ctxt0 = Proof_Context.init_global thy; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 471 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 472 | val all = names @ [more]; | 
| 32744 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 473 | 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 | 474 | val upds = map (suffix updateN) all ~~ all; | 
| 
50406c4951d9
Record patch imported and working.
 Thomas Sewell <tsewell@nicta.com.au> parents: 
32743diff
changeset | 475 | |
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 476 |     val {records, sel_upd = {selectors, updates, simpset, defset},
 | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 477 | equalities, extinjects, extsplit, splits, extfields, fieldext} = Data.get thy; | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 478 | val data = | 
| 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 479 | make_data records | 
| 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 480 |         {selectors = fold Symtab.update_new sels selectors,
 | 
| 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 481 | updates = fold Symtab.update_new upds updates, | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 482 | simpset = simpset_map ctxt0 (fn ctxt => ctxt addsimps simps) simpset, | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 483 | defset = simpset_map ctxt0 (fn ctxt => ctxt addsimps defs) defset} | 
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 484 | equalities extinjects extsplit splits extfields fieldext; | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 485 | in Data.put data thy end; | 
| 22846 | 486 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 487 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 488 | (* access 'equalities' *) | 
| 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 489 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 490 | fun add_equalities name thm = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 491 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 492 | make_data records sel_upd | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 493 | (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 494 | |
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 495 | val get_equalities = Symtab.lookup o #equalities o Data.get; | 
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 496 | |
| 22846 | 497 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 498 | (* access 'extinjects' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 499 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 500 | fun add_extinjects thm = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 501 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 502 | make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects) | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 503 | extsplit splits extfields fieldext); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 504 | |
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 505 | val get_extinjects = rev o #extinjects o Data.get; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 506 | |
| 22846 | 507 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 508 | (* access 'extsplit' *) | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 509 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 510 | fun add_extsplit name thm = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 511 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 512 | make_data records sel_upd equalities extinjects | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 513 | (Symtab.update_new (name, thm) extsplit) splits extfields fieldext); | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 514 | |
| 26088 | 515 | |
| 14255 | 516 | (* access 'splits' *) | 
| 517 | ||
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 518 | fun add_splits name thmP = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 519 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 520 | make_data records sel_upd equalities extinjects extsplit | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 521 | (Symtab.update_new (name, thmP) splits) extfields fieldext); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 522 | |
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 523 | val get_splits = Symtab.lookup o #splits o Data.get; | 
| 14255 | 524 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 525 | |
| 26088 | 526 | (* parent/extension of named record *) | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 527 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 528 | val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Data.get); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 529 | val get_extension = Option.map #extension oo (Symtab.lookup o #records o Data.get); | 
| 17261 | 530 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 531 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 532 | (* access 'extfields' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 533 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 534 | fun add_extfields name fields = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 535 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 536 | make_data records sel_upd equalities extinjects extsplit splits | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 537 | (Symtab.update_new (name, fields) extfields) fieldext); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 538 | |
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 539 | val get_extfields = Symtab.lookup o #extfields o Data.get; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 540 | |
| 18858 | 541 | fun get_extT_fields thy T = | 
| 15059 | 542 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 543 | val ((name, Ts), moreT) = dest_recT T; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 544 | val recname = | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 545 | let val (nm :: _ :: rst) = rev (Long_Name.explode name) (* FIXME !? *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 546 | in Long_Name.implode (rev (nm :: rst)) end; | 
| 45424 
01d75cf04497
localized Record.decode_type: use standard Proof_Context.get_sort;
 wenzelm parents: 
44653diff
changeset | 547 | val varifyT = varifyT (maxidx_of_typs (moreT :: Ts) + 1); | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 548 |     val {records, extfields, ...} = Data.get thy;
 | 
| 35149 | 549 | val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name); | 
| 17412 | 550 | val args = map varifyT (snd (#extension (the (Symtab.lookup records recname)))); | 
| 15058 | 551 | |
| 41489 
8e2b8649507d
standardized split_last/last_elem towards List.last;
 wenzelm parents: 
40845diff
changeset | 552 | val subst = fold (Sign.typ_match thy) (#1 (split_last args) ~~ #1 (split_last Ts)) Vartab.empty; | 
| 35149 | 553 | val fields' = map (apsnd (Envir.norm_type subst o varifyT)) fields; | 
| 554 | in (fields', (more, moreT)) end; | |
| 15058 | 555 | |
| 18858 | 556 | fun get_recT_fields thy T = | 
| 17261 | 557 | let | 
| 35149 | 558 | val (root_fields, (root_more, root_moreT)) = get_extT_fields thy T; | 
| 559 | val (rest_fields, rest_more) = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 560 | if is_recT root_moreT then get_recT_fields thy root_moreT | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 561 | else ([], (root_more, root_moreT)); | 
| 35149 | 562 | in (root_fields @ rest_fields, rest_more) end; | 
| 15059 | 563 | |
| 15058 | 564 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 565 | (* access 'fieldext' *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 566 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 567 | fun add_fieldext extname_types fields = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 568 |   Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
 | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 569 | let | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 570 | val fieldext' = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 571 | fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext; | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 572 | in make_data records sel_upd equalities extinjects extsplit splits extfields fieldext' end); | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 573 | |
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 574 | val get_fieldext = Symtab.lookup o #fieldext o Data.get; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 575 | |
| 21962 | 576 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 577 | (* parent records *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 578 | |
| 41577 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 579 | local | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 580 | |
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 581 | fun add_parents _ NONE = I | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 582 | | add_parents thy (SOME (types, name)) = | 
| 12247 | 583 | let | 
| 584 | fun err msg = error (msg ^ " parent record " ^ quote name); | |
| 12255 | 585 | |
| 41577 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 586 |         val {args, parent, ...} =
 | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 587 | (case get_info thy name of SOME info => info | NONE => err "Unknown"); | 
| 12247 | 588 | val _ = if length types <> length args then err "Bad number of arguments for" else (); | 
| 12255 | 589 | |
| 12247 | 590 | fun bad_inst ((x, S), T) = | 
| 22578 | 591 | if Sign.of_sort thy (T, S) then NONE else SOME x | 
| 32952 | 592 | val bads = map_filter bad_inst (args ~~ types); | 
| 21962 | 593 |         val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
 | 
| 12255 | 594 | |
| 41591 | 595 | val inst = args ~~ types; | 
| 596 | val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); | |
| 15570 | 597 | val parent' = Option.map (apfst (map subst)) parent; | 
| 41577 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 598 | in cons (name, inst) #> add_parents thy parent' end; | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 599 | |
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 600 | in | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 601 | |
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 602 | fun get_hierarchy thy (name, types) = add_parents thy (SOME (types, name)) []; | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 603 | |
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 604 | fun get_parent_info thy parent = | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 605 | add_parents thy parent [] |> map (fn (name, inst) => | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 606 | let | 
| 41591 | 607 | val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); | 
| 41577 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 608 |       val {fields, extension, induct_scheme, ext_def, ...} = the_info thy name;
 | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 609 | val fields' = map (apsnd subst) fields; | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 610 | val extension' = apsnd (map subst) extension; | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 611 | in make_parent_info name fields' extension' ext_def induct_scheme end); | 
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 612 | |
| 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 613 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 614 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 615 | |
| 21962 | 616 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 617 | (** concrete syntax for records **) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 618 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 619 | (* parse translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 620 | |
| 35144 | 621 | local | 
| 622 | ||
| 43681 | 623 | fun split_args (field :: fields) ((name, arg) :: fargs) = | 
| 624 | if can (unsuffix name) field then | |
| 625 | let val (args, rest) = split_args fields fargs | |
| 626 | in (arg :: args, rest) end | |
| 627 |       else raise Fail ("expecting field " ^ quote field ^ " but got " ^ quote name)
 | |
| 628 | | split_args [] (fargs as (_ :: _)) = ([], fargs) | |
| 629 | | split_args (_ :: _) [] = raise Fail "expecting more fields" | |
| 630 | | split_args _ _ = ([], []); | |
| 631 | ||
| 67149 | 632 | fun field_type_tr ((Const (\<^syntax_const>\<open>_field_type\<close>, _) $ Const (name, _) $ arg)) = | 
| 35146 | 633 | (name, arg) | 
| 634 |   | field_type_tr t = raise TERM ("field_type_tr", [t]);
 | |
| 635 | ||
| 67149 | 636 | fun field_types_tr (Const (\<^syntax_const>\<open>_field_types\<close>, _) $ t $ u) = | 
| 35146 | 637 | field_type_tr t :: field_types_tr u | 
| 638 | | field_types_tr t = [field_type_tr t]; | |
| 639 | ||
| 640 | fun record_field_types_tr more ctxt t = | |
| 17261 | 641 | let | 
| 42361 | 642 | val thy = Proof_Context.theory_of ctxt; | 
| 35146 | 643 |     fun err msg = raise TERM ("Error in record-type input: " ^ msg, [t]);
 | 
| 644 | ||
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 645 | fun mk_ext (fargs as (name, _) :: _) = | 
| 42361 | 646 | (case get_fieldext thy (Proof_Context.intern_const ctxt name) of | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 647 | SOME (ext, alphas) => | 
| 18858 | 648 | (case get_extfields thy ext of | 
| 35146 | 649 | SOME fields => | 
| 650 | let | |
| 41489 
8e2b8649507d
standardized split_last/last_elem towards List.last;
 wenzelm parents: 
40845diff
changeset | 651 | val (fields', _) = split_last fields; | 
| 35146 | 652 | val types = map snd fields'; | 
| 43681 | 653 | val (args, rest) = split_args (map fst fields') fargs | 
| 654 | handle Fail msg => err msg; | |
| 45434 | 655 | val argtypes = Syntax.check_typs ctxt (map Syntax_Phases.decode_typ args); | 
| 656 | val varifyT = varifyT (fold Term.maxidx_typ argtypes ~1 + 1); | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 657 | val vartypes = map varifyT types; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 658 | |
| 36159 
bffb04bf4e83
more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
 wenzelm parents: 
36153diff
changeset | 659 | val subst = Type.raw_matches (vartypes, argtypes) Vartab.empty | 
| 35146 | 660 | handle Type.TYPE_MATCH => err "type is no proper record (extension)"; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 661 | val alphas' = | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 662 | map (Syntax_Phases.term_of_typ ctxt o Envir.norm_type subst o varifyT) | 
| 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 663 | (#1 (split_last alphas)); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 664 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 665 | val more' = mk_ext rest; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 666 | in | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 667 | list_comb | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 668 | (Syntax.const (Lexicon.mark_type (suffix ext_typeN ext)), alphas' @ [more']) | 
| 35146 | 669 | end | 
| 43683 | 670 |               | NONE => err ("no fields defined for " ^ quote ext))
 | 
| 671 | | NONE => err (quote name ^ " is no proper field")) | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 672 | | mk_ext [] = more; | 
| 35146 | 673 | in | 
| 674 | mk_ext (field_types_tr t) | |
| 675 | end; | |
| 676 | ||
| 67149 | 677 | fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const \<^type_syntax>\<open>unit\<close>) ctxt t | 
| 35146 | 678 |   | record_type_tr _ ts = raise TERM ("record_type_tr", ts);
 | 
| 679 | ||
| 680 | fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t | |
| 681 |   | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts);
 | |
| 682 | ||
| 35147 | 683 | |
| 67149 | 684 | fun field_tr ((Const (\<^syntax_const>\<open>_field\<close>, _) $ Const (name, _) $ arg)) = (name, arg) | 
| 35147 | 685 |   | field_tr t = raise TERM ("field_tr", [t]);
 | 
| 686 | ||
| 67149 | 687 | fun fields_tr (Const (\<^syntax_const>\<open>_fields\<close>, _) $ t $ u) = field_tr t :: fields_tr u | 
| 35147 | 688 | | fields_tr t = [field_tr t]; | 
| 689 | ||
| 690 | fun record_fields_tr more ctxt t = | |
| 691 | let | |
| 42361 | 692 | val thy = Proof_Context.theory_of ctxt; | 
| 35147 | 693 |     fun err msg = raise TERM ("Error in record input: " ^ msg, [t]);
 | 
| 694 | ||
| 695 | fun mk_ext (fargs as (name, _) :: _) = | |
| 42361 | 696 | (case get_fieldext thy (Proof_Context.intern_const ctxt name) of | 
| 35147 | 697 | SOME (ext, _) => | 
| 698 | (case get_extfields thy ext of | |
| 699 | SOME fields => | |
| 700 | let | |
| 43681 | 701 | val (args, rest) = split_args (map fst (fst (split_last fields))) fargs | 
| 702 | handle Fail msg => err msg; | |
| 35147 | 703 | val more' = mk_ext rest; | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 704 | in list_comb (Syntax.const (Lexicon.mark_const (ext ^ extN)), args @ [more']) end | 
| 43683 | 705 |               | NONE => err ("no fields defined for " ^ quote ext))
 | 
| 706 | | NONE => err (quote name ^ " is no proper field")) | |
| 35147 | 707 | | mk_ext [] = more; | 
| 708 | in mk_ext (fields_tr t) end; | |
| 709 | ||
| 67149 | 710 | fun record_tr ctxt [t] = record_fields_tr (Syntax.const \<^const_syntax>\<open>Unity\<close>) ctxt t | 
| 35147 | 711 |   | record_tr _ ts = raise TERM ("record_tr", ts);
 | 
| 712 | ||
| 713 | fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t | |
| 714 |   | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts);
 | |
| 715 | ||
| 716 | ||
| 67149 | 717 | fun field_update_tr (Const (\<^syntax_const>\<open>_field_update\<close>, _) $ Const (name, _) $ arg) = | 
| 43683 | 718 | Syntax.const (suffix updateN name) $ Abs (Name.uu_, dummyT, arg) | 
| 35147 | 719 |   | field_update_tr t = raise TERM ("field_update_tr", [t]);
 | 
| 720 | ||
| 67149 | 721 | fun field_updates_tr (Const (\<^syntax_const>\<open>_field_updates\<close>, _) $ t $ u) = | 
| 35147 | 722 | field_update_tr t :: field_updates_tr u | 
| 723 | | field_updates_tr t = [field_update_tr t]; | |
| 724 | ||
| 725 | fun record_update_tr [t, u] = fold (curry op $) (field_updates_tr u) t | |
| 726 |   | record_update_tr ts = raise TERM ("record_update_tr", ts);
 | |
| 727 | ||
| 35146 | 728 | in | 
| 15215 | 729 | |
| 56732 | 730 | val _ = | 
| 731 | Theory.setup | |
| 732 | (Sign.parse_translation | |
| 67149 | 733 | [(\<^syntax_const>\<open>_record_update\<close>, K record_update_tr), | 
| 734 | (\<^syntax_const>\<open>_record\<close>, record_tr), | |
| 735 | (\<^syntax_const>\<open>_record_scheme\<close>, record_scheme_tr), | |
| 736 | (\<^syntax_const>\<open>_record_type\<close>, record_type_tr), | |
| 737 | (\<^syntax_const>\<open>_record_type_scheme\<close>, record_type_scheme_tr)]); | |
| 35146 | 738 | |
| 739 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 740 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
21109diff
changeset | 741 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 742 | (* print translations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 743 | |
| 67149 | 744 | val type_abbr = Attrib.setup_config_bool \<^binding>\<open>record_type_abbr\<close> (K true); | 
| 745 | val type_as_fields = Attrib.setup_config_bool \<^binding>\<open>record_type_as_fields\<close> (K true); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 746 | |
| 35149 | 747 | |
| 748 | local | |
| 749 | ||
| 750 | (* FIXME early extern (!??) *) | |
| 751 | (* FIXME Syntax.free (??) *) | |
| 67149 | 752 | fun field_type_tr' (c, t) = Syntax.const \<^syntax_const>\<open>_field_type\<close> $ Syntax.const c $ t; | 
| 753 | ||
| 754 | fun field_types_tr' (t, u) = Syntax.const \<^syntax_const>\<open>_field_types\<close> $ t $ u; | |
| 35149 | 755 | |
| 756 | fun record_type_tr' ctxt t = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 757 | let | 
| 42361 | 758 | val thy = Proof_Context.theory_of ctxt; | 
| 35149 | 759 | |
| 45427 
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
 wenzelm parents: 
45424diff
changeset | 760 | val T = Syntax_Phases.decode_typ t; | 
| 45424 
01d75cf04497
localized Record.decode_type: use standard Proof_Context.get_sort;
 wenzelm parents: 
44653diff
changeset | 761 | val varifyT = varifyT (Term.maxidx_of_typ T + 1); | 
| 35149 | 762 | |
| 763 | fun strip_fields T = | |
| 764 | (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 | 765 | Type (ext, args as _ :: _) => | 
| 35149 | 766 | (case try (unsuffix ext_typeN) ext of | 
| 767 | SOME ext' => | |
| 768 | (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 | 769 | 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 | 770 | (case get_fieldext thy x of | 
| 35149 | 771 | SOME (_, alphas) => | 
| 772 | (let | |
| 41489 
8e2b8649507d
standardized split_last/last_elem towards List.last;
 wenzelm parents: 
40845diff
changeset | 773 | val (f :: fs, _) = split_last fields; | 
| 35149 | 774 | val fields' = | 
| 42361 | 775 | apfst (Proof_Context.extern_const ctxt) f :: | 
| 42359 | 776 | map (apfst Long_Name.base_name) fs; | 
| 35149 | 777 | val (args', more) = split_last args; | 
| 41489 
8e2b8649507d
standardized split_last/last_elem towards List.last;
 wenzelm parents: 
40845diff
changeset | 778 | val alphavars = map varifyT (#1 (split_last alphas)); | 
| 36159 
bffb04bf4e83
more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
 wenzelm parents: 
36153diff
changeset | 779 | val subst = Type.raw_matches (alphavars, args') Vartab.empty; | 
| 35149 | 780 | val fields'' = (map o apsnd) (Envir.norm_type subst o varifyT) fields'; | 
| 781 | in fields'' @ strip_fields more end | |
| 36159 
bffb04bf4e83
more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
 wenzelm parents: 
36153diff
changeset | 782 |                       handle Type.TYPE_MATCH => [("", 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 | 783 |                   | _ => [("", 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 | 784 |               | _ => [("", 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 | 785 |           | _ => [("", T)])
 | 
| 35149 | 786 |       | _ => [("", T)]);
 | 
| 787 | ||
| 788 | val (fields, (_, moreT)) = split_last (strip_fields T); | |
| 789 | val _ = null fields andalso raise Match; | |
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 790 | val u = | 
| 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 791 | foldr1 field_types_tr' | 
| 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 792 | (map (field_type_tr' o apsnd (Syntax_Phases.term_of_typ ctxt)) fields); | 
| 35149 | 793 | in | 
| 43682 | 794 | if not (Config.get ctxt type_as_fields) orelse null fields then raise Match | 
| 67149 | 795 | else if moreT = HOLogic.unitT then Syntax.const \<^syntax_const>\<open>_record_type\<close> $ u | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 796 | else | 
| 67149 | 797 | Syntax.const \<^syntax_const>\<open>_record_type_scheme\<close> $ u $ | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 798 | Syntax_Phases.term_of_typ ctxt moreT | 
| 35149 | 799 | end; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 800 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 801 | (*try to reconstruct the record name type abbreviation from | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 802 | the (nested) extension types*) | 
| 35149 | 803 | 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 | 804 | let | 
| 45427 
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
 wenzelm parents: 
45424diff
changeset | 805 | val T = Syntax_Phases.decode_typ tm; | 
| 45424 
01d75cf04497
localized Record.decode_type: use standard Proof_Context.get_sort;
 wenzelm parents: 
44653diff
changeset | 806 | val varifyT = varifyT (maxidx_of_typ T + 1); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 807 | |
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 808 | fun mk_type_abbr subst name args = | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 809 | let val abbrT = Type (name, map (varifyT o TFree) args) | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 810 | in Syntax_Phases.term_of_typ ctxt (Envir.norm_type subst abbrT) end; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 811 | |
| 36159 
bffb04bf4e83
more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
 wenzelm parents: 
36153diff
changeset | 812 | fun match rT T = Type.raw_match (varifyT rT, T) Vartab.empty; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 813 | in | 
| 43682 | 814 | if Config.get ctxt type_abbr then | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 815 | (case last_extT T of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 816 | SOME (name, _) => | 
| 35148 | 817 | if name = last_ext then | 
| 35149 | 818 | let val subst = match schemeT T in | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 819 | if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree zeta))) | 
| 32335 | 820 | then mk_type_abbr subst abbr alphas | 
| 821 | else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta]) | |
| 35149 | 822 | end handle Type.TYPE_MATCH => record_type_tr' ctxt tm | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 823 | else raise Match (*give print translation of specialised record a chance*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 824 | | _ => raise Match) | 
| 35149 | 825 | else record_type_tr' ctxt tm | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 826 | end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 827 | |
| 35149 | 828 | in | 
| 829 | ||
| 830 | fun record_ext_type_tr' name = | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 831 | let | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 832 | val ext_type_name = Lexicon.mark_type (suffix ext_typeN name); | 
| 35149 | 833 | fun tr' ctxt ts = | 
| 834 | record_type_tr' ctxt (list_comb (Syntax.const ext_type_name, ts)); | |
| 835 | in (ext_type_name, tr') end; | |
| 836 | ||
| 837 | 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 | 838 | let | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 839 | val ext_type_name = Lexicon.mark_type (suffix ext_typeN name); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 840 | fun tr' ctxt ts = | 
| 35149 | 841 | record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt | 
| 842 | (list_comb (Syntax.const ext_type_name, ts)); | |
| 843 | in (ext_type_name, tr') end; | |
| 844 | ||
| 845 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 846 | |
| 32335 | 847 | |
| 35240 | 848 | local | 
| 849 | ||
| 850 | (* FIXME Syntax.free (??) *) | |
| 67149 | 851 | fun field_tr' (c, t) = Syntax.const \<^syntax_const>\<open>_field\<close> $ Syntax.const c $ t; | 
| 852 | fun fields_tr' (t, u) = Syntax.const \<^syntax_const>\<open>_fields\<close> $ t $ u; | |
| 35240 | 853 | |
| 854 | fun record_tr' ctxt t = | |
| 855 | let | |
| 42361 | 856 | val thy = Proof_Context.theory_of ctxt; | 
| 35240 | 857 | |
| 858 | fun strip_fields t = | |
| 859 | (case strip_comb t of | |
| 860 | (Const (ext, _), args as (_ :: _)) => | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 861 | (case try (Lexicon.unmark_const o unsuffix extN) ext of | 
| 35240 | 862 | SOME ext' => | 
| 863 | (case get_extfields thy ext' of | |
| 864 | SOME fields => | |
| 865 | (let | |
| 41489 
8e2b8649507d
standardized split_last/last_elem towards List.last;
 wenzelm parents: 
40845diff
changeset | 866 | val (f :: fs, _) = split_last (map fst fields); | 
| 42361 | 867 | val fields' = Proof_Context.extern_const ctxt f :: map Long_Name.base_name fs; | 
| 35240 | 868 | val (args', more) = split_last args; | 
| 869 | in (fields' ~~ args') @ strip_fields more end | |
| 40722 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 wenzelm parents: 
40315diff
changeset | 870 |                   handle ListPair.UnequalLengths => [("", t)])
 | 
| 35240 | 871 |               | NONE => [("", t)])
 | 
| 872 |           | NONE => [("", t)])
 | |
| 873 |        | _ => [("", t)]);
 | |
| 874 | ||
| 875 | val (fields, (_, more)) = split_last (strip_fields t); | |
| 876 | val _ = null fields andalso raise Match; | |
| 877 | val u = foldr1 fields_tr' (map field_tr' fields); | |
| 878 | in | |
| 46046 | 879 | (case more of | 
| 67149 | 880 | Const (\<^const_syntax>\<open>Unity\<close>, _) => Syntax.const \<^syntax_const>\<open>_record\<close> $ u | 
| 881 | | _ => Syntax.const \<^syntax_const>\<open>_record_scheme\<close> $ u $ more) | |
| 35240 | 882 | end; | 
| 883 | ||
| 884 | in | |
| 885 | ||
| 886 | fun record_ext_tr' name = | |
| 887 | let | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 888 | val ext_name = Lexicon.mark_const (name ^ extN); | 
| 35240 | 889 | fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts)); | 
| 890 | in (ext_name, tr') end; | |
| 891 | ||
| 892 | end; | |
| 893 | ||
| 894 | ||
| 895 | local | |
| 896 | ||
| 41578 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 897 | fun dest_update ctxt c = | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 898 | (case try Lexicon.unmark_const c of | 
| 42361 | 899 | SOME d => try (unsuffix updateN) (Proof_Context.extern_const ctxt d) | 
| 41578 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 900 | | NONE => NONE); | 
| 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 901 | |
| 35240 | 902 | fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) = | 
| 41578 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 903 | (case dest_update ctxt c of | 
| 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 904 | SOME name => | 
| 42284 | 905 | (case try Syntax_Trans.const_abs_tr' k of | 
| 42086 
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
 wenzelm parents: 
41928diff
changeset | 906 | SOME t => | 
| 67149 | 907 | apfst (cons (Syntax.const \<^syntax_const>\<open>_field_update\<close> $ Syntax.free name $ t)) | 
| 42086 
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
 wenzelm parents: 
41928diff
changeset | 908 | (field_updates_tr' ctxt u) | 
| 
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
 wenzelm parents: 
41928diff
changeset | 909 | | NONE => ([], tm)) | 
| 41578 
f5e7f6712815
recovered printing of record updates over compound terms, e.g. "(|x = a|)(|x := b|)", which was apparently broken in 45a2ffc5911e;
 wenzelm parents: 
41577diff
changeset | 910 | | NONE => ([], tm)) | 
| 35240 | 911 | | field_updates_tr' _ tm = ([], tm); | 
| 912 | ||
| 913 | fun record_update_tr' ctxt tm = | |
| 914 | (case field_updates_tr' ctxt tm of | |
| 915 | ([], _) => raise Match | |
| 916 | | (ts, u) => | |
| 67149 | 917 | Syntax.const \<^syntax_const>\<open>_record_update\<close> $ u $ | 
| 918 | foldr1 (fn (v, w) => Syntax.const \<^syntax_const>\<open>_field_updates\<close> $ v $ w) (rev ts)); | |
| 35240 | 919 | |
| 920 | in | |
| 921 | ||
| 922 | fun field_update_tr' name = | |
| 923 | let | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42284diff
changeset | 924 | val update_name = Lexicon.mark_const (name ^ updateN); | 
| 35240 | 925 | fun tr' ctxt [t, u] = record_update_tr' ctxt (Syntax.const update_name $ t $ u) | 
| 926 | | tr' _ _ = raise Match; | |
| 927 | in (update_name, tr') end; | |
| 928 | ||
| 929 | end; | |
| 930 | ||
| 931 | ||
| 32335 | 932 | |
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 933 | (** record simprocs **) | 
| 14358 | 934 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 935 | fun is_sel_upd_pair thy (Const (s, _)) (Const (u, t')) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 936 | (case get_updates thy u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 937 | SOME u_name => u_name = s | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 938 |   | NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')]));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 939 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 940 | fun mk_comp_id f = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 941 | let val T = range_type (fastype_of f) | 
| 67149 | 942 | in HOLogic.mk_comp (Const (\<^const_name>\<open>Fun.id\<close>, T --> T), f) end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 943 | |
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32749diff
changeset | 944 | fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 945 | | get_upd_funs _ = []; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 946 | |
| 60796 | 947 | fun get_accupd_simps ctxt term defset = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 948 | let | 
| 60796 | 949 | val thy = Proof_Context.theory_of ctxt; | 
| 950 | ||
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 951 | val (acc, [body]) = strip_comb term; | 
| 35408 | 952 | 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 | 953 | fun get_simp upd = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 954 | let | 
| 35133 | 955 | (* FIXME fresh "f" (!?) *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 956 | val T = domain_type (fastype_of upd); | 
| 40845 | 957 |         val lhs = HOLogic.mk_comp (acc, upd $ Free ("f", T));
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 958 | val rhs = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 959 | if is_sel_upd_pair thy acc upd | 
| 40845 | 960 |           then HOLogic.mk_comp (Free ("f", T), acc)
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 961 | else mk_comp_id acc; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 962 | val prop = lhs === rhs; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 963 | val othm = | 
| 60796 | 964 | Goal.prove ctxt [] [] prop | 
| 965 |             (fn {context = ctxt', ...} =>
 | |
| 966 | simp_tac (put_simpset defset ctxt') 1 THEN | |
| 967 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN | |
| 968 |               TRY (simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms id_apply id_o o_id}) 1));
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 969 | val dest = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 970 | if is_sel_upd_pair thy acc upd | 
| 46043 | 971 |           then @{thm o_eq_dest}
 | 
| 972 |           else @{thm 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 | 973 | 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 | 974 | 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 | 975 | |
| 60796 | 976 | fun get_updupd_simp ctxt defset u u' comp = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 977 | let | 
| 35133 | 978 | (* FIXME fresh "f" (!?) *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 979 |     val f = Free ("f", domain_type (fastype_of u));
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 980 |     val f' = Free ("f'", domain_type (fastype_of u'));
 | 
| 40845 | 981 | val lhs = HOLogic.mk_comp (u $ f, u' $ f'); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 982 | val rhs = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 983 | if comp | 
| 40845 | 984 | then u $ HOLogic.mk_comp (f, f') | 
| 985 | else HOLogic.mk_comp (u' $ f', u $ f); | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 986 | val prop = lhs === rhs; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 987 | val othm = | 
| 60796 | 988 | Goal.prove ctxt [] [] prop | 
| 989 |         (fn {context = ctxt', ...} =>
 | |
| 990 | simp_tac (put_simpset defset ctxt') 1 THEN | |
| 991 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN | |
| 992 |           TRY (simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms id_apply}) 1));
 | |
| 46043 | 993 |     val dest = if comp then @{thm o_eq_dest_lhs} else @{thm 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 | 994 | 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 | 995 | |
| 60796 | 996 | fun get_updupd_simps ctxt term defset = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 997 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 998 | val upd_funs = get_upd_funs term; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 999 | val cname = fst o dest_Const; | 
| 60796 | 1000 | fun getswap u u' = get_updupd_simp ctxt defset u u' (cname u = cname u'); | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1001 | fun build_swaps_to_eq _ [] swaps = swaps | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1002 | | build_swaps_to_eq upd (u :: us) swaps = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1003 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1004 | val key = (cname u, cname upd); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1005 | val newswaps = | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1006 | if Symreltab.defined swaps key then swaps | 
| 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1007 | else Symreltab.insert (K true) (key, getswap u upd) swaps; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1008 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1009 | if cname u = cname upd then newswaps | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1010 | else build_swaps_to_eq upd us newswaps | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1011 | end; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1012 | fun swaps_needed [] _ _ swaps = map snd (Symreltab.dest swaps) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1013 | | swaps_needed (u :: us) prev seen swaps = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1014 | if Symtab.defined seen (cname u) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1015 | then swaps_needed us prev seen (build_swaps_to_eq u prev swaps) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1016 | 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 | 1017 | 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 | 1018 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1019 | 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 | 1020 | let | 
| 60796 | 1021 | val ctxt = Proof_Context.init_global thy; | 
| 1022 | ||
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1023 | val defset = get_sel_upd_defs thy; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1024 | val prop' = Envir.beta_eta_contract prop; | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1025 | val (lhs, _) = Logic.dest_equals (Logic.strip_assums_concl prop'); | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1026 | val (_, args) = strip_comb lhs; | 
| 60796 | 1027 | val simps = (if length args = 1 then get_accupd_simps else get_updupd_simps) ctxt lhs defset; | 
| 16973 | 1028 | in | 
| 60796 | 1029 | Goal.prove ctxt [] [] prop' | 
| 1030 |       (fn {context = ctxt', ...} =>
 | |
| 1031 |         simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (simps @ @{thms K_record_comp})) 1 THEN
 | |
| 1032 | TRY (simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ex_simps addsimprocs ex_simprs) 1)) | |
| 15203 | 1033 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1034 | |
| 15215 | 1035 | |
| 15059 | 1036 | local | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1037 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1038 | fun eq (s1: string) (s2: string) = (s1 = s2); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1039 | |
| 16822 | 1040 | fun has_field extfields f T = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1041 | 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 | 1042 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1043 | 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 | 1044 | 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 | 1045 | | K_skeleton n T b _ = ((n, T), b); | 
| 25705 | 1046 | |
| 15059 | 1047 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1048 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1049 | (* simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1050 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1051 | (* | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1052 | Simplify selections of an record update: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1053 | (1) S (S_update k r) = k (S r) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1054 | (2) S (X_update k r) = S r | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1055 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1056 | The simproc skips multiple updates at once, eg: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1057 | 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 | 1058 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1059 | But be careful in (2) because of the extensibility of records. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1060 | - 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 | 1061 | X does not affect the selected subrecord. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1062 | - 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 | 1063 | subrecord. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1064 | *) | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1065 | val simproc = | 
| 69593 | 1066 | Simplifier.make_simproc \<^context> "record" | 
| 67149 | 1067 |    {lhss = [\<^term>\<open>x::'a::{}\<close>],
 | 
| 61144 | 1068 | proc = fn _ => fn ctxt => fn ct => | 
| 1069 | let | |
| 1070 | val thy = Proof_Context.theory_of ctxt; | |
| 1071 | val t = Thm.term_of ct; | |
| 1072 | in | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1073 | (case t of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1074 | (sel as Const (s, Type (_, [_, rangeS]))) $ | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1075 | ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1076 | if is_selector thy s andalso is_some (get_updates thy u) then | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1077 | let | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1078 |                 val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy;
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1079 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1080 | fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1081 | (case Symtab.lookup updates u of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1082 | NONE => NONE | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1083 | | SOME u_name => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1084 | if u_name = s then | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1085 | (case mk_eq_terms r of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1086 | NONE => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1087 | let | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1088 |                                   val rv = ("r", rT);
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1089 | val rb = Bound 0; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1090 | val (kv, kb) = K_skeleton "k" kT (Bound 1) k; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1091 | in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1092 | | SOME (trm, trm', vars) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1093 | let | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1094 | val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1095 | in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1096 | else if has_field extfields u_name rangeS orelse | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1097 | has_field extfields s (domain_type kT) then NONE | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1098 | else | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1099 | (case mk_eq_terms r of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1100 | SOME (trm, trm', vars) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1101 | let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1102 | in SOME (upd $ kb $ trm, trm', kv :: vars) end | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1103 | | NONE => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1104 | let | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1105 |                                   val rv = ("r", rT);
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1106 | val rb = Bound 0; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1107 | val (kv, kb) = K_skeleton "k" kT (Bound 1) k; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1108 | in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1109 | | mk_eq_terms _ = NONE; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1110 | in | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1111 | (case mk_eq_terms (upd $ k $ r) of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1112 | SOME (trm, trm', vars) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1113 | SOME | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1114 | (prove_unfold_defs thy [] [] | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1115 | (Logic.list_all (vars, Logic.mk_equals (sel $ trm, trm')))) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1116 | | NONE => NONE) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1117 | end | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1118 | else NONE | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1119 | | _ => NONE) | 
| 62913 | 1120 | end}; | 
| 7178 | 1121 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1122 | fun get_upd_acc_cong_thm upd acc thy ss = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1123 | let | 
| 60796 | 1124 | val ctxt = Proof_Context.init_global thy; | 
| 1125 | val prop = | |
| 1126 | infer_instantiate ctxt | |
| 1127 |         [(("upd", 0), Thm.cterm_of ctxt upd), (("ac", 0), Thm.cterm_of ctxt acc)]
 | |
| 1128 | updacc_cong_triv | |
| 1129 | |> Thm.concl_of; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1130 | in | 
| 60796 | 1131 | Goal.prove ctxt [] [] prop | 
| 1132 |       (fn {context = ctxt', ...} =>
 | |
| 1133 | simp_tac (put_simpset ss ctxt') 1 THEN | |
| 1134 | REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN | |
| 1135 | TRY (resolve_tac ctxt' [updacc_cong_idI] 1)) | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1136 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1137 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1138 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1139 | (* upd_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1140 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1141 | (*Simplify multiple updates: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1142 | (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 | 1143 | (N_update (y o x) (M_update (g o f) r))" | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1144 | (2) "r(|M:= M r|) = r" | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1145 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1146 | In both cases "more" updates complicate matters: for this reason | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1147 | we omit considering further updates if doing so would introduce | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1148 | both a more update and an update to a field within it.*) | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1149 | val upd_simproc = | 
| 69593 | 1150 | Simplifier.make_simproc \<^context> "record_upd" | 
| 67149 | 1151 |    {lhss = [\<^term>\<open>x::'a::{}\<close>],
 | 
| 61144 | 1152 | proc = fn _ => fn ctxt => fn ct => | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1153 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1154 | val thy = Proof_Context.theory_of ctxt; | 
| 61144 | 1155 | val t = Thm.term_of ct; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1156 | (*We can use more-updators with other updators as long | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1157 | as none of the other updators go deeper than any more | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1158 | updator. min here is the depth of the deepest other | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1159 | 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 | 1160 | fun include_depth (dep, true) (min, max) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1161 | if min <= dep | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1162 | then SOME (min, if dep <= max orelse max = ~1 then dep else max) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1163 | else NONE | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1164 | | include_depth (dep, false) (min, max) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1165 | if dep <= max orelse max = ~1 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1166 | then SOME (if min <= dep then dep else min, max) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1167 | else NONE; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1168 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1169 | fun getupdseq (term as (upd as Const (u, _)) $ f $ tm) min max = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1170 | (case get_update_details u thy of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1171 | SOME (s, dep, ismore) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1172 | (case include_depth (dep, ismore) (min, max) of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1173 | SOME (min', max') => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1174 | 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 | 1175 | in ((upd, s, f) :: us, bs, fastype_of term) end | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1176 | | NONE => ([], term, HOLogic.unitT)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1177 | | NONE => ([], term, HOLogic.unitT)) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1178 | | getupdseq term _ _ = ([], term, HOLogic.unitT); | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1179 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1180 | 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 | 1181 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1182 | 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 | 1183 | if s = s' andalso null (loose_bnos tm') | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1184 | andalso subst_bound (HOLogic.unit, tm') = tm | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1185 | then (true, Abs (n, T, Const (s', T') $ Bound 1)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1186 | else (false, HOLogic.unit) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1187 | | is_upd_noop _ _ _ = (false, HOLogic.unit); | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1188 | |
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1189 | fun get_noop_simps (upd as Const _) (Abs (_, _, (acc as Const _) $ _)) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1190 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1191 | 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 | 1192 | val uathm = get_upd_acc_cong_thm upd acc thy ss; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1193 | 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 | 1194 | [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 | 1195 | Drule.export_without_context (uathm RS updacc_noop_compE)] | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1196 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1197 | |
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1198 | (*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 | 1199 | only returns constant abstractions thus when we see an | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1200 | abstraction we can discard inner updates.*) | 
| 57225 
ff69e42ccf92
register record extensions as freely generated types
 blanchet parents: 
56732diff
changeset | 1201 | fun add_upd (f as Abs _) _ = [f] | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1202 | | 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 | 1203 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1204 | (*mk_updterm returns | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1205 | (orig-term-skeleton, simplified-skeleton, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1206 | variables, duplicate-updates, simp-flag, noop-simps) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1207 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1208 | where duplicate-updates is a table used to pass upward | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1209 | the list of update functions which can be composed | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1210 | into an update above them, simp-flag indicates whether | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1211 | any simplification was achieved, and noop-simps are | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1212 | used for eliminating case (2) defined above*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1213 | fun mk_updterm ((upd as Const (u, T), s, f) :: upds) above term = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1214 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1215 | 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 | 1216 | mk_updterm upds (Symtab.update (u, ()) above) term; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1217 | val (fvar, skelf) = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1218 | 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 | 1219 | val (isnoop, skelf') = is_upd_noop s f term; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1220 | val funT = domain_type T; | 
| 35133 | 1221 | fun mk_comp_local (f, f') = | 
| 67149 | 1222 | Const (\<^const_name>\<open>Fun.comp\<close>, funT --> funT --> funT) $ f $ f'; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1223 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1224 | if isnoop then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1225 | (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 | 1226 | Symtab.update (u, []) dups, true, | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1227 | 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 | 1228 | else Symtab.update (u, get_noop_simps upd skelf') noops) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1229 | else if Symtab.defined above u then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1230 | (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 | 1231 | 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 | 1232 | true, noops) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1233 | else | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1234 | (case Symtab.lookup dups u of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1235 | SOME fs => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1236 | (upd $ skelf $ lhs, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1237 | upd $ foldr1 mk_comp_local (add_upd skelf fs) $ rhs, | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1238 | fvar :: vars, dups, true, noops) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1239 | | NONE => (upd $ skelf $ lhs, upd $ skelf $ rhs, fvar :: vars, dups, simp, noops)) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1240 | end | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1241 | | mk_updterm [] _ _ = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1242 |               (Bound 0, Bound 0, [("r", baseT)], Symtab.empty, false, Symtab.empty)
 | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1243 |           | 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 | 1244 | |
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1245 | val (lhs, rhs, vars, _, simp, noops) = mk_updterm upds Symtab.empty base; | 
| 32952 | 1246 | 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 | 1247 | in | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1248 | if simp then | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1249 | SOME | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1250 | (prove_unfold_defs thy noops' [simproc] | 
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
46215diff
changeset | 1251 | (Logic.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 | 1252 | else NONE | 
| 62913 | 1253 | end}; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1254 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1255 | end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1256 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1257 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1258 | (* eq_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1259 | |
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1260 | (*Look up the most specific record-equality. | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1261 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1262 | Note on efficiency: | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1263 | 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 | 1264 | Therefore the complexity is: #components * complexity for single component. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1265 | Especially if a record has a lot of components it may be better to split up | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1266 | the record first and do simplification on that (split_simp_tac). | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1267 | e.g. r(|lots of updates|) = x | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1268 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1269 | eq_simproc split_simp_tac | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1270 | Complexity: #components * #updates #updates | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1271 | *) | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1272 | val eq_simproc = | 
| 69593 | 1273 | Simplifier.make_simproc \<^context> "record_eq" | 
| 67149 | 1274 |    {lhss = [\<^term>\<open>r = s\<close>],
 | 
| 61144 | 1275 | proc = fn _ => fn ctxt => fn ct => | 
| 1276 | (case Thm.term_of ct of | |
| 67149 | 1277 | Const (\<^const_name>\<open>HOL.eq\<close>, Type (_, [T, _])) $ _ $ _ => | 
| 61144 | 1278 | (case rec_id ~1 T of | 
| 1279 | "" => NONE | |
| 1280 | | name => | |
| 1281 | (case get_equalities (Proof_Context.theory_of ctxt) name of | |
| 1282 | NONE => NONE | |
| 1283 |               | SOME thm => SOME (thm RS @{thm Eq_TrueI})))
 | |
| 62913 | 1284 | | _ => NONE)}; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1285 | |
| 14079 
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
 berghofe parents: 
13904diff
changeset | 1286 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1287 | (* split_simproc *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1288 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1289 | (*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 | 1290 | subterm starting at the quantified occurrence of the record (including the quantifier): | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1291 | P t = 0: do not split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1292 | P t = ~1: completely split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1293 | P t > 0: split up to given bound of record extensions.*) | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1294 | fun split_simproc P = | 
| 69593 | 1295 | Simplifier.make_simproc \<^context> "record_split" | 
| 67149 | 1296 |    {lhss = [\<^term>\<open>x::'a::{}\<close>],
 | 
| 61144 | 1297 | proc = fn _ => fn ctxt => fn ct => | 
| 1298 | (case Thm.term_of ct of | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1299 | Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ => | 
| 67149 | 1300 | if quantifier = \<^const_name>\<open>Pure.all\<close> orelse | 
| 1301 | quantifier = \<^const_name>\<open>All\<close> orelse | |
| 1302 | quantifier = \<^const_name>\<open>Ex\<close> | |
| 35133 | 1303 | then | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1304 | (case rec_id ~1 T of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1305 | "" => NONE | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1306 | | _ => | 
| 61144 | 1307 | let val split = P (Thm.term_of ct) in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1308 | if split <> 0 then | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1309 | (case get_splits (Proof_Context.theory_of ctxt) (rec_id split T) of | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1310 | NONE => NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1311 | | SOME (all_thm, All_thm, Ex_thm, _) => | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1312 | SOME | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1313 | (case quantifier of | 
| 67149 | 1314 | \<^const_name>\<open>Pure.all\<close> => all_thm | 
| 1315 |                           | \<^const_name>\<open>All\<close> => All_thm RS @{thm eq_reflection}
 | |
| 1316 |                           | \<^const_name>\<open>Ex\<close> => Ex_thm RS @{thm eq_reflection}
 | |
| 43683 | 1317 | | _ => raise Fail "split_simproc")) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1318 | else NONE | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1319 | end) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1320 | else NONE | 
| 62913 | 1321 | | _ => NONE)}; | 
| 7178 | 1322 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1323 | val ex_sel_eq_simproc = | 
| 69593 | 1324 | Simplifier.make_simproc \<^context> "ex_sel_eq" | 
| 67149 | 1325 |    {lhss = [\<^term>\<open>Ex t\<close>],
 | 
| 61144 | 1326 | proc = fn _ => fn ctxt => fn ct => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1327 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1328 | val thy = Proof_Context.theory_of ctxt; | 
| 61144 | 1329 | val t = Thm.term_of ct; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1330 | fun mkeq (lr, Teq, (sel, Tsel), x) i = | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1331 | if is_selector thy sel then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1332 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1333 | val x' = | 
| 47234 | 1334 | if not (Term.is_dependent x) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1335 |                 then Free ("x" ^ string_of_int i, range_type Tsel)
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1336 |                 else raise TERM ("", [x]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1337 | val sel' = Const (sel, Tsel) $ Bound 0; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1338 | val (l, r) = if lr then (sel', x') else (x', sel'); | 
| 67149 | 1339 | in Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ l $ r end | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1340 |           else raise TERM ("", [Const (sel, Tsel)]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1341 | |
| 67149 | 1342 | fun dest_sel_eq (Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1343 | (true, Teq, (sel, Tsel), X) | 
| 67149 | 1344 | | dest_sel_eq (Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1345 | (false, Teq, (sel, Tsel), X) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1346 |           | dest_sel_eq _ = raise TERM ("", []);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1347 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1348 | (case t of | 
| 67149 | 1349 | Const (\<^const_name>\<open>Ex\<close>, Tex) $ Abs (s, T, t) => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1350 | (let | 
| 35133 | 1351 | val eq = mkeq (dest_sel_eq t) 0; | 
| 1352 | val prop = | |
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
46215diff
changeset | 1353 |                Logic.list_all ([("r", T)],
 | 
| 35133 | 1354 | Logic.mk_equals | 
| 67149 | 1355 | (Const (\<^const_name>\<open>Ex\<close>, Tex) $ Abs (s, T, eq), \<^term>\<open>True\<close>)); | 
| 46049 | 1356 | in | 
| 51551 | 1357 | SOME (Goal.prove_sorry_global thy [] [] prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1358 | (fn _ => simp_tac (put_simpset (get_simpset thy) ctxt | 
| 46049 | 1359 |                     addsimps @{thms simp_thms} addsimprocs [split_simproc (K ~1)]) 1))
 | 
| 1360 | end handle TERM _ => NONE) | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1361 | | _ => NONE) | 
| 62913 | 1362 | end}; | 
| 14427 | 1363 | |
| 5698 | 1364 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1365 | (* split_simp_tac *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1366 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1367 | (*Split (and simplify) all records in the goal for which P holds. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1368 | For quantified occurrences of a record | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1369 | 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 | 1370 | can only peek on the variable itself. | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1371 | P t = 0: do not split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1372 | P t = ~1: completely split | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1373 | P t > 0: split up to given bound of record extensions.*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1374 | fun split_simp_tac ctxt thms P = CSUBGOAL (fn (cgoal, i) => | 
| 14255 | 1375 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1376 | val thy = Proof_Context.theory_of ctxt; | 
| 32975 | 1377 | |
| 59582 | 1378 | val goal = Thm.term_of cgoal; | 
| 32975 | 1379 | val frees = filter (is_recT o #2) (Term.add_frees goal []); | 
| 14255 | 1380 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1381 | val has_rec = exists_Const | 
| 14255 | 1382 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 67149 | 1383 | (s = \<^const_name>\<open>Pure.all\<close> orelse s = \<^const_name>\<open>All\<close> orelse s = \<^const_name>\<open>Ex\<close>) | 
| 56245 | 1384 | andalso is_recT T | 
| 14255 | 1385 | | _ => false); | 
| 1386 | ||
| 17261 | 1387 | fun mk_split_free_tac free induct_thm i = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1388 | let | 
| 60796 | 1389 | val _ $ (_ $ Var (r, _)) = Thm.concl_of induct_thm; | 
| 1390 | val thm = infer_instantiate ctxt [(r, Thm.cterm_of ctxt free)] induct_thm; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1391 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1392 |         simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) i THEN
 | 
| 60752 | 1393 | resolve_tac ctxt [thm] i THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1394 |         simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_rulify}) i
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1395 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1396 | |
| 32975 | 1397 | val split_frees_tacs = | 
| 1398 | frees |> map_filter (fn (x, T) => | |
| 1399 | (case rec_id ~1 T of | |
| 1400 | "" => NONE | |
| 1401 | | _ => | |
| 1402 | let | |
| 1403 | val free = Free (x, T); | |
| 1404 | val split = P free; | |
| 1405 | in | |
| 1406 | if split <> 0 then | |
| 1407 | (case get_splits thy (rec_id split T) of | |
| 1408 | NONE => NONE | |
| 1409 | | SOME (_, _, _, induct_thm) => | |
| 1410 | SOME (mk_split_free_tac free induct_thm i)) | |
| 1411 | else NONE | |
| 1412 | end)); | |
| 17261 | 1413 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1414 | val simprocs = if has_rec goal then [split_simproc P] else []; | 
| 46043 | 1415 |     val thms' = @{thms o_apply K_record_comp} @ thms;
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1416 | in | 
| 32975 | 1417 | EVERY split_frees_tacs THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1418 | full_simp_tac (put_simpset (get_simpset thy) ctxt addsimps thms' addsimprocs simprocs) i | 
| 32975 | 1419 | end); | 
| 14255 | 1420 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1421 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1422 | (* split_tac *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1423 | |
| 35147 | 1424 | (*Split all records in the goal, which are quantified by !! or ALL.*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1425 | fun split_tac ctxt = CSUBGOAL (fn (cgoal, i) => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1426 | let | 
| 59582 | 1427 | val goal = Thm.term_of cgoal; | 
| 32975 | 1428 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1429 | val has_rec = exists_Const | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1430 | (fn (s, Type (_, [Type (_, [T, _]), _])) => | 
| 67149 | 1431 | (s = \<^const_name>\<open>Pure.all\<close> orelse s = \<^const_name>\<open>All\<close>) andalso is_recT T | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1432 | | _ => false); | 
| 17261 | 1433 | |
| 67149 | 1434 | fun is_all (Const (\<^const_name>\<open>Pure.all\<close>, _) $ _) = ~1 | 
| 1435 | | is_all (Const (\<^const_name>\<open>All\<close>, _) $ _) = ~1 | |
| 35240 | 1436 | | is_all _ = 0; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1437 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1438 | if has_rec goal then | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1439 | full_simp_tac (put_simpset HOL_basic_ss ctxt addsimprocs [split_simproc is_all]) i | 
| 32975 | 1440 | else no_tac | 
| 1441 | end); | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1442 | |
| 56732 | 1443 | val _ = | 
| 1444 | Theory.setup | |
| 1445 | (map_theory_simpset (fn ctxt => ctxt addsimprocs [simproc, upd_simproc, eq_simproc])); | |
| 1446 | ||
| 32335 | 1447 | |
| 6358 | 1448 | (* wrapper *) | 
| 1449 | ||
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1450 | val split_name = "record_split_tac"; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1451 | val split_wrapper = (split_name, fn ctxt => fn tac => split_tac ctxt ORELSE' tac); | 
| 5698 | 1452 | |
| 16330 | 1453 | |
| 1454 | ||
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1455 | (** theory extender interface **) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 1456 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1457 | (* attributes *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1458 | |
| 33368 | 1459 | 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 | 1460 | 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 | 1461 | 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 | 1462 | |
| 32335 | 1463 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1464 | (* tactics *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1465 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1466 | (*Do case analysis / induction according to rule on last parameter of ith subgoal | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1467 | (or on s if there are no parameters). | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1468 | Instatiation of record variable (and predicate) in rule is calculated to | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1469 | avoid problems with higher order unification.*) | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 1470 | fun try_param_tac ctxt s rule = CSUBGOAL (fn (cgoal, i) => | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1471 | let | 
| 32975 | 1472 | val g = Thm.term_of cgoal; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1473 | val params = Logic.strip_params g; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1474 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g); | 
| 32975 | 1475 | val rule' = Thm.lift_rule cgoal rule; | 
| 59582 | 1476 | val (P, ys) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_assums_concl (Thm.prop_of rule'))); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1477 | (*ca indicates if rule is a case analysis or induction rule*) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1478 | val (x, ca) = | 
| 33957 | 1479 | (case rev (drop (length params) ys) of | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1480 | [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop | 
| 59582 | 1481 | (hd (rev (Logic.strip_assums_hyp (hd (Thm.prems_of rule')))))))), true) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1482 | | [x] => (head_of x, false)); | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1483 | val rule'' = | 
| 60796 | 1484 | infer_instantiate ctxt | 
| 1485 | (map (apsnd (Thm.cterm_of ctxt)) | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1486 | (case rev params of | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1487 | [] => | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1488 | (case AList.lookup (op =) (Term.add_frees g []) s of | 
| 40315 
11846d9800de
try_param_tac: plain user error appears more appropriate;
 wenzelm parents: 
39557diff
changeset | 1489 | NONE => error "try_param_tac: no such variable" | 
| 60796 | 1490 | | SOME T => | 
| 1491 | [(#1 (dest_Var P), if ca then concl else lambda (Free (s, T)) concl), | |
| 1492 | (#1 (dest_Var x), Free (s, T))]) | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1493 | | (_, T) :: _ => | 
| 60796 | 1494 | [(#1 (dest_Var P), | 
| 1495 | fold_rev Term.abs params | |
| 1496 | (if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))), | |
| 1497 | (#1 (dest_Var x), fold_rev Term.abs params (Bound 0))])) rule'; | |
| 59582 | 1498 | in compose_tac ctxt (false, rule'', Thm.nprems_of rule) i end); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1499 | |
| 15215 | 1500 | |
| 61260 | 1501 | fun extension_definition overloaded name fields alphas zeta moreT more vars thy = | 
| 17261 | 1502 | let | 
| 43682 | 1503 | val ctxt = Proof_Context.init_global thy; | 
| 1504 | ||
| 35239 | 1505 | val base_name = Long_Name.base_name name; | 
| 1506 | ||
| 32977 | 1507 | val fieldTs = map snd fields; | 
| 35239 | 1508 | val fields_moreTs = fieldTs @ [moreT]; | 
| 1509 | ||
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1510 | val alphas_zeta = alphas @ [zeta]; | 
| 35239 | 1511 | |
| 1512 | val ext_binding = Binding.name (suffix extN base_name); | |
| 1513 | val ext_name = suffix extN name; | |
| 57225 
ff69e42ccf92
register record extensions as freely generated types
 blanchet parents: 
56732diff
changeset | 1514 | val ext_tyco = suffix ext_typeN name; | 
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1515 | val extT = Type (ext_tyco, map TFree alphas_zeta); | 
| 35239 | 1516 | val ext_type = fields_moreTs ---> extT; | 
| 1517 | ||
| 1518 | ||
| 1519 | (* the tree of new types that will back the record extension *) | |
| 32767 | 1520 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1521 | val mktreeV = Balanced_Tree.make Iso_Tuple_Support.mk_cons_tuple; | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1522 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1523 | fun mk_iso_tuple (left, right) (thy, i) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1524 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1525 | val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i; | 
| 35239 | 1526 | val ((_, cons), thy') = thy | 
| 61260 | 1527 | |> Iso_Tuple_Support.add_iso_tuple_type overloaded | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1528 | (Binding.suffix_name suff (Binding.name base_name), alphas_zeta) | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1529 | (fastype_of left, fastype_of right); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1530 | in | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1531 | (cons $ left $ right, (thy', i + 1)) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1532 | end; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1533 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1534 | (*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 | 1535 | fun mk_even_iso_tuple [arg] = pair arg | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1536 | | 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 | 1537 | |
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1538 | fun build_meta_tree_type i thy vars more = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1539 | let val len = length vars in | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1540 |         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 | 1541 | else if len > 16 then | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1542 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1543 | fun group16 [] = [] | 
| 33957 | 1544 | | group16 xs = take 16 xs :: group16 (drop 16 xs); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1545 | val vars' = group16 vars; | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1546 | 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 | 1547 | in | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1548 | build_meta_tree_type i' thy' composites more | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1549 | end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1550 | else | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1551 | let val (term, (thy', _)) = mk_iso_tuple (mktreeV vars, more) (thy, 0) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1552 | in (term, thy') end | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1553 | end; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1554 | |
| 43682 | 1555 | val _ = timing_msg ctxt "record extension preparing definitions"; | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1556 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1557 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1558 | (* 1st stage part 1: introduce the tree of new types *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1559 | |
| 46044 | 1560 | val (ext_body, typ_thy) = timeit_msg ctxt "record extension nested type def:" (fn () => | 
| 1561 | build_meta_tree_type 1 thy vars more); | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1562 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1563 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1564 | (* prepare declarations and definitions *) | 
| 17261 | 1565 | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1566 | (* 1st stage part 2: define the ext constant *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1567 | |
| 35239 | 1568 | fun mk_ext args = list_comb (Const (ext_name, ext_type), args); | 
| 1569 | val ext_spec = Logic.mk_equals (mk_ext (vars @ [more]), ext_body); | |
| 1570 | ||
| 46044 | 1571 | val ([ext_def], defs_thy) = timeit_msg ctxt "record extension constructor def:" (fn () => | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1572 | typ_thy | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 1573 | |> Sign.declare_const_global ((ext_binding, ext_type), NoSyn) |> snd | 
| 52788 | 1574 | |> Global_Theory.add_defs false [((Thm.def_binding ext_binding, ext_spec), [])]); | 
| 54895 | 1575 | val defs_ctxt = Proof_Context.init_global defs_thy; | 
| 17261 | 1576 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1577 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1578 | (* prepare propositions *) | 
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1579 | |
| 43682 | 1580 | val _ = timing_msg ctxt "record extension preparing propositions"; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1581 | val vars_more = vars @ [more]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1582 | val variants = map (fn Free (x, _) => x) vars_more; | 
| 15215 | 1583 | val ext = mk_ext vars_more; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1584 | val s = Free (rN, extT); | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 1585 | val P = Free (singleton (Name.variant_list variants) "P", extT --> HOLogic.boolT); | 
| 17261 | 1586 | |
| 43683 | 1587 | val inject_prop = (* FIXME local x x' *) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1588 | 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 | 1589 | HOLogic.mk_conj (HOLogic.eq_const extT $ | 
| 67149 | 1590 | mk_ext vars_more $ mk_ext vars_more', \<^term>\<open>True\<close>) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1591 | === | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1592 | foldr1 HOLogic.mk_conj | 
| 67149 | 1593 | (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [\<^term>\<open>True\<close>]) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1594 | end; | 
| 17261 | 1595 | |
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 1596 | val induct_prop = (fold_rev Logic.all vars_more (Trueprop (P $ ext)), Trueprop (P $ s)); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 1597 | |
| 43683 | 1598 | val split_meta_prop = (* FIXME local P *) | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 1599 | let val P = Free (singleton (Name.variant_list variants) "P", extT --> propT) | 
| 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 1600 | in Logic.mk_equals (Logic.all s (P $ s), fold_rev Logic.all vars_more (P $ ext)) end; | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 1601 | |
| 46044 | 1602 | val inject = timeit_msg ctxt "record extension inject proof:" (fn () => | 
| 54895 | 1603 | simplify (put_simpset HOL_ss defs_ctxt) | 
| 51551 | 1604 | (Goal.prove_sorry_global defs_thy [] [] inject_prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1605 |           (fn {context = ctxt, ...} =>
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1606 | simp_tac (put_simpset HOL_basic_ss ctxt addsimps [ext_def]) 1 THEN | 
| 32975 | 1607 | REPEAT_DETERM | 
| 60752 | 1608 |               (resolve_tac ctxt @{thms refl_conj_eq} 1 ORELSE
 | 
| 59164 | 1609 | Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE | 
| 60752 | 1610 | resolve_tac ctxt [refl] 1)))); | 
| 46044 | 1611 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1612 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1613 | (*We need a surjection property r = (| f = f r, g = g r ... |) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1614 | to prove other theorems. We haven't given names to the accessors | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1615 | f, g etc yet however, so we generate an ext structure with | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1616 | 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 | 1617 | 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 | 1618 | to convert the free variables into unifiable variables and unify them with | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1619 | (roughly) the definition of the accessor.*) | 
| 46044 | 1620 | val surject = timeit_msg ctxt "record extension surjective proof:" (fn () => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1621 | let | 
| 60796 | 1622 | val start = | 
| 1623 |           infer_instantiate defs_ctxt [(("y", 0), Thm.cterm_of defs_ctxt ext)] surject_assist_idE;
 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1624 | val tactic1 = | 
| 54895 | 1625 | simp_tac (put_simpset HOL_basic_ss defs_ctxt addsimps [ext_def]) 1 THEN | 
| 59164 | 1626 | REPEAT_ALL_NEW (Iso_Tuple_Support.iso_tuple_intros_tac defs_ctxt) 1; | 
| 60752 | 1627 | val tactic2 = | 
| 1628 | REPEAT (resolve_tac defs_ctxt [surject_assistI] 1 THEN resolve_tac defs_ctxt [refl] 1); | |
| 32972 | 1629 | 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 | 1630 | 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 | 1631 | in | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1632 | surject | 
| 46044 | 1633 | end); | 
| 1634 | ||
| 1635 | val split_meta = timeit_msg ctxt "record extension split_meta proof:" (fn () => | |
| 51551 | 1636 | Goal.prove_sorry_global defs_thy [] [] split_meta_prop | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 1637 |         (fn {context = ctxt, ...} =>
 | 
| 32975 | 1638 | EVERY1 | 
| 60752 | 1639 |            [resolve_tac ctxt @{thms equal_intr_rule},
 | 
| 1640 | Goal.norm_hhf_tac ctxt, | |
| 1641 |             eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt,
 | |
| 1642 |             resolve_tac ctxt [@{thm prop_subst} OF [surject]],
 | |
| 1643 |             REPEAT o eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt]));
 | |
| 46044 | 1644 | |
| 1645 | val induct = timeit_msg ctxt "record extension induct proof:" (fn () => | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1646 | let val (assm, concl) = induct_prop in | 
| 51551 | 1647 | Goal.prove_sorry_global defs_thy [] [assm] concl | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1648 |           (fn {context = ctxt, prems, ...} =>
 | 
| 46708 
b138dee7bed3
prefer cut_tac, where it is clear that the special variants cut_rules_tac or cut_facts_tac are not required;
 wenzelm parents: 
46223diff
changeset | 1649 | cut_tac (split_meta RS Drule.equal_elim_rule2) 1 THEN | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 1650 | resolve_tac ctxt prems 2 THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 1651 | asm_simp_tac (put_simpset HOL_ss ctxt) 1) | 
| 46044 | 1652 | end); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1653 | |
| 46055 | 1654 | val ([(_, [induct']), (_, [inject']), (_, [surjective']), (_, [split_meta'])], thm_thy) = | 
| 17261 | 1655 | defs_thy | 
| 46055 | 1656 | |> Global_Theory.note_thmss "" | 
| 1657 | [((Binding.name "ext_induct", []), [([induct], [])]), | |
| 1658 | ((Binding.name "ext_inject", []), [([inject], [])]), | |
| 1659 | ((Binding.name "ext_surjective", []), [([surject], [])]), | |
| 1660 | ((Binding.name "ext_split", []), [([split_meta], [])])]; | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1661 | in | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1662 | (((ext_name, ext_type), (ext_tyco, alphas_zeta), | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1663 | extT, induct', inject', surjective', split_meta', ext_def), thm_thy) | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1664 | end; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1665 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1666 | fun chunks [] [] = [] | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1667 | | chunks [] xs = [xs] | 
| 33957 | 1668 | | chunks (l :: ls) xs = take l xs :: chunks ls (drop l xs); | 
| 17261 | 1669 | |
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1670 | fun chop_last [] = error "chop_last: list should not be empty" | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1671 | | chop_last [x] = ([], x) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1672 | | 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 | 1673 | |
| 32799 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1674 | fun subst_last _ [] = error "subst_last: list should not be empty" | 
| 
7478ea535416
eliminated dead code, redundant bindings and parameters;
 wenzelm parents: 
32770diff
changeset | 1675 | | subst_last s [_] = [s] | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1676 | | subst_last s (x :: xs) = x :: subst_last s xs; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1677 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1678 | |
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1679 | (* mk_recordT *) | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1680 | |
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1681 | (*build up the record type from the current extension tpye extT and a list | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1682 | of parent extensions, starting with the root of the record hierarchy*) | 
| 21078 | 1683 | fun mk_recordT extT = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1684 | fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT; | 
| 15215 | 1685 | |
| 1686 | ||
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1687 | (* code generation *) | 
| 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1688 | |
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1689 | fun mk_random_eq tyco vs extN Ts = | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1690 | let | 
| 43683 | 1691 | (* FIXME local i etc. *) | 
| 67149 | 1692 | val size = \<^term>\<open>i::natural\<close>; | 
| 1693 | fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> term\<close>); | |
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1694 | val T = Type (tyco, map TFree vs); | 
| 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1695 | val Tm = termifyT T; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43324diff
changeset | 1696 | val params = Name.invent_names Name.context "x" Ts; | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1697 | val lhs = HOLogic.mk_random T size; | 
| 67149 | 1698 | val tc = HOLogic.mk_return Tm \<^typ>\<open>Random.seed\<close> | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1699 | (HOLogic.mk_valtermify_app extN params T); | 
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1700 | val rhs = | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1701 | HOLogic.mk_ST | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1702 | (map (fn (v, T') => | 
| 67149 | 1703 | ((HOLogic.mk_random T' size, \<^typ>\<open>Random.seed\<close>), SOME (v, termifyT T'))) params) | 
| 1704 | tc \<^typ>\<open>Random.seed\<close> (SOME Tm, \<^typ>\<open>Random.seed\<close>); | |
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1705 | in | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1706 | (lhs, rhs) | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1707 | end | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1708 | |
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1709 | fun mk_full_exhaustive_eq tyco vs extN Ts = | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1710 | let | 
| 43683 | 1711 | (* FIXME local i etc. *) | 
| 67149 | 1712 | val size = \<^term>\<open>i::natural\<close>; | 
| 1713 | fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> term\<close>); | |
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1714 | val T = Type (tyco, map TFree vs); | 
| 67149 | 1715 |     val test_function = Free ("f", termifyT T --> \<^typ>\<open>(bool \<times> term list) option\<close>);
 | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43324diff
changeset | 1716 | val params = Name.invent_names Name.context "x" Ts; | 
| 43683 | 1717 | fun full_exhaustiveT T = | 
| 67149 | 1718 | (termifyT T --> \<^typ>\<open>(bool \<times> Code_Evaluation.term list) option\<close>) --> | 
| 1719 | \<^typ>\<open>natural\<close> --> \<^typ>\<open>(bool \<times> Code_Evaluation.term list) option\<close>; | |
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1720 | fun mk_full_exhaustive T = | 
| 67149 | 1721 | Const (\<^const_name>\<open>Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive\<close>, | 
| 43683 | 1722 | full_exhaustiveT T); | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1723 | val lhs = mk_full_exhaustive T $ test_function $ size; | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1724 | val tc = test_function $ (HOLogic.mk_valtermify_app extN params T); | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1725 | val rhs = fold_rev (fn (v, T) => fn cont => | 
| 43683 | 1726 | mk_full_exhaustive T $ (lambda (Free (v, termifyT T)) cont) $ size) params tc; | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1727 | in | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1728 | (lhs, rhs) | 
| 43683 | 1729 | end; | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1730 | |
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1731 | fun instantiate_sort_record (sort, mk_eq) tyco vs extN Ts thy = | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1732 | let | 
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1733 | val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_eq tyco vs extN Ts)); | 
| 41576 | 1734 | in | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1735 | thy | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1736 | |> Class.instantiation ([tyco], vs, sort) | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1737 | |> `(fn lthy => Syntax.check_term lthy eq) | 
| 63180 | 1738 | |-> (fn eq => Specification.definition NONE [] [] ((Binding.concealed Binding.empty, []), eq)) | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1739 | |> snd | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 1740 | |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) | 
| 43683 | 1741 | end; | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1742 | |
| 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1743 | fun ensure_sort_record (sort, mk_eq) ext_tyco vs extN Ts thy = | 
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1744 | let | 
| 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1745 | val algebra = Sign.classes_of thy; | 
| 48272 | 1746 | val has_inst = Sorts.has_instance algebra ext_tyco sort; | 
| 41576 | 1747 | in | 
| 1748 | if has_inst then thy | |
| 1749 | else | |
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1750 | (case Quickcheck_Common.perhaps_constrain thy (map (rpair sort) Ts) vs of | 
| 41576 | 1751 | SOME constrain => | 
| 42695 
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
 bulwahn parents: 
42381diff
changeset | 1752 | instantiate_sort_record (sort, mk_eq) ext_tyco (map constrain vs) extN | 
| 41576 | 1753 | ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy | 
| 1754 | | NONE => thy) | |
| 38544 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1755 | end; | 
| 
ac554311b1b9
re-added instantiation of type class random for records
 haftmann parents: 
38534diff
changeset | 1756 | |
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1757 | fun add_code ext_tyco vs extT ext simps inject thy = | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1758 | if Config.get_global thy codegen then | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1759 | let | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1760 | val eq = | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1761 | HOLogic.mk_Trueprop (HOLogic.mk_eq | 
| 67149 | 1762 | (Const (\<^const_name>\<open>HOL.equal\<close>, extT --> extT --> HOLogic.boolT), | 
| 1763 | Const (\<^const_name>\<open>HOL.eq\<close>, extT --> extT --> HOLogic.boolT))); | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 1764 | fun tac ctxt eq_def = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 1765 | Class.intro_classes_tac ctxt [] | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 1766 | THEN rewrite_goals_tac ctxt [Simpdata.mk_eq eq_def] | 
| 60752 | 1767 |         THEN ALLGOALS (resolve_tac ctxt @{thms refl});
 | 
| 63073 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
63064diff
changeset | 1768 | fun mk_eq ctxt eq_def = | 
| 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
63064diff
changeset | 1769 | rewrite_rule ctxt | 
| 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
63064diff
changeset | 1770 | [Axclass.unoverload ctxt (Thm.symmetric (Simpdata.mk_eq eq_def))] inject; | 
| 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
63064diff
changeset | 1771 | fun mk_eq_refl ctxt = | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1772 |         @{thm equal_refl}
 | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1773 | |> Thm.instantiate | 
| 74290 | 1774 | (TVars.make [(\<^tvar>\<open>?'a::equal\<close>, Thm.ctyp_of ctxt (Logic.varifyT_global extT))], Vars.empty) | 
| 63073 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
63064diff
changeset | 1775 | |> Axclass.unoverload ctxt; | 
| 67149 | 1776 | val ensure_random_record = ensure_sort_record (\<^sort>\<open>random\<close>, mk_random_eq); | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1777 | val ensure_exhaustive_record = | 
| 67149 | 1778 | ensure_sort_record (\<^sort>\<open>full_exhaustive\<close>, mk_full_exhaustive_eq); | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1779 | fun add_code eq_def thy = | 
| 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1780 | let | 
| 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1781 | val ctxt = Proof_Context.init_global thy; | 
| 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1782 | in | 
| 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1783 | thy | 
| 74292 | 1784 | |> Code.declare_default_eqns_global [(mk_eq ctxt eq_def, true), (mk_eq_refl ctxt, false)] | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1785 | end; | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1786 | in | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1787 | thy | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1788 | |> Code.declare_datatype_global [ext] | 
| 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1789 | |> Code.declare_default_eqns_global (map (rpair true) simps) | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1790 | |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_equal]) | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1791 | |> `(fn lthy => Syntax.check_term lthy eq) | 
| 63352 | 1792 | |-> (fn eq => Specification.definition NONE [] [] (Binding.empty_atts, eq)) | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1793 | |-> (fn (_, (_, eq_def)) => | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 1794 | Class.prove_instantiation_exit_result Morphism.thm tac eq_def) | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 1795 | |-> add_code | 
| 47842 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1796 | |> ensure_random_record ext_tyco vs (fst ext) (binder_types (snd ext)) | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1797 | |> ensure_exhaustive_record ext_tyco vs (fst ext) (binder_types (snd ext)) | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1798 | end | 
| 
bfc08ce7b7b9
provide [[record_codegen]] option for skipping codegen setup for records
 Gerwin Klein <gerwin.klein@nicta.com.au> parents: 
47783diff
changeset | 1799 | else thy; | 
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1800 | |
| 58363 
a5c08cd60a3f
take out selectors for records -- for derived records, these don't quite have the right type
 blanchet parents: 
58239diff
changeset | 1801 | fun add_ctr_sugar ctr exhaust inject sel_thms = | 
| 58188 | 1802 | Ctr_Sugar.default_register_ctr_sugar_global (K true) | 
| 58675 | 1803 |     {kind = Ctr_Sugar.Record, T = body_type (fastype_of ctr), ctrs = [ctr], casex = Term.dummy,
 | 
| 1804 | discs = [], selss = [], exhaust = exhaust, nchotomy = Drule.dummy_thm, injects = [inject], | |
| 1805 | distincts = [], case_thms = [], case_cong = Drule.dummy_thm, case_cong_weak = Drule.dummy_thm, | |
| 59266 | 1806 | case_distribs = [], split = Drule.dummy_thm, split_asm = Drule.dummy_thm, disc_defs = [], | 
| 59271 | 1807 | disc_thmss = [], discIs = [], disc_eq_cases = [], sel_defs = [], sel_thmss = [sel_thms], | 
| 1808 | distinct_discsss = [], exhaust_discs = [], exhaust_sels = [], collapses = [], expands = [], | |
| 1809 | split_sels = [], split_sel_asms = [], case_eq_ifs = []}; | |
| 57225 
ff69e42ccf92
register record extensions as freely generated types
 blanchet parents: 
56732diff
changeset | 1810 | |
| 67149 | 1811 | fun lhs_of_equation (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t $ _) = t | 
| 1812 | | lhs_of_equation (\<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t $ _)) = t; | |
| 61861 | 1813 | |
| 1814 | fun add_spec_rule rule = | |
| 71179 | 1815 | let val head = head_of (lhs_of_equation (Thm.prop_of rule)) | 
| 71214 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 wenzelm parents: 
71179diff
changeset | 1816 | in Spec_Rules.add_global Binding.empty Spec_Rules.equational [head] [rule] end; | 
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 1817 | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1818 | (* definition *) | 
| 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 1819 | |
| 61260 | 1820 | fun definition overloaded (alphas, binding) parent (parents: parent_info list) raw_fields thy0 = | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1821 | let | 
| 60796 | 1822 | val ctxt0 = Proof_Context.init_global thy0; | 
| 43682 | 1823 | |
| 37470 | 1824 | val prefix = Binding.name_of binding; | 
| 60796 | 1825 | val name = Sign.full_name thy0 binding; | 
| 1826 | val full = Sign.full_name_path thy0 prefix; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1827 | |
| 35136 | 1828 | val bfields = map (fn (x, T, _) => (x, T)) raw_fields; | 
| 1829 | val field_syntax = map #3 raw_fields; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1830 | |
| 32952 | 1831 | val parent_fields = maps #fields parents; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1832 | val parent_chunks = map (length o #fields) parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1833 | val parent_names = map fst parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1834 | val parent_types = map snd parent_fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1835 | val parent_fields_len = length parent_fields; | 
| 35239 | 1836 | val parent_variants = | 
| 1837 | Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names); | |
| 37470 | 1838 | val parent_vars = map2 (curry Free) parent_variants parent_types; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1839 | val parent_len = length parents; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1840 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1841 | val fields = map (apfst full) bfields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1842 | val names = map fst fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1843 | val types = map snd fields; | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 1844 | val alphas_fields = fold Term.add_tfreesT types []; | 
| 33049 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 haftmann parents: 
33039diff
changeset | 1845 | val alphas_ext = inter (op =) alphas_fields alphas; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1846 | 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 | 1847 | val variants = | 
| 32764 
690f9cccf232
replaced meta_iffD2 by existing Drule.equal_elim_rule2;
 wenzelm parents: 
32763diff
changeset | 1848 | Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) | 
| 35136 | 1849 | (map (Binding.name_of o fst) bfields); | 
| 37470 | 1850 | val vars = map2 (curry Free) variants types; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1851 | val named_vars = names ~~ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1852 | val idxms = 0 upto len; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1853 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1854 | val all_fields = parent_fields @ fields; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1855 | val all_types = parent_types @ types; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1856 | val all_variants = parent_variants @ variants; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1857 | val all_vars = parent_vars @ vars; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1858 | 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 | 1859 | |
| 67149 | 1860 | val zeta = (singleton (Name.variant_list (map #1 alphas)) "'z", \<^sort>\<open>type\<close>); | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 1861 | val moreT = TFree zeta; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1862 | val more = Free (moreN, moreT); | 
| 35136 | 1863 | val full_moreN = full (Binding.name moreN); | 
| 1864 | val bfields_more = bfields @ [(Binding.name moreN, moreT)]; | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1865 | val fields_more = fields @ [(full_moreN, moreT)]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1866 | 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 | 1867 | val all_vars_more = all_vars @ [more]; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1868 | val all_named_vars_more = all_named_vars @ [(full_moreN, more)]; | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1869 | |
| 17261 | 1870 | |
| 35239 | 1871 | (* 1st stage: ext_thy *) | 
| 1872 | ||
| 1873 | val extension_name = full binding; | |
| 1874 | ||
| 38758 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1875 | val ((ext, (ext_tyco, vs), | 
| 
f2cfb2cc03e8
misc tuning and simplification, notably theory data;
 wenzelm parents: 
38715diff
changeset | 1876 | extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) = | 
| 60796 | 1877 | thy0 | 
| 35239 | 1878 | |> Sign.qualified_path false binding | 
| 61260 | 1879 | |> extension_definition overloaded extension_name fields alphas_ext zeta moreT more vars; | 
| 54895 | 1880 | val ext_ctxt = Proof_Context.init_global ext_thy; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1881 | |
| 60796 | 1882 | val _ = timing_msg ext_ctxt "record preparing definitions"; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1883 | val Type extension_scheme = extT; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1884 | val extension_name = unsuffix ext_typeN (fst extension_scheme); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1885 | val extension = let val (n, Ts) = extension_scheme in (n, subst_last HOLogic.unitT Ts) end; | 
| 35239 | 1886 | 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 | 1887 | val extension_id = implode extension_names; | 
| 17261 | 1888 | |
| 33957 | 1889 | 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 | 1890 | val rec_schemeT0 = rec_schemeT 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1891 | |
| 17261 | 1892 | fun recT n = | 
| 32972 | 1893 | let val (c, Ts) = extension in | 
| 33957 | 1894 | mk_recordT (map #extension (drop n parents)) | 
| 32972 | 1895 | (Type (c, subst_last HOLogic.unitT Ts)) | 
| 1896 | end; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1897 | val recT0 = recT 0; | 
| 17261 | 1898 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1899 | fun mk_rec args n = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1900 | let | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1901 | val (args', more) = chop_last args; | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 1902 | 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 | 1903 | fun build Ts = | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1904 | fold_rev mk_ext' (drop n ((extension_names ~~ Ts) ~~ chunks parent_chunks args')) more; | 
| 17261 | 1905 | in | 
| 1906 | if more = HOLogic.unit | |
| 33063 | 1907 | then build (map_range recT (parent_len + 1)) | 
| 1908 | 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 | 1909 | end; | 
| 17261 | 1910 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1911 | val r_rec0 = mk_rec all_vars_more 0; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1912 | 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 | 1913 | |
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1914 | fun r n = Free (rN, rec_schemeT n); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1915 | val r0 = r 0; | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1916 | fun r_unit n = Free (rN, recT n); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1917 | val r_unit0 = r_unit 0; | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1918 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1919 | |
| 35239 | 1920 | (* print translations *) | 
| 1921 | ||
| 35149 | 1922 | val record_ext_type_abbr_tr's = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1923 | let | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1924 | val trname = hd extension_names; | 
| 35148 | 1925 | val last_ext = unsuffix ext_typeN (fst extension); | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1926 | in [record_ext_type_abbr_tr' name alphas zeta last_ext rec_schemeT0 trname] end; | 
| 35149 | 1927 | |
| 1928 | val record_ext_type_tr's = | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1929 | let | 
| 35149 | 1930 | (*avoid conflict with record_type_abbr_tr's*) | 
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35410diff
changeset | 1931 | val trnames = if parent_len > 0 then [extension_name] else []; | 
| 35149 | 1932 | in map record_ext_type_tr' trnames end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1933 | |
| 52143 | 1934 | val print_translation = | 
| 35239 | 1935 | map field_update_tr' (full_moreN :: names) @ [record_ext_tr' extension_name] @ | 
| 1936 | record_ext_type_tr's @ record_ext_type_abbr_tr's; | |
| 1937 | ||
| 17261 | 1938 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1939 | (* prepare declarations *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1940 | |
| 35136 | 1941 | val sel_decls = map (mk_selC rec_schemeT0 o apfst Binding.name_of) bfields_more; | 
| 1942 | 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 | 1943 | val make_decl = (makeN, all_types ---> recT0); | 
| 17261 | 1944 | val fields_decl = (fields_selN, types ---> Type extension); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1945 | val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1946 | val truncate_decl = (truncateN, rec_schemeT0 --> recT0); | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1947 | |
| 35133 | 1948 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 1949 | (* prepare definitions *) | 
| 17261 | 1950 | |
| 35138 | 1951 | 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 | 1952 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33957diff
changeset | 1953 | (*Theorems from the iso_tuple intros. | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1954 | By unfolding ext_defs from r_rec0 we create a tree of constructor | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1955 | calls (many of them Pair, but others as well). The introduction | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1956 | rules for update_accessor_eq_assist can unify two different ways | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1957 | on these constructors. If we take the complete result sequence of | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1958 | running a the introduction tactic, we get one theorem for each upd/acc | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1959 | pair, from which we can derive the bodies of our selector and | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1960 | updator and their convs.*) | 
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1961 | val (accessor_thms, updator_thms, upd_acc_cong_assists) = | 
| 60796 | 1962 | timeit_msg ext_ctxt "record getting tree access/updates:" (fn () => | 
| 46044 | 1963 | let | 
| 1964 | val r_rec0_Vars = | |
| 1965 | let | |
| 1966 | (*pick variable indices of 1 to avoid possible variable | |
| 1967 | collisions with existing variables in updacc_eq_triv*) | |
| 1968 | fun to_Var (Free (c, T)) = Var ((c, 1), T); | |
| 1969 | in mk_rec (map to_Var all_vars_more) 0 end; | |
| 1970 | ||
| 60796 | 1971 | val init_thm = | 
| 1972 | infer_instantiate ext_ctxt | |
| 1973 |               [(("v", 0), Thm.cterm_of ext_ctxt r_rec0),
 | |
| 1974 |                (("v'", 0), Thm.cterm_of ext_ctxt r_rec0_Vars)]
 | |
| 1975 | updacc_eq_triv; | |
| 60752 | 1976 | val terminal = | 
| 1977 | resolve_tac ext_ctxt [updacc_eq_idI] 1 THEN resolve_tac ext_ctxt [refl] 1; | |
| 46044 | 1978 | val tactic = | 
| 54895 | 1979 | simp_tac (put_simpset HOL_basic_ss ext_ctxt addsimps ext_defs) 1 THEN | 
| 59164 | 1980 | REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac ext_ctxt 1 ORELSE terminal); | 
| 46044 | 1981 | val updaccs = Seq.list_of (tactic init_thm); | 
| 1982 | in | |
| 1983 | (updaccs RL [updacc_accessor_eqE], | |
| 1984 | updaccs RL [updacc_updator_eqE], | |
| 1985 | updaccs RL [updacc_cong_from_eq]) | |
| 1986 | end); | |
| 32743 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 1987 | |
| 33957 | 1988 | 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 | 1989 | |
| 17261 | 1990 | (*selectors*) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1991 | 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 | 1992 | let | 
| 35239 | 1993 | val (acc $ arg, _) = | 
| 1994 | 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 | 1995 | val _ = | 
| 35239 | 1996 | if arg aconv r_rec0 then () | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1997 |           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 | 1998 | in | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 1999 | 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 | 2000 | end; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2001 | 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 | 2002 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2003 | (*updates*) | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2004 | fun mk_upd_spec ((c, T), thm) = | 
| 17261 | 2005 | let | 
| 35239 | 2006 | val (upd $ _ $ arg, _) = | 
| 2007 | 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 | 2008 | val _ = | 
| 35135 | 2009 | if arg aconv r_rec0 then () | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2010 |           else raise TERM ("mk_sel_spec: different arg", [arg]);
 | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2011 | 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 | 2012 | 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 | 2013 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2014 | (*derived operations*) | 
| 35144 | 2015 | val make_spec = | 
| 2016 | list_comb (Const (full (Binding.name makeN), all_types ---> recT0), all_vars) :== | |
| 2017 | mk_rec (all_vars @ [HOLogic.unit]) 0; | |
| 2018 | val fields_spec = | |
| 2019 | list_comb (Const (full (Binding.name fields_selN), types ---> Type extension), vars) :== | |
| 2020 | mk_rec (all_vars @ [HOLogic.unit]) parent_len; | |
| 17261 | 2021 | val extend_spec = | 
| 35136 | 2022 | Const (full (Binding.name extendN), recT0 --> moreT --> rec_schemeT0) $ r_unit0 $ more :== | 
| 35144 | 2023 | mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0; | 
| 2024 | val truncate_spec = | |
| 2025 | Const (full (Binding.name truncateN), rec_schemeT0 --> recT0) $ r0 :== | |
| 2026 | 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 | 2027 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2028 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2029 | (* 2st stage: defs_thy *) | 
| 17261 | 2030 | |
| 22747 
0c9c413b4678
add definitions explicitly to code generator table
 haftmann parents: 
22693diff
changeset | 2031 | val (((sel_defs, upd_defs), derived_defs), defs_thy) = | 
| 60796 | 2032 | timeit_msg ext_ctxt "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:" | 
| 46044 | 2033 | (fn () => | 
| 2034 | ext_thy | |
| 52143 | 2035 | |> Sign.print_translation print_translation | 
| 60796 | 2036 | |> Sign.restore_naming thy0 | 
| 46044 | 2037 | |> Typedecl.abbrev_global (binding, map #1 alphas, NoSyn) recT0 |> snd | 
| 2038 | |> Typedecl.abbrev_global | |
| 46056 | 2039 | (Binding.suffix_name schemeN binding, map #1 (alphas @ [zeta]), NoSyn) rec_schemeT0 | 
| 2040 | |> snd | |
| 46044 | 2041 | |> Sign.qualified_path false binding | 
| 2042 | |> fold (fn ((x, T), mx) => snd o Sign.declare_const_global ((Binding.name x, T), mx)) | |
| 2043 | (sel_decls ~~ (field_syntax @ [NoSyn])) | |
| 2044 | |> fold (fn (x, T) => snd o Sign.declare_const_global ((Binding.name x, T), NoSyn)) | |
| 2045 | (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl]) | |
| 2046 | |> (Global_Theory.add_defs false o | |
| 59859 | 2047 | map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) sel_specs | 
| 46044 | 2048 | ||>> (Global_Theory.add_defs false o | 
| 59859 | 2049 | map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) upd_specs | 
| 46044 | 2050 | ||>> (Global_Theory.add_defs false o | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 2051 | map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) | 
| 52788 | 2052 | [make_spec, fields_spec, extend_spec, truncate_spec]); | 
| 54895 | 2053 | val defs_ctxt = Proof_Context.init_global defs_thy; | 
| 17261 | 2054 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2055 | (* prepare propositions *) | 
| 60796 | 2056 | val _ = timing_msg defs_ctxt "record preparing propositions"; | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2057 | val P = Free (singleton (Name.variant_list all_variants) "P", rec_schemeT0 --> HOLogic.boolT); | 
| 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2058 | val C = Free (singleton (Name.variant_list all_variants) "C", HOLogic.boolT); | 
| 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2059 | val P_unit = Free (singleton (Name.variant_list all_variants) "P", recT0 --> HOLogic.boolT); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2060 | |
| 17261 | 2061 | (*selectors*) | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2062 | val sel_conv_props = | 
| 61861 | 2063 | 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 | 2064 | |
| 17261 | 2065 | (*updates*) | 
| 37470 | 2066 | fun mk_upd_prop i (c, T) = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2067 | let | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2068 | val x' = | 
| 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2069 | Free (singleton (Name.variant_list all_variants) (Long_Name.base_name c ^ "'"), T --> T); | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2070 | val n = parent_fields_len + i; | 
| 35239 | 2071 | 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 | 2072 | in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end; | 
| 37470 | 2073 | val upd_conv_props = map2 mk_upd_prop idxms fields_more; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2074 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2075 | (*induct*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2076 | val induct_scheme_prop = | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2077 | fold_rev Logic.all all_vars_more (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0); | 
| 17261 | 2078 | val induct_prop = | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2079 | (fold_rev Logic.all all_vars (Trueprop (P_unit $ r_rec_unit0)), Trueprop (P_unit $ r_unit0)); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2080 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2081 | (*surjective*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2082 | val surjective_prop = | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2083 | 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 | 2084 | in r0 === mk_rec args 0 end; | 
| 17261 | 2085 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2086 | (*cases*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2087 | val cases_scheme_prop = | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2088 | (fold_rev Logic.all all_vars_more ((r0 === r_rec0) ==> Trueprop C), Trueprop C); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2089 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2090 | val cases_prop = | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2091 | fold_rev Logic.all all_vars ((r_unit0 === r_rec_unit0) ==> Trueprop C) ==> Trueprop C; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2092 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2093 | (*split*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2094 | val split_meta_prop = | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2095 | let | 
| 46056 | 2096 | val P = Free (singleton (Name.variant_list all_variants) "P", rec_schemeT0 --> propT); | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42795diff
changeset | 2097 | in | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2098 | Logic.mk_equals (Logic.all r0 (P $ r0), fold_rev Logic.all all_vars_more (P $ r_rec0)) | 
| 17261 | 2099 | end; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2100 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2101 | val split_object_prop = | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2102 | 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 | 2103 | 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 | 2104 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2105 | val split_ex_prop = | 
| 32974 
2a1aaa2d9e64
eliminated old List.foldr and OldTerm operations;
 wenzelm parents: 
32973diff
changeset | 2106 | 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 | 2107 | 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 | 2108 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2109 | (*equality*) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2110 | val equality_prop = | 
| 17261 | 2111 | let | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2112 | val s' = Free (rN ^ "'", rec_schemeT0); | 
| 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2113 | 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 | 2114 | val seleqs = map mk_sel_eq all_named_vars_more; | 
| 46215 
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
 wenzelm parents: 
46186diff
changeset | 2115 | in Logic.all r0 (Logic.all s' (Logic.list_implies (seleqs, r0 === s'))) end; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2116 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2117 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2118 | (* 3rd stage: thms_thy *) | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2119 | |
| 46054 
3458b0e955ac
simplified proof -- avoid res_inst_tac, afford plain asm_full_simp_tac;
 wenzelm parents: 
46053diff
changeset | 2120 | val record_ss = get_simpset defs_thy; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2121 | val sel_upd_ctxt = | 
| 54895 | 2122 | put_simpset record_ss defs_ctxt | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2123 | addsimps (sel_defs @ accessor_thms @ upd_defs @ updator_thms); | 
| 46050 | 2124 | |
| 2125 | val (sel_convs, upd_convs) = | |
| 60796 | 2126 | timeit_msg defs_ctxt "record sel_convs/upd_convs proof:" (fn () => | 
| 46893 
d5bb4c212df1
some grouping of Par_List operations, to adjust granularity;
 wenzelm parents: 
46708diff
changeset | 2127 | grouped 10 Par_List.map (fn prop => | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2128 | Goal.prove_sorry_global defs_thy [] [] prop | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2129 | (fn _ => ALLGOALS (asm_full_simp_tac sel_upd_ctxt))) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2130 | (sel_conv_props @ upd_conv_props)) | 
| 46050 | 2131 | |> chop (length sel_conv_props); | 
| 46044 | 2132 | |
| 60796 | 2133 | val (fold_congs, unfold_congs) = timeit_msg defs_ctxt "record upd fold/unfold congs:" (fn () => | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2134 | let | 
| 36945 | 2135 | val symdefs = map Thm.symmetric (sel_defs @ upd_defs); | 
| 54895 | 2136 | val fold_ctxt = put_simpset HOL_basic_ss defs_ctxt addsimps symdefs; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2137 | val ua_congs = map (Drule.export_without_context o simplify fold_ctxt) upd_acc_cong_assists; | 
| 46044 | 2138 | in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end); | 
| 15012 
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
 schirmer parents: 
14981diff
changeset | 2139 | |
| 35138 | 2140 | 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 | 2141 | |
| 60796 | 2142 | val induct_scheme = timeit_msg defs_ctxt "record induct_scheme proof:" (fn () => | 
| 51551 | 2143 | Goal.prove_sorry_global defs_thy [] [] induct_scheme_prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2144 |         (fn {context = ctxt, ...} =>
 | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2145 | EVERY | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2146 | [case parent_induct of NONE => all_tac | SOME ind => try_param_tac ctxt rN ind 1, | 
| 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2147 | try_param_tac ctxt rN ext_induct 1, | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2148 | asm_simp_tac (put_simpset HOL_basic_ss ctxt) 1])); | 
| 46044 | 2149 | |
| 60796 | 2150 | val induct = timeit_msg defs_ctxt "record induct proof:" (fn () => | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2151 | Goal.prove_sorry_global defs_thy [] [#1 induct_prop] (#2 induct_prop) | 
| 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2152 |         (fn {context = ctxt, prems, ...} =>
 | 
| 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2153 | try_param_tac ctxt rN induct_scheme 1 | 
| 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2154 |           THEN try_param_tac ctxt "more" @{thm unit.induct} 1
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 2155 | THEN resolve_tac ctxt prems 1)); | 
| 46044 | 2156 | |
| 60796 | 2157 | val surjective = timeit_msg defs_ctxt "record surjective proof:" (fn () => | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2158 | Goal.prove_sorry_global defs_thy [] [] surjective_prop | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2159 |         (fn {context = ctxt, ...} =>
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2160 | EVERY | 
| 60752 | 2161 | [resolve_tac ctxt [surject_assist_idE] 1, | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2162 | simp_tac (put_simpset HOL_basic_ss ctxt addsimps ext_defs) 1, | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2163 | REPEAT | 
| 59164 | 2164 | (Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE | 
| 60752 | 2165 | (resolve_tac ctxt [surject_assistI] 1 THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2166 | simp_tac (put_simpset (get_sel_upd_defs defs_thy) ctxt | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2167 |                     addsimps (sel_defs @ @{thms o_assoc id_apply id_o o_id})) 1))]));
 | 
| 46044 | 2168 | |
| 60796 | 2169 | val cases_scheme = timeit_msg defs_ctxt "record cases_scheme proof:" (fn () => | 
| 51551 | 2170 | Goal.prove_sorry_global defs_thy [] [#1 cases_scheme_prop] (#2 cases_scheme_prop) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 2171 |         (fn {context = ctxt, prems, ...} =>
 | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59271diff
changeset | 2172 | resolve_tac ctxt prems 1 THEN | 
| 60752 | 2173 | resolve_tac ctxt [surjective] 1)); | 
| 46053 | 2174 | |
| 60796 | 2175 | val cases = timeit_msg defs_ctxt "record cases proof:" (fn () => | 
| 51551 | 2176 | Goal.prove_sorry_global defs_thy [] [] cases_prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2177 |         (fn {context = ctxt, ...} =>
 | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58936diff
changeset | 2178 | try_param_tac ctxt rN cases_scheme 1 THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2179 | ALLGOALS (asm_full_simp_tac | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2180 |             (put_simpset HOL_basic_ss ctxt addsimps @{thms unit_all_eq1}))));
 | 
| 46053 | 2181 | |
| 60796 | 2182 | val split_meta = timeit_msg defs_ctxt "record split_meta proof:" (fn () => | 
| 51551 | 2183 | Goal.prove_sorry_global defs_thy [] [] split_meta_prop | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 2184 |         (fn {context = ctxt', ...} =>
 | 
| 32975 | 2185 | EVERY1 | 
| 60752 | 2186 |            [resolve_tac ctxt' @{thms equal_intr_rule}, Goal.norm_hhf_tac ctxt',
 | 
| 2187 |             eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt',
 | |
| 2188 |             resolve_tac ctxt' [@{thm prop_subst} OF [surjective]],
 | |
| 2189 |             REPEAT o eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt']));
 | |
| 46044 | 2190 | |
| 60796 | 2191 | val split_object = timeit_msg defs_ctxt "record split_object proof:" (fn () => | 
| 51551 | 2192 | Goal.prove_sorry_global defs_thy [] [] split_object_prop | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 2193 |         (fn {context = ctxt, ...} =>
 | 
| 60752 | 2194 |           resolve_tac ctxt [@{lemma "Trueprop A \<equiv> Trueprop B \<Longrightarrow> A = B" by (rule iffI) unfold}] 1 THEN
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54707diff
changeset | 2195 |           rewrite_goals_tac ctxt @{thms atomize_all [symmetric]} THEN
 | 
| 60752 | 2196 | resolve_tac ctxt [split_meta] 1)); | 
| 46044 | 2197 | |
| 60796 | 2198 | val split_ex = timeit_msg defs_ctxt "record split_ex proof:" (fn () => | 
| 51551 | 2199 | Goal.prove_sorry_global defs_thy [] [] split_ex_prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2200 |         (fn {context = ctxt, ...} =>
 | 
| 46052 | 2201 | simp_tac | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2202 | (put_simpset HOL_basic_ss ctxt addsimps | 
| 60752 | 2203 |               (@{lemma "\<exists>x. P x \<equiv> \<not> (\<forall>x. \<not> P x)" by simp} ::
 | 
| 46052 | 2204 |                 @{thms not_not Not_eq_iff})) 1 THEN
 | 
| 60752 | 2205 | resolve_tac ctxt [split_object] 1)); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2206 | |
| 60796 | 2207 | val equality = timeit_msg defs_ctxt "record equality proof:" (fn () => | 
| 51551 | 2208 | Goal.prove_sorry_global defs_thy [] [] equality_prop | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2209 |         (fn {context = ctxt, ...} =>
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 2210 | asm_full_simp_tac (put_simpset record_ss ctxt addsimps (split_meta :: sel_convs)) 1)); | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2211 | |
| 46055 | 2212 | val ([(_, sel_convs'), (_, upd_convs'), (_, sel_defs'), (_, upd_defs'), | 
| 2213 | (_, fold_congs'), (_, unfold_congs'), | |
| 2214 | (_, splits' as [split_meta', split_object', split_ex']), (_, derived_defs'), | |
| 2215 | (_, [surjective']), (_, [equality']), (_, [induct_scheme']), (_, [induct']), | |
| 2216 | (_, [cases_scheme']), (_, [cases'])], thms_thy) = defs_thy | |
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
63352diff
changeset | 2217 | |> Code.declare_default_eqns_global (map (rpair true) derived_defs) | 
| 46055 | 2218 | |> Global_Theory.note_thmss "" | 
| 2219 | [((Binding.name "select_convs", []), [(sel_convs, [])]), | |
| 2220 | ((Binding.name "update_convs", []), [(upd_convs, [])]), | |
| 2221 | ((Binding.name "select_defs", []), [(sel_defs, [])]), | |
| 2222 | ((Binding.name "update_defs", []), [(upd_defs, [])]), | |
| 2223 | ((Binding.name "fold_congs", []), [(fold_congs, [])]), | |
| 2224 | ((Binding.name "unfold_congs", []), [(unfold_congs, [])]), | |
| 2225 | ((Binding.name "splits", []), [([split_meta, split_object, split_ex], [])]), | |
| 2226 | ((Binding.name "defs", []), [(derived_defs, [])]), | |
| 2227 | ((Binding.name "surjective", []), [([surjective], [])]), | |
| 2228 | ((Binding.name "equality", []), [([equality], [])]), | |
| 2229 | ((Binding.name "induct_scheme", induct_type_global (suffix schemeN name)), | |
| 2230 | [([induct_scheme], [])]), | |
| 2231 | ((Binding.name "induct", induct_type_global name), [([induct], [])]), | |
| 2232 | ((Binding.name "cases_scheme", cases_type_global (suffix schemeN name)), | |
| 2233 | [([cases_scheme], [])]), | |
| 2234 | ((Binding.name "cases", cases_type_global name), [([cases], [])])]; | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2235 | |
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2236 | 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 | 2237 | val sel_upd_defs = sel_defs' @ upd_defs'; | 
| 
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32335diff
changeset | 2238 | val depth = parent_len + 1; | 
| 35138 | 2239 | |
| 46055 | 2240 | val ([(_, simps'), (_, iffs')], thms_thy') = thms_thy | 
| 2241 | |> Global_Theory.note_thmss "" | |
| 2242 | [((Binding.name "simps", [Simplifier.simp_add]), [(sel_upd_simps, [])]), | |
| 2243 | ((Binding.name "iffs", [iff_add]), [([ext_inject], [])])]; | |
| 35138 | 2244 | |
| 2245 | val info = | |
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 2246 | make_info alphas parent fields extension | 
| 35138 | 2247 | ext_induct ext_inject ext_surjective ext_split ext_def | 
| 2248 | sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs' | |
| 2249 | surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs'; | |
| 2250 | ||
| 2251 | val final_thy = | |
| 2252 | thms_thy' | |
| 2253 | |> put_record name info | |
| 46221 
6dcb2cea827d
eliminated dead code, together with spurious warning about congruence rule for "Fun.comp";
 wenzelm parents: 
46219diff
changeset | 2254 | |> put_sel_upd names full_moreN depth sel_upd_simps sel_upd_defs | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 2255 | |> add_equalities extension_id equality' | 
| 15015 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2256 | |> add_extinjects ext_inject | 
| 
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
 schirmer parents: 
15012diff
changeset | 2257 | |> add_extsplit extension_name ext_split | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 2258 | |> add_splits extension_id (split_meta', split_object', split_ex', induct_scheme') | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2259 | |> add_extfields extension_name (fields @ [(full_moreN, moreT)]) | 
| 43685 
5c9160f420d5
clarified record syntax: fieldext excludes the "more" pseudo-field (unlike 2f885b7e5ba7), so that errors like (| x = a, more = b |) are reported less confusingly;
 wenzelm parents: 
43683diff
changeset | 2260 | |> add_fieldext (extension_name, snd extension) names | 
| 38533 
8d23c7403699
use extension constant as formal constructor of logical record type
 haftmann parents: 
38531diff
changeset | 2261 | |> add_code ext_tyco vs extT ext simps' ext_inject | 
| 58363 
a5c08cd60a3f
take out selectors for records -- for derived records, these don't quite have the right type
 blanchet parents: 
58239diff
changeset | 2262 | |> add_ctr_sugar (Const ext) cases_scheme' ext_inject sel_convs' | 
| 61861 | 2263 | |> fold add_spec_rule (sel_convs' @ upd_convs' @ derived_defs') | 
| 60796 | 2264 | |> Sign.restore_naming thy0; | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14643diff
changeset | 2265 | |
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2266 | in final_thy end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2267 | |
| 27278 
129574589734
export read_typ/cert_typ -- version with regular context operations;
 wenzelm parents: 
27239diff
changeset | 2268 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2269 | (* add_record *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2270 | |
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2271 | local | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2272 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2273 | fun read_parent NONE ctxt = (NONE, ctxt) | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2274 | | read_parent (SOME raw_T) ctxt = | 
| 42361 | 2275 | (case Proof_Context.read_typ_abbrev ctxt raw_T of | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2276 | Type (name, Ts) => (SOME (Ts, name), fold Variable.declare_typ Ts ctxt) | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2277 |       | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
 | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2278 | |
| 46990 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2279 | fun read_fields raw_fields ctxt = | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2280 | let | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2281 | val Ts = Syntax.read_typs ctxt (map (fn (_, raw_T, _) => raw_T) raw_fields); | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2282 | val fields = map2 (fn (x, _, mx) => fn T => (x, T, mx)) raw_fields Ts; | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2283 | val ctxt' = fold Variable.declare_typ Ts ctxt; | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2284 | in (fields, ctxt') end; | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2285 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2286 | in | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2287 | |
| 61260 | 2288 | fun add_record overloaded (params, binding) raw_parent raw_fields thy = | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2289 | let | 
| 42361 | 2290 | val ctxt = Proof_Context.init_global thy; | 
| 2291 | fun cert_typ T = Type.no_tvars (Proof_Context.cert_typ ctxt T) | |
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2292 | handle TYPE (msg, _, _) => error msg; | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2293 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2294 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2295 | (* specification *) | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2296 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2297 | val parent = Option.map (apfst (map cert_typ)) raw_parent | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2298 | handle ERROR msg => | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2299 |         cat_error msg ("The error(s) above occurred in parent record specification");
 | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2300 | val parent_args = (case parent of SOME (Ts, _) => Ts | NONE => []); | 
| 41577 
9a64c4007864
export Record.get_hierarchy -- external tools typically need this information;
 wenzelm parents: 
41576diff
changeset | 2301 | val parents = get_parent_info thy parent; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2302 | |
| 46990 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2303 | val bfields = | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2304 | raw_fields |> map (fn (x, raw_T, mx) => (x, cert_typ raw_T, mx) | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2305 | handle ERROR msg => | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2306 |           cat_error msg ("The error(s) above occurred in record field " ^ Binding.print x));
 | 
| 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2307 | |
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2308 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2309 | (* errors *) | 
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2310 | |
| 35239 | 2311 | val name = Sign.full_name thy binding; | 
| 17261 | 2312 | val err_dup_record = | 
| 38012 
3ca193a6ae5a
delete structure Basic_Record; avoid `record` in names in structure Record
 haftmann parents: 
37781diff
changeset | 2313 | if is_none (get_info thy name) then [] | 
| 4890 | 2314 | else ["Duplicate definition of record " ^ quote name]; | 
| 2315 | ||
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2316 | val spec_frees = fold Term.add_tfreesT (parent_args @ map #2 bfields) []; | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2317 | val err_extra_frees = | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2318 | (case subtract (op =) params spec_frees of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2319 | [] => [] | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2320 | | extras => ["Extra free type variable(s) " ^ | 
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2321 | commas (map (Syntax.string_of_typ ctxt o TFree) extras)]); | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2322 | |
| 4890 | 2323 | 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 | 2324 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2325 | val err_dup_fields = | 
| 35136 | 2326 | (case duplicates Binding.eq_name (map #1 bfields) of | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2327 | [] => [] | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 2328 | | dups => ["Duplicate field(s) " ^ commas (map Binding.print dups)]); | 
| 4890 | 2329 | |
| 2330 | val err_bad_fields = | |
| 35136 | 2331 | if forall (not_equal moreN o Binding.name_of o #1) bfields then [] | 
| 4890 | 2332 | else ["Illegal field name " ^ quote moreN]; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2333 | |
| 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2334 | val errs = | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2335 | err_dup_record @ err_extra_frees @ err_no_fields @ err_dup_fields @ err_bad_fields; | 
| 32761 
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
 wenzelm parents: 
32760diff
changeset | 2336 | 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 | 2337 | in | 
| 61260 | 2338 | thy |> definition overloaded (params, binding) parent parents bfields | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2339 | end | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 2340 |   handle ERROR msg => cat_error msg ("Failed to define record " ^ Binding.print binding);
 | 
| 35136 | 2341 | |
| 61260 | 2342 | fun add_record_cmd overloaded (raw_params, binding) raw_parent raw_fields thy = | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2343 | let | 
| 42361 | 2344 | val ctxt = Proof_Context.init_global thy; | 
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36151diff
changeset | 2345 | val params = map (apsnd (Typedecl.read_constraint ctxt)) raw_params; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36151diff
changeset | 2346 | val ctxt1 = fold (Variable.declare_typ o TFree) params ctxt; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36151diff
changeset | 2347 | val (parent, ctxt2) = read_parent raw_parent ctxt1; | 
| 46990 
63fae4a2cc65
simultaneous read_fields -- e.g. relevant for sort assignment;
 wenzelm parents: 
46961diff
changeset | 2348 | val (fields, ctxt3) = read_fields raw_fields ctxt2; | 
| 42361 | 2349 | val params' = map (Proof_Context.check_tfree ctxt3) params; | 
| 61260 | 2350 | in thy |> add_record overloaded (params', binding) parent fields end; | 
| 36151 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2351 | |
| 
b89a2a05a3ce
modernized treatment of sort constraints in specification;
 wenzelm parents: 
36137diff
changeset | 2352 | end; | 
| 4867 
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
 wenzelm parents: diff
changeset | 2353 | |
| 32335 | 2354 | |
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2355 | (* printing *) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2356 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2357 | local | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2358 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2359 | fun the_parent_recT (Type (parent, [Type (_, [unit as Type (_,[])])])) = Type (parent, [unit]) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2360 | | the_parent_recT (Type (extT, [T])) = Type (extT, [the_parent_recT T]) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2361 |   | the_parent_recT T = raise TYPE ("Not a unit record scheme with parent: ", [T], [])
 | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2362 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2363 | in | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2364 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2365 | fun pretty_recT ctxt typ = | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2366 | let | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2367 | val thy = Proof_Context.theory_of ctxt | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2368 | val (fs, (_, moreT)) = get_recT_fields thy typ | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2369 |     val _ = if moreT = HOLogic.unitT then () else raise TYPE ("Not a unit record scheme: ", [typ], [])
 | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2370 | val parent = if length (dest_recTs typ) >= 2 then SOME (the_parent_recT typ) else NONE | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2371 | val pfs = case parent of SOME p => fst (get_recT_fields thy p) | NONE => [] | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2372 | val fs' = drop (length pfs) fs | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2373 | fun pretty_field (name, typ) = Pretty.block [ | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2374 | Syntax.pretty_term ctxt (Const (name, typ)), | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2375 | Pretty.brk 1, | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2376 | Pretty.str "::", | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2377 | Pretty.brk 1, | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2378 | Syntax.pretty_typ ctxt typ | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2379 | ] | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2380 | in | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2381 | Pretty.block (Library.separate (Pretty.brk 1) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2382 | ([Pretty.keyword1 "record", Syntax.pretty_typ ctxt typ, Pretty.str "="] @ | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2383 | (case parent of | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2384 | SOME p => [Syntax.pretty_typ ctxt p, Pretty.str "+"] | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2385 | | NONE => [])) @ | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2386 | Pretty.fbrk :: | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2387 | Pretty.fbreaks (map pretty_field fs')) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2388 | end | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2389 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2390 | end | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2391 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2392 | fun string_of_record ctxt s = | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2393 | let | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2394 | val T = Syntax.read_typ ctxt s | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2395 | in | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2396 | Pretty.string_of (pretty_recT ctxt T) | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2397 |       handle TYPE _ => error ("Unknown record: " ^ Syntax.string_of_typ ctxt T)
 | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2398 | end | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2399 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2400 | val print_record = | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2401 | let | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2402 | fun print_item string_of (modes, arg) = Toplevel.keep (fn state => | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2403 | Print_Mode.with_modes modes (fn () => Output.writeln (string_of state arg)) ()); | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2404 | in print_item (string_of_record o Toplevel.context_of) end | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2405 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2406 | |
| 6358 | 2407 | (* outer syntax *) | 
| 2408 | ||
| 24867 | 2409 | val _ = | 
| 67149 | 2410 | Outer_Syntax.command \<^command_keyword>\<open>record\<close> "define extensible record" | 
| 61260 | 2411 | (Parse_Spec.overloaded -- (Parse.type_args_constrained -- Parse.binding) -- | 
| 67149 | 2412 | (\<^keyword>\<open>=\<close> |-- Scan.option (Parse.typ --| \<^keyword>\<open>+\<close>) -- | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36945diff
changeset | 2413 | Scan.repeat1 Parse.const_binding) | 
| 61260 | 2414 | >> (fn ((overloaded, x), (y, z)) => | 
| 2415 |         Toplevel.theory (add_record_cmd {overloaded = overloaded} x y z)));
 | |
| 6358 | 2416 | |
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2417 | val opt_modes = | 
| 67149 | 2418 | Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) [] | 
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2419 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2420 | val _ = | 
| 67149 | 2421 | Outer_Syntax.command \<^command_keyword>\<open>print_record\<close> "print record definiton" | 
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2422 | (opt_modes -- Parse.typ >> print_record); | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2423 | |
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61861diff
changeset | 2424 | end |