author | wenzelm |
Mon, 05 Sep 2005 17:38:18 +0200 | |
changeset 17261 | 193b84a70ca4 |
parent 17257 | 0ab67cb765da |
child 17337 | 5e3dde342840 |
permissions | -rw-r--r-- |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/record_package.ML |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
14579 | 3 |
Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
4 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
5 |
Extensible records with structural subtyping in HOL. |
5698 | 6 |
*) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
7 |
|
5698 | 8 |
signature BASIC_RECORD_PACKAGE = |
9 |
sig |
|
7178 | 10 |
val record_simproc: simproc |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
11 |
val record_eq_simproc: simproc |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
12 |
val record_upd_simproc: simproc |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
13 |
val record_split_simproc: (term -> int) -> simproc |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
14 |
val record_ex_sel_eq_simproc: simproc |
5698 | 15 |
val record_split_tac: int -> tactic |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
16 |
val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic |
5713 | 17 |
val record_split_name: string |
5698 | 18 |
val record_split_wrapper: string * wrapper |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
19 |
val print_record_type_abbr: bool ref |
17261 | 20 |
val print_record_type_as_fields: bool ref |
5698 | 21 |
end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
22 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
23 |
signature RECORD_PACKAGE = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
24 |
sig |
5698 | 25 |
include BASIC_RECORD_PACKAGE |
26 |
val quiet_mode: bool ref |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
27 |
val record_quick_and_dirty_sensitive: bool ref |
8574 | 28 |
val updateN: string |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
29 |
val ext_typeN: string |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
30 |
val last_extT: typ -> (string * typ list) option |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
31 |
val dest_recTs : typ -> (string * typ list) list |
16458 | 32 |
val get_extT_fields: theory -> typ -> ((string * typ) list * (string * typ)) |
33 |
val get_recT_fields: theory -> typ -> ((string * typ) list * (string * typ)) |
|
34 |
val get_extension: theory -> Symtab.key -> (string * typ list) option |
|
35 |
val get_extinjects: theory -> thm list |
|
36 |
val get_simpset: theory -> simpset |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
37 |
val print_records: theory -> unit |
16458 | 38 |
val add_record: string list * string -> string option -> (string * string * mixfix) list |
39 |
-> theory -> theory |
|
40 |
val add_record_i: string list * string -> (typ list * string) option |
|
41 |
-> (string * typ * mixfix) list -> theory -> theory |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
42 |
val setup: (theory -> theory) list |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
43 |
end; |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
44 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
45 |
|
17261 | 46 |
structure RecordPackage:RECORD_PACKAGE = |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
47 |
struct |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
48 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
49 |
val rec_UNIV_I = thm "rec_UNIV_I"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
50 |
val rec_True_simp = thm "rec_True_simp"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
51 |
val Pair_eq = thm "Product_Type.Pair_eq"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
52 |
val atomize_all = thm "HOL.atomize_all"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
53 |
val atomize_imp = thm "HOL.atomize_imp"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
54 |
val triv_goal = thm "triv_goal"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
55 |
val prop_subst = thm "prop_subst"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
56 |
val Pair_sel_convs = [fst_conv,snd_conv]; |
11832 | 57 |
|
58 |
||
59 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
60 |
(** name components **) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
61 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
62 |
val rN = "r"; |
15215 | 63 |
val wN = "w"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
64 |
val moreN = "more"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
65 |
val schemeN = "_scheme"; |
17261 | 66 |
val ext_typeN = "_ext_type"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
67 |
val extN ="_ext"; |
15215 | 68 |
val casesN = "_cases"; |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
69 |
val ext_dest = "_sel"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
70 |
val updateN = "_update"; |
15215 | 71 |
val updN = "_upd"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
72 |
val schemeN = "_scheme"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
73 |
val makeN = "make"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
74 |
val fields_selN = "fields"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
75 |
val extendN = "extend"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
76 |
val truncateN = "truncate"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
77 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
78 |
(*see typedef_package.ML*) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
79 |
val RepN = "Rep_"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
80 |
val AbsN = "Abs_"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
81 |
|
4894 | 82 |
(*** utilities ***) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
83 |
|
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
84 |
fun but_last xs = fst (split_last xs); |
15531 | 85 |
fun list NONE = [] |
86 |
| list (SOME x) = [x] |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
87 |
|
5698 | 88 |
(* messages *) |
89 |
||
90 |
val quiet_mode = ref false; |
|
91 |
fun message s = if ! quiet_mode then () else writeln s; |
|
92 |
||
15012
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents:
14981
diff
changeset
|
93 |
(* timing *) |
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents:
14981
diff
changeset
|
94 |
|
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents:
14981
diff
changeset
|
95 |
fun timeit_msg s x = if !timing then (warning s; timeit x) else x (); |
16379 | 96 |
fun timing_msg s = if !timing then warning s else (); |
17261 | 97 |
|
12255 | 98 |
(* syntax *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
99 |
|
12247 | 100 |
fun prune n xs = Library.drop (n, xs); |
11832 | 101 |
fun prefix_base s = NameSpace.map_base (fn bname => s ^ bname); |
102 |
||
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
103 |
val Trueprop = HOLogic.mk_Trueprop; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
104 |
fun All xs t = Term.list_all_free (xs, t); |
4894 | 105 |
|
11934
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
106 |
infix 9 $$; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
107 |
infix 0 :== ===; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
108 |
infixr 0 ==>; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
109 |
|
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
110 |
val (op $$) = Term.list_comb; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
111 |
val (op :==) = Logic.mk_defpair; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
112 |
val (op ===) = Trueprop o HOLogic.mk_eq; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
113 |
val (op ==>) = Logic.mk_implies; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
114 |
|
11832 | 115 |
(* morphisms *) |
116 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
117 |
fun mk_RepN name = suffix ext_typeN (prefix_base RepN name); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
118 |
fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name); |
11832 | 119 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
120 |
fun mk_Rep name repT absT = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
121 |
Const (suffix ext_typeN (prefix_base RepN name),absT --> repT); |
11832 | 122 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
123 |
fun mk_Abs name repT absT = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
124 |
Const (mk_AbsN name,repT --> absT); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
125 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
126 |
(* constructor *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
127 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
128 |
fun mk_extC (name,T) Ts = (suffix extN name, Ts ---> T); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
129 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
130 |
fun mk_ext (name,T) ts = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
131 |
let val Ts = map fastype_of ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
132 |
in list_comb (Const (mk_extC (name,T) Ts),ts) end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
133 |
|
15215 | 134 |
(* cases *) |
135 |
||
136 |
fun mk_casesC (name,T,vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT) |
|
137 |
||
138 |
fun mk_cases (name,T,vT) f = |
|
17261 | 139 |
let val Ts = binder_types (fastype_of f) |
15215 | 140 |
in Const (mk_casesC (name,T,vT) Ts) $ f end; |
17261 | 141 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
142 |
(* selector *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
143 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
144 |
fun mk_selC sT (c,T) = (c,sT --> T); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
145 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
146 |
fun mk_sel s (c,T) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
147 |
let val sT = fastype_of s |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
148 |
in Const (mk_selC sT (c,T)) $ s end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
149 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
150 |
(* updates *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
151 |
|
15215 | 152 |
fun mk_updC sfx sT (c,T) = (suffix sfx c, T --> sT --> sT); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
153 |
|
15215 | 154 |
fun mk_upd sfx c v s = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
155 |
let val sT = fastype_of s; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
156 |
val vT = fastype_of v; |
15215 | 157 |
in Const (mk_updC sfx sT (c, vT)) $ v $ s end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
158 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
159 |
(* types *) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
160 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
161 |
fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
162 |
(case try (unsuffix ext_typeN) c_ext_type of |
15531 | 163 |
NONE => raise TYPE ("RecordPackage.dest_recT", [typ], []) |
15570 | 164 |
| SOME c => ((c, Ts), List.last Ts)) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
165 |
| dest_recT typ = raise TYPE ("RecordPackage.dest_recT", [typ], []); |
5197 | 166 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
167 |
fun is_recT T = |
17261 | 168 |
(case try dest_recT T of NONE => false | SOME _ => true); |
11833 | 169 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
170 |
fun dest_recTs T = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
171 |
let val ((c, Ts), U) = dest_recT T |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
172 |
in (c, Ts) :: dest_recTs U |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
173 |
end handle TYPE _ => []; |
14255 | 174 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
175 |
fun last_extT T = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
176 |
let val ((c, Ts), U) = dest_recT T |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
177 |
in (case last_extT U of |
15531 | 178 |
NONE => SOME (c,Ts) |
179 |
| SOME l => SOME l) |
|
180 |
end handle TYPE _ => NONE |
|
14255 | 181 |
|
17261 | 182 |
fun rec_id i T = |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
183 |
let val rTs = dest_recTs T |
15570 | 184 |
val rTs' = if i < 0 then rTs else Library.take (i,rTs) |
185 |
in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end; |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
186 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
187 |
(*** extend theory by record definition ***) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
188 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
189 |
(** record info **) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
190 |
|
14255 | 191 |
(* type record_info and parent_info *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
192 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
193 |
type record_info = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
194 |
{args: (string * sort) list, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
195 |
parent: (typ list * string) option, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
196 |
fields: (string * typ) list, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
197 |
extension: (string * typ list), |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
198 |
induct: thm |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
199 |
}; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
200 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
201 |
fun make_record_info args parent fields extension induct = |
17261 | 202 |
{args = args, parent = parent, fields = fields, extension = extension, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
203 |
induct = induct}: record_info; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
204 |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
205 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
206 |
type parent_info = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
207 |
{name: string, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
208 |
fields: (string * typ) list, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
209 |
extension: (string * typ list), |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
210 |
induct: thm |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
211 |
}; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
212 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
213 |
fun make_parent_info name fields extension induct = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
214 |
{name = name, fields = fields, extension = extension, induct = induct}: parent_info; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
215 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
216 |
(* data kind 'HOL/record' *) |
5001 | 217 |
|
7178 | 218 |
type record_data = |
219 |
{records: record_info Symtab.table, |
|
220 |
sel_upd: |
|
221 |
{selectors: unit Symtab.table, |
|
222 |
updates: string Symtab.table, |
|
223 |
simpset: Simplifier.simpset}, |
|
14255 | 224 |
equalities: thm Symtab.table, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
225 |
extinjects: thm list, |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
226 |
extsplit: thm Symtab.table, (* maps extension name to split rule *) |
17261 | 227 |
splits: (thm*thm*thm*thm) Symtab.table, (* !!,!,EX - split-equalities,induct rule *) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
228 |
extfields: (string*typ) list Symtab.table, (* maps extension to its fields *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
229 |
fieldext: (string*typ list) Symtab.table (* maps field to its extension *) |
14255 | 230 |
}; |
7178 | 231 |
|
17261 | 232 |
fun make_record_data |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
233 |
records sel_upd equalities extinjects extsplit splits extfields fieldext = |
17261 | 234 |
{records = records, sel_upd = sel_upd, |
235 |
equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits, |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
236 |
extfields = extfields, fieldext = fieldext }: record_data; |
7178 | 237 |
|
16458 | 238 |
structure RecordsData = TheoryDataFun |
239 |
(struct |
|
17261 | 240 |
val name = "HOL/records"; |
7178 | 241 |
type T = record_data; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
242 |
|
7178 | 243 |
val empty = |
244 |
make_record_data Symtab.empty |
|
245 |
{selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss} |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
246 |
Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; |
7178 | 247 |
|
6556 | 248 |
val copy = I; |
16458 | 249 |
val extend = I; |
250 |
fun merge _ |
|
7178 | 251 |
({records = recs1, |
252 |
sel_upd = {selectors = sels1, updates = upds1, simpset = ss1}, |
|
14255 | 253 |
equalities = equalities1, |
17261 | 254 |
extinjects = extinjects1, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
255 |
extsplit = extsplit1, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
256 |
splits = splits1, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
257 |
extfields = extfields1, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
258 |
fieldext = fieldext1}, |
7178 | 259 |
{records = recs2, |
260 |
sel_upd = {selectors = sels2, updates = upds2, simpset = ss2}, |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
261 |
equalities = equalities2, |
17261 | 262 |
extinjects = extinjects2, |
263 |
extsplit = extsplit2, |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
264 |
splits = splits2, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
265 |
extfields = extfields2, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
266 |
fieldext = fieldext2}) = |
17261 | 267 |
make_record_data |
7178 | 268 |
(Symtab.merge (K true) (recs1, recs2)) |
269 |
{selectors = Symtab.merge (K true) (sels1, sels2), |
|
270 |
updates = Symtab.merge (K true) (upds1, upds2), |
|
271 |
simpset = Simplifier.merge_ss (ss1, ss2)} |
|
14255 | 272 |
(Symtab.merge Thm.eq_thm (equalities1, equalities2)) |
15248
b436486091a6
record_split_simp_tac now can get simp rules as parameter
schirmer
parents:
15215
diff
changeset
|
273 |
(gen_merge_lists Thm.eq_thm extinjects1 extinjects2) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
274 |
(Symtab.merge Thm.eq_thm (extsplit1,extsplit2)) |
17261 | 275 |
(Symtab.merge (fn ((a,b,c,d),(w,x,y,z)) |
276 |
=> Thm.eq_thm (a,w) andalso Thm.eq_thm (b,x) andalso |
|
277 |
Thm.eq_thm (c,y) andalso Thm.eq_thm (d,z)) |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
278 |
(splits1, splits2)) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
279 |
(Symtab.merge (K true) (extfields1,extfields2)) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
280 |
(Symtab.merge (K true) (fieldext1,fieldext2)); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
281 |
|
7178 | 282 |
fun print sg ({records = recs, ...}: record_data) = |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
283 |
let |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
284 |
val prt_typ = Sign.pretty_typ sg; |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
285 |
|
15531 | 286 |
fun pretty_parent NONE = [] |
287 |
| pretty_parent (SOME (Ts, name)) = |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
288 |
[Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
289 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
290 |
fun pretty_field (c, T) = Pretty.block |
16364 | 291 |
[Pretty.str (Sign.extern_const sg c), Pretty.str " ::", |
12129 | 292 |
Pretty.brk 1, Pretty.quote (prt_typ T)]; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
293 |
|
12247 | 294 |
fun pretty_record (name, {args, parent, fields, ...}: record_info) = |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
295 |
Pretty.block (Pretty.fbreaks (Pretty.block |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
296 |
[prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
297 |
pretty_parent parent @ map pretty_field fields)); |
12129 | 298 |
in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end; |
16458 | 299 |
end); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
300 |
|
5006 | 301 |
val print_records = RecordsData.print; |
302 |
||
16458 | 303 |
|
7178 | 304 |
(* access 'records' *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
305 |
|
17261 | 306 |
val get_record = Symtab.curried_lookup o #records o RecordsData.get; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
307 |
|
4890 | 308 |
fun put_record name info thy = |
7178 | 309 |
let |
17261 | 310 |
val {records, sel_upd, equalities, extinjects,extsplit,splits,extfields,fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
311 |
RecordsData.get thy; |
17261 | 312 |
val data = make_record_data (Symtab.curried_update (name, info) records) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
313 |
sel_upd equalities extinjects extsplit splits extfields fieldext; |
7178 | 314 |
in RecordsData.put data thy end; |
315 |
||
316 |
(* access 'sel_upd' *) |
|
317 |
||
16458 | 318 |
val get_sel_upd = #sel_upd o RecordsData.get; |
7178 | 319 |
|
17261 | 320 |
val get_selectors = Symtab.curried_lookup o #selectors o get_sel_upd; |
321 |
fun is_selector sg name = is_some (get_selectors sg (Sign.intern_const sg name)); |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
322 |
|
17261 | 323 |
val get_updates = Symtab.curried_lookup o #updates o get_sel_upd; |
324 |
val get_simpset = #simpset o get_sel_upd; |
|
7178 | 325 |
|
326 |
fun put_sel_upd names simps thy = |
|
327 |
let |
|
328 |
val sels = map (rpair ()) names; |
|
329 |
val upds = map (suffix updateN) names ~~ names; |
|
330 |
||
17261 | 331 |
val {records, sel_upd = {selectors, updates, simpset}, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
332 |
equalities, extinjects, extsplit, splits, extfields,fieldext} = RecordsData.get thy; |
7178 | 333 |
val data = make_record_data records |
334 |
{selectors = Symtab.extend (selectors, sels), |
|
335 |
updates = Symtab.extend (updates, upds), |
|
336 |
simpset = Simplifier.addsimps (simpset, simps)} |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
337 |
equalities extinjects extsplit splits extfields fieldext; |
7178 | 338 |
in RecordsData.put data thy end; |
339 |
||
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
340 |
(* access 'equalities' *) |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
341 |
|
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
342 |
fun add_record_equalities name thm thy = |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
343 |
let |
17261 | 344 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
345 |
RecordsData.get thy; |
17261 | 346 |
val data = make_record_data records sel_upd |
347 |
(Symtab.curried_update_new (name, thm) equalities) extinjects extsplit |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
348 |
splits extfields fieldext; |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
349 |
in RecordsData.put data thy end; |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
350 |
|
17261 | 351 |
val get_equalities =Symtab.curried_lookup o #equalities o RecordsData.get; |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
352 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
353 |
(* access 'extinjects' *) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
354 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
355 |
fun add_extinjects thm thy = |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
356 |
let |
17261 | 357 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
358 |
RecordsData.get thy; |
17261 | 359 |
val data = make_record_data records sel_upd equalities (extinjects@[thm]) extsplit |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
360 |
splits extfields fieldext; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
361 |
in RecordsData.put data thy end; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
362 |
|
16458 | 363 |
fun get_extinjects sg = #extinjects (RecordsData.get sg); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
364 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
365 |
(* access 'extsplit' *) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
366 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
367 |
fun add_extsplit name thm thy = |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
368 |
let |
17261 | 369 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
370 |
RecordsData.get thy; |
17261 | 371 |
val data = make_record_data records sel_upd |
372 |
equalities extinjects (Symtab.curried_update_new (name, thm) extsplit) splits |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
373 |
extfields fieldext; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
374 |
in RecordsData.put data thy end; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
375 |
|
17261 | 376 |
val get_extsplit = Symtab.curried_lookup o #extsplit o RecordsData.get; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
377 |
|
14255 | 378 |
(* access 'splits' *) |
379 |
||
380 |
fun add_record_splits name thmP thy = |
|
381 |
let |
|
17261 | 382 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
383 |
RecordsData.get thy; |
17261 | 384 |
val data = make_record_data records sel_upd |
385 |
equalities extinjects extsplit (Symtab.curried_update_new (name, thmP) splits) |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
386 |
extfields fieldext; |
14255 | 387 |
in RecordsData.put data thy end; |
388 |
||
17261 | 389 |
val get_splits = Symtab.curried_lookup o #splits o RecordsData.get; |
14255 | 390 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
391 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
392 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
393 |
(* extension of a record name *) |
17261 | 394 |
val get_extension = |
395 |
Option.map #extension oo (Symtab.curried_lookup o #records o RecordsData.get); |
|
396 |
||
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
397 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
398 |
(* access 'extfields' *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
399 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
400 |
fun add_extfields name fields thy = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
401 |
let |
17261 | 402 |
val {records, sel_upd, equalities, extinjects, extsplit,splits, extfields, fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
403 |
RecordsData.get thy; |
17261 | 404 |
val data = make_record_data records sel_upd |
405 |
equalities extinjects extsplit splits |
|
406 |
(Symtab.curried_update_new (name, fields) extfields) fieldext; |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
407 |
in RecordsData.put data thy end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
408 |
|
17261 | 409 |
val get_extfields = Symtab.curried_lookup o #extfields o RecordsData.get; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
410 |
|
17261 | 411 |
fun get_extT_fields sg T = |
15059 | 412 |
let |
413 |
val ((name,Ts),moreT) = dest_recT T; |
|
17261 | 414 |
val recname = let val (nm::recn::rst) = rev (NameSpace.unpack name) |
15059 | 415 |
in NameSpace.pack (rev (nm::rst)) end; |
416 |
val midx = maxidx_of_typs (moreT::Ts); |
|
417 |
fun varify (a, S) = TVar ((a, midx), S); |
|
418 |
val varifyT = map_type_tfree varify; |
|
16458 | 419 |
val {records,extfields,...} = RecordsData.get sg; |
17261 | 420 |
val (flds,(more,_)) = split_last (Symtab.curried_lookup_multi extfields name); |
421 |
val args = map varifyT (snd (#extension (the (Symtab.curried_lookup records recname)))); |
|
15058 | 422 |
|
16934 | 423 |
val (subst,_) = fold (Sign.typ_unify sg) (but_last args ~~ but_last Ts) (Vartab.empty,0); |
15059 | 424 |
val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds; |
425 |
in (flds',(more,moreT)) end; |
|
15058 | 426 |
|
17261 | 427 |
fun get_recT_fields sg T = |
428 |
let |
|
15059 | 429 |
val (root_flds,(root_more,root_moreT)) = get_extT_fields sg T; |
17261 | 430 |
val (rest_flds,rest_more) = |
431 |
if is_recT root_moreT then get_recT_fields sg root_moreT |
|
15059 | 432 |
else ([],(root_more,root_moreT)); |
433 |
in (root_flds@rest_flds,rest_more) end; |
|
434 |
||
15058 | 435 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
436 |
(* access 'fieldext' *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
437 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
438 |
fun add_fieldext extname_types fields thy = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
439 |
let |
17261 | 440 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
441 |
RecordsData.get thy; |
17261 | 442 |
val fieldext' = |
443 |
fold (fn field => Symtab.curried_update_new (field, extname_types)) fields fieldext; |
|
444 |
val data=make_record_data records sel_upd equalities extinjects extsplit |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
445 |
splits extfields fieldext'; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
446 |
in RecordsData.put data thy end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
447 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
448 |
|
17261 | 449 |
val get_fieldext = Symtab.curried_lookup o #fieldext o RecordsData.get; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
450 |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
451 |
(* parent records *) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
452 |
|
15531 | 453 |
fun add_parents thy NONE parents = parents |
454 |
| add_parents thy (SOME (types, name)) parents = |
|
12247 | 455 |
let |
456 |
val sign = Theory.sign_of thy; |
|
457 |
fun err msg = error (msg ^ " parent record " ^ quote name); |
|
12255 | 458 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
459 |
val {args, parent, fields, extension, induct} = |
15531 | 460 |
(case get_record thy name of SOME info => info | NONE => err "Unknown"); |
12247 | 461 |
val _ = if length types <> length args then err "Bad number of arguments for" else (); |
12255 | 462 |
|
12247 | 463 |
fun bad_inst ((x, S), T) = |
15531 | 464 |
if Sign.of_sort sign (T, S) then NONE else SOME x |
15570 | 465 |
val bads = List.mapPartial bad_inst (args ~~ types); |
12255 | 466 |
|
12247 | 467 |
val inst = map fst args ~~ types; |
15570 | 468 |
val subst = Term.map_type_tfree (fn (x, _) => valOf (assoc (inst, x))); |
469 |
val parent' = Option.map (apfst (map subst)) parent; |
|
12247 | 470 |
val fields' = map (apsnd subst) fields; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
471 |
val extension' = apsnd (map subst) extension; |
12247 | 472 |
in |
12255 | 473 |
conditional (not (null bads)) (fn () => |
474 |
err ("Ill-sorted instantiation of " ^ commas bads ^ " in")); |
|
475 |
add_parents thy parent' |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
476 |
(make_parent_info name fields' extension' induct::parents) |
12247 | 477 |
end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
478 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
479 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
480 |
(** concrete syntax for records **) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
481 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
482 |
(* parse translations *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
483 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
484 |
fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
485 |
if c = mark then Syntax.const (suffix sfx name) $ arg |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
486 |
else raise TERM ("gen_field_tr: " ^ mark, [t]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
487 |
| gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
488 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
489 |
fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
490 |
if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
491 |
else [gen_field_tr mark sfx tm] |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
492 |
| gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
493 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
494 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
495 |
fun record_update_tr [t, u] = |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
496 |
foldr (op $) t (rev (gen_fields_tr "_updates" "_update" updateN u)) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
497 |
| record_update_tr ts = raise TERM ("record_update_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
498 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
499 |
fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
500 |
| update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
501 |
| update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
502 |
(c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
503 |
| update_name_tr ts = raise TERM ("update_name_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
504 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
505 |
fun dest_ext_field mark (t as (Const (c,_) $ Const (name,_) $ arg)) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
506 |
if c = mark then (name,arg) else raise TERM ("dest_ext_field: " ^ mark, [t]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
507 |
| dest_ext_field _ t = raise TERM ("dest_ext_field", [t]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
508 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
509 |
fun dest_ext_fields sep mark (trm as (Const (c,_) $ t $ u)) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
510 |
if c = sep then dest_ext_field mark t::dest_ext_fields sep mark u |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
511 |
else [dest_ext_field mark trm] |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
512 |
| dest_ext_fields _ mark t = [dest_ext_field mark t] |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
513 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
514 |
fun gen_ext_fields_tr sep mark sfx more sg t = |
17261 | 515 |
let |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
516 |
val msg = "error in record input: "; |
17261 | 517 |
val fieldargs = dest_ext_fields sep mark t; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
518 |
fun splitargs (field::fields) ((name,arg)::fargs) = |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
519 |
if can (unsuffix name) field |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
520 |
then let val (args,rest) = splitargs fields fargs |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
521 |
in (arg::args,rest) end |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
522 |
else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
523 |
| splitargs [] (fargs as (_::_)) = ([],fargs) |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
524 |
| splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t]) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
525 |
| splitargs _ _ = ([],[]); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
526 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
527 |
fun mk_ext (fargs as (name,arg)::_) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
528 |
(case get_fieldext sg (Sign.intern_const sg name) of |
15531 | 529 |
SOME (ext,_) => (case get_extfields sg ext of |
17261 | 530 |
SOME flds |
531 |
=> let val (args,rest) = |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
532 |
splitargs (map fst (but_last flds)) fargs; |
17261 | 533 |
val more' = mk_ext rest; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
534 |
in list_comb (Syntax.const (suffix sfx ext),args@[more']) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
535 |
end |
15531 | 536 |
| NONE => raise TERM(msg ^ "no fields defined for " |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
537 |
^ ext,[t])) |
15531 | 538 |
| NONE => raise TERM (msg ^ name ^" is no proper field",[t])) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
539 |
| mk_ext [] = more |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
540 |
|
17261 | 541 |
in mk_ext fieldargs end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
542 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
543 |
fun gen_ext_type_tr sep mark sfx more sg t = |
17261 | 544 |
let |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
545 |
val msg = "error in record-type input: "; |
17261 | 546 |
val fieldargs = dest_ext_fields sep mark t; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
547 |
fun splitargs (field::fields) ((name,arg)::fargs) = |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
548 |
if can (unsuffix name) field |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
549 |
then let val (args,rest) = splitargs fields fargs |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
550 |
in (arg::args,rest) end |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
551 |
else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
552 |
| splitargs [] (fargs as (_::_)) = ([],fargs) |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
553 |
| splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t]) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
554 |
| splitargs _ _ = ([],[]); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
555 |
|
17261 | 556 |
fun get_sort xs n = (case assoc (xs,n) of |
557 |
SOME s => s |
|
15531 | 558 |
| NONE => Sign.defaultS sg); |
17261 | 559 |
|
560 |
||
16934 | 561 |
fun to_type t = Sign.certify_typ sg |
17261 | 562 |
(Sign.intern_typ sg |
15957 | 563 |
(Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t)); |
16934 | 564 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
565 |
fun mk_ext (fargs as (name,arg)::_) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
566 |
(case get_fieldext sg (Sign.intern_const sg name) of |
17261 | 567 |
SOME (ext,alphas) => |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
568 |
(case get_extfields sg ext of |
17261 | 569 |
SOME flds |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
570 |
=> (let |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
571 |
val flds' = but_last flds; |
17261 | 572 |
val types = map snd flds'; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
573 |
val (args,rest) = splitargs (map fst flds') fargs; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
574 |
val vartypes = map Type.varifyT types; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
575 |
val argtypes = map to_type args; |
16934 | 576 |
val (subst,_) = fold (Sign.typ_unify sg) (vartypes ~~ argtypes) |
577 |
(Vartab.empty,0); |
|
17261 | 578 |
val alphas' = map ((Syntax.term_of_typ (! Syntax.show_sorts)) o |
579 |
Envir.norm_type subst o Type.varifyT) |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
580 |
(but_last alphas); |
17261 | 581 |
|
582 |
val more' = mk_ext rest; |
|
583 |
in list_comb (Syntax.const (suffix sfx ext),alphas'@[more']) |
|
584 |
end handle TUNIFY => raise |
|
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
585 |
TERM (msg ^ "type is no proper record (extension)", [t])) |
15531 | 586 |
| NONE => raise TERM (msg ^ "no fields defined for " ^ ext,[t])) |
587 |
| NONE => raise TERM (msg ^ name ^" is no proper field",[t])) |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
588 |
| mk_ext [] = more |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
589 |
|
17261 | 590 |
in mk_ext fieldargs end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
591 |
|
17261 | 592 |
fun gen_adv_record_tr sep mark sfx unit sg [t] = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
593 |
gen_ext_fields_tr sep mark sfx unit sg t |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
594 |
| gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
595 |
|
17261 | 596 |
fun gen_adv_record_scheme_tr sep mark sfx sg [t, more] = |
597 |
gen_ext_fields_tr sep mark sfx more sg t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
598 |
| gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
599 |
|
17261 | 600 |
fun gen_adv_record_type_tr sep mark sfx unit sg [t] = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
601 |
gen_ext_type_tr sep mark sfx unit sg t |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
602 |
| gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
603 |
|
17261 | 604 |
fun gen_adv_record_type_scheme_tr sep mark sfx sg [t, more] = |
605 |
gen_ext_type_tr sep mark sfx more sg t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
606 |
| gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
607 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
608 |
val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
609 |
val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
610 |
|
17261 | 611 |
val adv_record_type_tr = |
612 |
gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
613 |
(Syntax.term_of_typ false (HOLogic.unitT)); |
17261 | 614 |
val adv_record_type_scheme_tr = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
615 |
gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
616 |
|
15215 | 617 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
618 |
val parse_translation = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
619 |
[("_record_update", record_update_tr), |
17261 | 620 |
("_update_name", update_name_tr)]; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
621 |
|
17261 | 622 |
val adv_parse_translation = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
623 |
[("_record",adv_record_tr), |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
624 |
("_record_scheme",adv_record_scheme_tr), |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
625 |
("_record_type",adv_record_type_tr), |
17261 | 626 |
("_record_type_scheme",adv_record_type_scheme_tr)]; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
627 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
628 |
(* print translations *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
629 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
630 |
val print_record_type_abbr = ref true; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
631 |
val print_record_type_as_fields = ref true; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
632 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
633 |
fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ t $ u) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
634 |
(case try (unsuffix sfx) name_field of |
15531 | 635 |
SOME name => |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
636 |
apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u) |
15531 | 637 |
| NONE => ([], tm)) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
638 |
| gen_field_upds_tr' _ _ tm = ([], tm); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
639 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
640 |
fun record_update_tr' tm = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
641 |
let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
642 |
Syntax.const "_record_update" $ u $ |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
643 |
foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
644 |
end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
645 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
646 |
fun gen_field_tr' sfx tr' name = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
647 |
let val name_sfx = suffix sfx name |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
648 |
in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
649 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
650 |
fun record_tr' sep mark record record_scheme unit sg t = |
17261 | 651 |
let |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
652 |
fun field_lst t = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
653 |
(case strip_comb t of |
17261 | 654 |
(Const (ext,_),args) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
655 |
=> (case try (unsuffix extN) (Sign.intern_const sg ext) of |
17261 | 656 |
SOME ext' |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
657 |
=> (case get_extfields sg ext' of |
17261 | 658 |
SOME flds |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
659 |
=> (let |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
660 |
val (f::fs) = but_last (map fst flds); |
17261 | 661 |
val flds' = Sign.extern_const sg f :: map NameSpace.base fs; |
662 |
val (args',more) = split_last args; |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
663 |
in (flds'~~args')@field_lst more end |
15570 | 664 |
handle UnequalLengths => [("",t)]) |
15531 | 665 |
| NONE => [("",t)]) |
666 |
| NONE => [("",t)]) |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
667 |
| _ => [("",t)]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
668 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
669 |
val (flds,(_,more)) = split_last (field_lst t); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
670 |
val flds' = map (fn (n,t)=>Syntax.const mark$Syntax.const n$t) flds; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
671 |
val flds'' = foldr1 (fn (x,y) => Syntax.const sep$x$y) flds'; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
672 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
673 |
in if null flds then raise Match |
17261 | 674 |
else if unit more |
675 |
then Syntax.const record$flds'' |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
676 |
else Syntax.const record_scheme$flds''$more |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
677 |
end |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
678 |
|
17261 | 679 |
fun gen_record_tr' name = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
680 |
let val name_sfx = suffix extN name; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
681 |
val unit = (fn Const ("Unity",_) => true | _ => false); |
17261 | 682 |
fun tr' sg ts = record_tr' "_fields" "_field" "_record" "_record_scheme" unit sg |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
683 |
(list_comb (Syntax.const name_sfx,ts)) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
684 |
in (name_sfx,tr') |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
685 |
end |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
686 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
687 |
fun print_translation names = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
688 |
map (gen_field_tr' updateN record_update_tr') names; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
689 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
690 |
(* record_type_abbr_tr' tries to reconstruct the record name type abbreviation from *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
691 |
(* the (nested) extension types. *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
692 |
fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT sg tm = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
693 |
let |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
694 |
(* tm is term representation of a (nested) field type. We first reconstruct the *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
695 |
(* type from tm so that we can continue on the type level rather then the term level.*) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
696 |
|
17261 | 697 |
fun get_sort xs n = (case assoc (xs,n) of |
698 |
SOME s => s |
|
15531 | 699 |
| NONE => Sign.defaultS sg); |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
700 |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
701 |
(* WORKAROUND: |
17261 | 702 |
* If a record type occurs in an error message of type inference there |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
703 |
* may be some internal frees donoted by ??: |
17261 | 704 |
* (Const "_tfree",_)$Free ("??'a",_). |
705 |
||
706 |
* This will unfortunately be translated to Type ("??'a",[]) instead of |
|
707 |
* TFree ("??'a",_) by typ_of_term, which will confuse unify below. |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
708 |
* fixT works around. |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
709 |
*) |
17261 | 710 |
fun fixT (T as Type (x,[])) = |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
711 |
if String.isPrefix "??'" x then TFree (x,Sign.defaultS sg) else T |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
712 |
| fixT (Type (x,xs)) = Type (x,map fixT xs) |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
713 |
| fixT T = T; |
17261 | 714 |
|
715 |
val T = fixT (Sign.intern_typ sg |
|
716 |
(Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts tm)) I tm)); |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
717 |
|
17261 | 718 |
fun mk_type_abbr subst name alphas = |
16281 | 719 |
let val abbrT = Type (name, map (fn a => TVar ((a, 0), Sign.defaultS sg)) alphas); |
17261 | 720 |
in Syntax.term_of_typ (! Syntax.show_sorts) |
721 |
(Sign.extern_typ sg (Envir.norm_type subst abbrT)) end; |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
722 |
|
16934 | 723 |
fun unify rT T = fst (Sign.typ_unify sg (Type.varifyT rT,T) (Vartab.empty,0)); |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
724 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
725 |
in if !print_record_type_abbr |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
726 |
then (case last_extT T of |
17261 | 727 |
SOME (name,_) |
728 |
=> if name = lastExt |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
729 |
then |
17261 | 730 |
(let |
731 |
val subst = unify schemeT T |
|
732 |
in |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
733 |
if HOLogic.is_unitT (Envir.norm_type subst (TVar((zeta,0),Sign.defaultS sg))) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
734 |
then mk_type_abbr subst abbr alphas |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
735 |
else mk_type_abbr subst (suffix schemeN abbr) (alphas@[zeta]) |
17261 | 736 |
end handle TUNIFY => default_tr' sg tm) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
737 |
else raise Match (* give print translation of specialised record a chance *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
738 |
| _ => raise Match) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
739 |
else default_tr' sg tm |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
740 |
end |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
741 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
742 |
fun record_type_tr' sep mark record record_scheme sg t = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
743 |
let |
17261 | 744 |
fun get_sort xs n = (case assoc (xs,n) of |
745 |
SOME s => s |
|
15531 | 746 |
| NONE => Sign.defaultS sg); |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
747 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
748 |
val T = Sign.intern_typ sg (Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
749 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
750 |
fun term_of_type T = Syntax.term_of_typ (!Syntax.show_sorts) (Sign.extern_typ sg T); |
17261 | 751 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
752 |
fun field_lst T = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
753 |
(case T of |
17261 | 754 |
Type (ext,args) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
755 |
=> (case try (unsuffix ext_typeN) ext of |
17261 | 756 |
SOME ext' |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
757 |
=> (case get_extfields sg ext' of |
17261 | 758 |
SOME flds |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
759 |
=> (case get_fieldext sg (fst (hd flds)) of |
17261 | 760 |
SOME (_,alphas) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
761 |
=> (let |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
762 |
val (f::fs) = but_last flds; |
16364 | 763 |
val flds' = apfst (Sign.extern_const sg) f |
17261 | 764 |
::map (apfst NameSpace.base) fs; |
765 |
val (args',more) = split_last args; |
|
766 |
val alphavars = map Type.varifyT (but_last alphas); |
|
16934 | 767 |
val (subst,_)= fold (Sign.typ_unify sg) (alphavars~~args') |
768 |
(Vartab.empty,0); |
|
769 |
val flds'' =map (apsnd (Envir.norm_type subst o Type.varifyT)) |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
770 |
flds'; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
771 |
in flds''@field_lst more end |
17261 | 772 |
handle TUNIFY => [("",T)] |
15570 | 773 |
| UnequalLengths => [("",T)]) |
15531 | 774 |
| NONE => [("",T)]) |
775 |
| NONE => [("",T)]) |
|
17261 | 776 |
| NONE => [("",T)]) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
777 |
| _ => [("",T)]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
778 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
779 |
val (flds,(_,moreT)) = split_last (field_lst T); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
780 |
val flds' = map (fn (n,T)=>Syntax.const mark$Syntax.const n$term_of_type T) flds; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
781 |
val flds'' = foldr1 (fn (x,y) => Syntax.const sep$x$y) flds'; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
782 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
783 |
in if not (!print_record_type_as_fields) orelse null flds then raise Match |
17261 | 784 |
else if moreT = HOLogic.unitT |
785 |
then Syntax.const record$flds'' |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
786 |
else Syntax.const record_scheme$flds''$term_of_type moreT |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
787 |
end |
17261 | 788 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
789 |
|
17261 | 790 |
fun gen_record_type_tr' name = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
791 |
let val name_sfx = suffix ext_typeN name; |
17261 | 792 |
fun tr' sg ts = record_type_tr' "_field_types" "_field_type" |
793 |
"_record_type" "_record_type_scheme" sg |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
794 |
(list_comb (Syntax.const name_sfx,ts)) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
795 |
in (name_sfx,tr') |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
796 |
end |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
797 |
|
17261 | 798 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
799 |
fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
800 |
let val name_sfx = suffix ext_typeN name; |
17261 | 801 |
val default_tr' = record_type_tr' "_field_types" "_field_type" |
802 |
"_record_type" "_record_type_scheme" |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
803 |
fun tr' sg ts = record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT sg |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
804 |
(list_comb (Syntax.const name_sfx,ts)) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
805 |
in (name_sfx, tr') end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
806 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
807 |
(** record simprocs **) |
14358 | 808 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
809 |
val record_quick_and_dirty_sensitive = ref false; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
810 |
|
15215 | 811 |
|
812 |
fun quick_and_dirty_prove stndrd sg asms prop tac = |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
813 |
if !record_quick_and_dirty_sensitive andalso !quick_and_dirty |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
814 |
then Tactic.prove sg [] [] (Logic.list_implies (map Logic.varify asms,Logic.varify prop)) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
815 |
(K (SkipProof.cheat_tac HOL.thy)) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
816 |
(* standard can take quite a while for large records, thats why |
17261 | 817 |
* we varify the proposition manually here.*) |
15215 | 818 |
else let val prf = Tactic.prove sg [] asms prop tac; |
17261 | 819 |
in if stndrd then standard prf else prf end; |
15215 | 820 |
|
17261 | 821 |
fun quick_and_dirty_prf noopt opt () = |
822 |
if !record_quick_and_dirty_sensitive andalso !quick_and_dirty |
|
15215 | 823 |
then noopt () |
824 |
else opt (); |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
825 |
|
14255 | 826 |
|
16973 | 827 |
fun prove_split_simp sg ss T prop = |
17261 | 828 |
let |
16458 | 829 |
val {sel_upd={simpset,...},extsplit,...} = RecordsData.get sg; |
17261 | 830 |
val extsplits = |
831 |
Library.foldl (fn (thms,(n,_)) => list (Symtab.curried_lookup extsplit n) @ thms) |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
832 |
([],dest_recTs T); |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
833 |
val thms = (case get_splits sg (rec_id (~1) T) of |
17261 | 834 |
SOME (all_thm,_,_,_) => |
15203 | 835 |
all_thm::(case extsplits of [thm] => [] | _ => extsplits) |
836 |
(* [thm] is the same as all_thm *) |
|
17261 | 837 |
| NONE => extsplits) |
16973 | 838 |
in |
839 |
quick_and_dirty_prove true sg [] prop |
|
840 |
(fn _ => simp_tac (Simplifier.inherit_bounds ss simpset addsimps thms) 1) |
|
15203 | 841 |
end; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
842 |
|
15215 | 843 |
|
15059 | 844 |
local |
16822 | 845 |
fun eq (s1:string) (s2:string) = (s1 = s2); |
846 |
fun has_field extfields f T = |
|
17261 | 847 |
exists (fn (eN,_) => exists (eq f o fst) (Symtab.curried_lookup_multi extfields eN)) |
16822 | 848 |
(dest_recTs T); |
15059 | 849 |
in |
14255 | 850 |
(* record_simproc *) |
851 |
(* Simplifies selections of an record update: |
|
852 |
* (1) S (r(|S:=k|)) = k respectively |
|
853 |
* (2) S (r(|X:=k|)) = S r |
|
854 |
* The simproc skips multiple updates at once, eg: |
|
855 |
* S (r (|S:=k,X:=2,Y:=3|)) = k |
|
856 |
* But be careful in (2) because of the extendibility of records. |
|
857 |
* - If S is a more-selector we have to make sure that the update on component |
|
858 |
* X does not affect the selected subrecord. |
|
859 |
* - If X is a more-selector we have to make sure that S is not in the updated |
|
17261 | 860 |
* subrecord. |
14255 | 861 |
*) |
13462 | 862 |
val record_simproc = |
863 |
Simplifier.simproc (Theory.sign_of HOL.thy) "record_simp" ["s (u k r)"] |
|
16973 | 864 |
(fn sg => fn ss => fn t => |
16872 | 865 |
(case t of (sel as Const (s, Type (_,[domS,rangeS])))$ |
866 |
((upd as Const (u,Type(_,[_,Type (_,[rT,_])]))) $ k $ r)=> |
|
15531 | 867 |
(case get_selectors sg s of SOME () => |
868 |
(case get_updates sg u of SOME u_name => |
|
13462 | 869 |
let |
16458 | 870 |
val {sel_upd={updates,...},extfields,...} = RecordsData.get sg; |
17261 | 871 |
|
16872 | 872 |
fun mk_eq_terms ((upd as Const (u,Type(_,[kT,_]))) $ k $ r) = |
17261 | 873 |
(case Symtab.curried_lookup updates u of |
15531 | 874 |
NONE => NONE |
17261 | 875 |
| SOME u_name |
14255 | 876 |
=> if u_name = s |
17261 | 877 |
then let |
16872 | 878 |
val rv = ("r",rT) |
14255 | 879 |
val rb = Bound 0 |
16872 | 880 |
val kv = ("k",kT) |
17261 | 881 |
val kb = Bound 1 |
15531 | 882 |
in SOME (upd$kb$rb,kb,[kv,rv],true) end |
16822 | 883 |
else if has_field extfields u_name rangeS |
16872 | 884 |
orelse has_field extfields s kT |
15531 | 885 |
then NONE |
17261 | 886 |
else (case mk_eq_terms r of |
887 |
SOME (trm,trm',vars,update_s) |
|
888 |
=> let |
|
889 |
val kv = ("k",kT) |
|
14255 | 890 |
val kb = Bound (length vars) |
17261 | 891 |
in SOME (upd$kb$trm,trm',kv::vars,update_s) end |
15531 | 892 |
| NONE |
17261 | 893 |
=> let |
894 |
val rv = ("r",rT) |
|
14255 | 895 |
val rb = Bound 0 |
16872 | 896 |
val kv = ("k",kT) |
17261 | 897 |
val kb = Bound 1 |
15531 | 898 |
in SOME (upd$kb$rb,rb,[kv,rv],false) end)) |
17261 | 899 |
| mk_eq_terms r = NONE |
13462 | 900 |
in |
17261 | 901 |
(case mk_eq_terms (upd$k$r) of |
902 |
SOME (trm,trm',vars,update_s) |
|
903 |
=> if update_s |
|
904 |
then SOME (prove_split_simp sg ss domS |
|
16872 | 905 |
(list_all(vars,(equals rangeS$(sel$trm)$trm')))) |
17261 | 906 |
else SOME (prove_split_simp sg ss domS |
16872 | 907 |
(list_all(vars,(equals rangeS$(sel$trm)$(sel$trm'))))) |
15531 | 908 |
| NONE => NONE) |
13462 | 909 |
end |
15531 | 910 |
| NONE => NONE) |
911 |
| NONE => NONE) |
|
912 |
| _ => NONE)); |
|
7178 | 913 |
|
17261 | 914 |
(* record_upd_simproc *) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
915 |
(* simplify multiple updates: |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
916 |
* (1) "r(|M:=3,N:=1,M:=2,N:=4|) == r(|M:=2,N:=4|)" |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
917 |
* (2) "r(|M:= M r|) = r" |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
918 |
* For (2) special care of "more" updates has to be taken: |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
919 |
* r(|more := m; A := A r|) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
920 |
* If A is contained in the fields of m we cannot remove the update A := A r! |
17261 | 921 |
* (But r(|more := r; A := A (r(|more := r|))|) = r(|more := r|) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
922 |
*) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
923 |
val record_upd_simproc = |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
924 |
Simplifier.simproc (Theory.sign_of HOL.thy) "record_upd_simp" ["(u k r)"] |
16973 | 925 |
(fn sg => fn ss => fn t => |
16872 | 926 |
(case t of ((upd as Const (u, Type(_,[_,Type(_,[rT,_])]))) $ k $ r) => |
17261 | 927 |
let datatype ('a,'b) calc = Init of 'b | Inter of 'a |
16458 | 928 |
val {sel_upd={selectors,updates,...},extfields,...} = RecordsData.get sg; |
17261 | 929 |
|
930 |
(*fun mk_abs_var x t = (x, fastype_of t);*) |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
931 |
fun sel_name u = NameSpace.base (unsuffix updateN u); |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
932 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
933 |
fun seed s (upd as Const (more,Type(_,[mT,_]))$ k $ r) = |
16822 | 934 |
if has_field extfields s mT then upd else seed s r |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
935 |
| seed _ r = r; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
936 |
|
17261 | 937 |
fun grow u uT k kT vars (sprout,skeleton) = |
938 |
if sel_name u = moreN |
|
16872 | 939 |
then let val kv = ("k", kT); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
940 |
val kb = Bound (length vars); |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
941 |
in ((Const (u,uT)$k$sprout,Const (u,uT)$kb$skeleton),kv::vars) end |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
942 |
else ((sprout,skeleton),vars); |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
943 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
944 |
fun is_upd_same (sprout,skeleton) u ((sel as Const (s,_))$r) = |
17261 | 945 |
if (unsuffix updateN u) = s andalso (seed s sprout) = r |
15531 | 946 |
then SOME (sel,seed s skeleton) |
947 |
else NONE |
|
948 |
| is_upd_same _ _ _ = NONE |
|
17261 | 949 |
|
16872 | 950 |
fun init_seed r = ((r,Bound 0), [("r", rT)]); |
17261 | 951 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
952 |
(* mk_updterm returns either |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
953 |
* - Init (orig-term, orig-term-skeleton, vars) if no optimisation can be made, |
17261 | 954 |
* where vars are the bound variables in the skeleton |
955 |
* - Inter (orig-term-skeleton,simplified-term-skeleton, |
|
16872 | 956 |
* vars, (term-sprout, skeleton-sprout)) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
957 |
* where "All vars. orig-term-skeleton = simplified-term-skeleton" is |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
958 |
* the desired simplification rule, |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
959 |
* the sprouts accumulate the "more-updates" on the way from the seed |
17261 | 960 |
* to the outermost update. It is only relevant to calculate the |
961 |
* possible simplification for (2) |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
962 |
* The algorithm first walks down the updates to the seed-record while |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
963 |
* memorising the updates in the already-table. While walking up the |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
964 |
* updates again, the optimised term is constructed. |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
965 |
*) |
17261 | 966 |
fun mk_updterm upds already |
16872 | 967 |
(t as ((upd as Const (u,uT as (Type (_,[kT,_])))) $ k $ r)) = |
17261 | 968 |
if Symtab.defined upds u |
969 |
then let |
|
970 |
fun rest already = mk_updterm upds already |
|
971 |
in if u mem_string already |
|
972 |
then (case (rest already r) of |
|
973 |
Init ((sprout,skel),vars) => |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
974 |
let |
17261 | 975 |
val kv = (sel_name u, kT); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
976 |
val kb = Bound (length vars); |
16872 | 977 |
val (sprout',vars')= grow u uT k kT (kv::vars) (sprout,skel); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
978 |
in Inter (upd$kb$skel,skel,vars',sprout') end |
17261 | 979 |
| Inter (trm,trm',vars,sprout) => |
980 |
let |
|
981 |
val kv = (sel_name u, kT); |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
982 |
val kb = Bound (length vars); |
16872 | 983 |
val (sprout',vars') = grow u uT k kT (kv::vars) sprout; |
17261 | 984 |
in Inter(upd$kb$trm,trm',kv::vars',sprout') end) |
985 |
else |
|
986 |
(case rest (u::already) r of |
|
987 |
Init ((sprout,skel),vars) => |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
988 |
(case is_upd_same (sprout,skel) u k of |
17261 | 989 |
SOME (sel,skel') => |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
990 |
let |
17261 | 991 |
val (sprout',vars') = grow u uT k kT vars (sprout,skel); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
992 |
in Inter(upd$(sel$skel')$skel,skel,vars',sprout') end |
17261 | 993 |
| NONE => |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
994 |
let |
17261 | 995 |
val kv = (sel_name u, kT); |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
996 |
val kb = Bound (length vars); |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
997 |
in Init ((upd$k$sprout,upd$kb$skel),kv::vars) end) |
17261 | 998 |
| Inter (trm,trm',vars,sprout) => |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
999 |
(case is_upd_same sprout u k of |
15531 | 1000 |
SOME (sel,skel) => |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1001 |
let |
16872 | 1002 |
val (sprout',vars') = grow u uT k kT vars sprout |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1003 |
in Inter(upd$(sel$skel)$trm,trm',vars',sprout') end |
15531 | 1004 |
| NONE => |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1005 |
let |
17261 | 1006 |
val kv = (sel_name u, kT) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1007 |
val kb = Bound (length vars) |
16872 | 1008 |
val (sprout',vars') = grow u uT k kT (kv::vars) sprout |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1009 |
in Inter (upd$kb$trm,upd$kb$trm',vars',sprout') end)) |
17261 | 1010 |
end |
1011 |
else Init (init_seed t) |
|
1012 |
| mk_updterm _ _ t = Init (init_seed t); |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1013 |
|
17261 | 1014 |
in (case mk_updterm updates [] t of |
1015 |
Inter (trm,trm',vars,_) |
|
1016 |
=> SOME (prove_split_simp sg ss rT |
|
16872 | 1017 |
(list_all(vars,(equals rT$trm$trm')))) |
15531 | 1018 |
| _ => NONE) |
17261 | 1019 |
end |
15531 | 1020 |
| _ => NONE)); |
15059 | 1021 |
end |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1022 |
|
14255 | 1023 |
(* record_eq_simproc *) |
1024 |
(* looks up the most specific record-equality. |
|
1025 |
* Note on efficiency: |
|
1026 |
* Testing equality of records boils down to the test of equality of all components. |
|
1027 |
* Therefore the complexity is: #components * complexity for single component. |
|
1028 |
* Especially if a record has a lot of components it may be better to split up |
|
1029 |
* the record first and do simplification on that (record_split_simp_tac). |
|
1030 |
* e.g. r(|lots of updates|) = x |
|
1031 |
* |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1032 |
* record_eq_simproc record_split_simp_tac |
17261 | 1033 |
* Complexity: #components * #updates #updates |
1034 |
* |
|
14255 | 1035 |
*) |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
1036 |
val record_eq_simproc = |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
1037 |
Simplifier.simproc (Theory.sign_of HOL.thy) "record_eq_simp" ["r = s"] |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
1038 |
(fn sg => fn _ => fn t => |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
1039 |
(case t of Const ("op =", Type (_, [T, _])) $ _ $ _ => |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1040 |
(case rec_id (~1) T of |
15531 | 1041 |
"" => NONE |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1042 |
| name => (case get_equalities sg name of |
15531 | 1043 |
NONE => NONE |
1044 |
| SOME thm => SOME (thm RS Eq_TrueI))) |
|
1045 |
| _ => NONE)); |
|
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
1046 |
|
14255 | 1047 |
(* record_split_simproc *) |
17261 | 1048 |
(* splits quantified occurrences of records, for which P holds. P can peek on the |
14255 | 1049 |
* subterm starting at the quantified occurrence of the record (including the quantifier) |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1050 |
* P t = 0: do not split |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1051 |
* P t = ~1: completely split |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1052 |
* P t > 0: split up to given bound of record extensions |
14255 | 1053 |
*) |
1054 |
fun record_split_simproc P = |
|
1055 |
Simplifier.simproc (Theory.sign_of HOL.thy) "record_split_simp" ["(a t)"] |
|
1056 |
(fn sg => fn _ => fn t => |
|
1057 |
(case t of (Const (quantifier, Type (_, [Type (_, [T, _]), _])))$trm => |
|
1058 |
if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex" |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1059 |
then (case rec_id (~1) T of |
15531 | 1060 |
"" => NONE |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1061 |
| name |
17261 | 1062 |
=> let val split = P t |
1063 |
in if split <> 0 then |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1064 |
(case get_splits sg (rec_id split T) of |
15531 | 1065 |
NONE => NONE |
17261 | 1066 |
| SOME (all_thm, All_thm, Ex_thm,_) |
15531 | 1067 |
=> SOME (case quantifier of |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1068 |
"all" => all_thm |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1069 |
| "All" => All_thm RS HOL.eq_reflection |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1070 |
| "Ex" => Ex_thm RS HOL.eq_reflection |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1071 |
| _ => error "record_split_simproc")) |
15531 | 1072 |
else NONE |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1073 |
end) |
15531 | 1074 |
else NONE |
1075 |
| _ => NONE)) |
|
7178 | 1076 |
|
14427 | 1077 |
val record_ex_sel_eq_simproc = |
1078 |
Simplifier.simproc (Theory.sign_of HOL.thy) "record_ex_sel_eq_simproc" ["Ex t"] |
|
16973 | 1079 |
(fn sg => fn ss => fn t => |
17261 | 1080 |
let |
16973 | 1081 |
fun prove prop = |
1082 |
quick_and_dirty_prove true sg [] prop |
|
1083 |
(fn _ => simp_tac (Simplifier.inherit_bounds ss (get_simpset sg) |
|
1084 |
addsimps simp_thms addsimprocs [record_split_simproc (K ~1)]) 1); |
|
14959 | 1085 |
|
1086 |
fun mkeq (lr,Teq,(sel,Tsel),x) i = |
|
15531 | 1087 |
(case get_selectors sg sel of SOME () => |
17261 | 1088 |
let val x' = if not (loose_bvar1 (x,0)) |
1089 |
then Free ("x" ^ string_of_int i, range_type Tsel) |
|
14959 | 1090 |
else raise TERM ("",[x]); |
1091 |
val sel' = Const (sel,Tsel)$Bound 0; |
|
1092 |
val (l,r) = if lr then (sel',x') else (x',sel'); |
|
1093 |
in Const ("op =",Teq)$l$r end |
|
15531 | 1094 |
| NONE => raise TERM ("",[Const (sel,Tsel)])); |
14959 | 1095 |
|
17261 | 1096 |
fun dest_sel_eq (Const ("op =",Teq)$(Const (sel,Tsel)$Bound 0)$X) = |
14959 | 1097 |
(true,Teq,(sel,Tsel),X) |
1098 |
| dest_sel_eq (Const ("op =",Teq)$X$(Const (sel,Tsel)$Bound 0)) = |
|
1099 |
(false,Teq,(sel,Tsel),X) |
|
1100 |
| dest_sel_eq _ = raise TERM ("",[]); |
|
1101 |
||
17261 | 1102 |
in |
1103 |
(case t of |
|
14959 | 1104 |
(Const ("Ex",Tex)$Abs(s,T,t)) => |
16872 | 1105 |
(let val eq = mkeq (dest_sel_eq t) 0; |
14959 | 1106 |
val prop = list_all ([("r",T)], |
1107 |
Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq), |
|
1108 |
HOLogic.true_const)); |
|
15531 | 1109 |
in SOME (prove prop) end |
16872 | 1110 |
handle TERM _ => NONE) |
17261 | 1111 |
| _ => NONE) |
14427 | 1112 |
end) |
1113 |
||
5698 | 1114 |
|
17261 | 1115 |
|
6358 | 1116 |
|
14255 | 1117 |
local |
1118 |
val inductive_atomize = thms "induct_atomize"; |
|
1119 |
val inductive_rulify1 = thms "induct_rulify1"; |
|
1120 |
in |
|
1121 |
(* record_split_simp_tac *) |
|
17261 | 1122 |
(* splits (and simplifies) all records in the goal for which P holds. |
14255 | 1123 |
* For quantified occurrences of a record |
1124 |
* P can peek on the whole subterm (including the quantifier); for free variables P |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1125 |
* can only peek on the variable itself. |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1126 |
* P t = 0: do not split |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1127 |
* P t = ~1: completely split |
17261 | 1128 |
* P t > 0: split up to given bound of record extensions |
14255 | 1129 |
*) |
15248
b436486091a6
record_split_simp_tac now can get simp rules as parameter
schirmer
parents:
15215
diff
changeset
|
1130 |
fun record_split_simp_tac thms P i st = |
14255 | 1131 |
let |
1132 |
val sg = Thm.sign_of_thm st; |
|
17261 | 1133 |
val {sel_upd={simpset,...},...} |
16458 | 1134 |
= RecordsData.get sg; |
14255 | 1135 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1136 |
val has_rec = exists_Const |
14255 | 1137 |
(fn (s, Type (_, [Type (_, [T, _]), _])) => |
17261 | 1138 |
(s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T |
14255 | 1139 |
| _ => false); |
1140 |
||
15570 | 1141 |
val goal = List.nth (Thm.prems_of st, i - 1); |
1142 |
val frees = List.filter (is_recT o type_of) (term_frees goal); |
|
14255 | 1143 |
|
17261 | 1144 |
fun mk_split_free_tac free induct_thm i = |
1145 |
let val cfree = cterm_of sg free; |
|
14255 | 1146 |
val (_$(_$r)) = concl_of induct_thm; |
1147 |
val crec = cterm_of sg r; |
|
1148 |
val thm = cterm_instantiate [(crec,cfree)] induct_thm; |
|
1149 |
in EVERY [simp_tac (HOL_basic_ss addsimps inductive_atomize) i, |
|
1150 |
rtac thm i, |
|
1151 |
simp_tac (HOL_basic_ss addsimps inductive_rulify1) i] |
|
17261 | 1152 |
end; |
14255 | 1153 |
|
17261 | 1154 |
fun split_free_tac P i (free as Free (n,T)) = |
1155 |
(case rec_id (~1) T of |
|
15531 | 1156 |
"" => NONE |
17261 | 1157 |
| name => let val split = P free |
1158 |
in if split <> 0 then |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1159 |
(case get_splits sg (rec_id split T) of |
15531 | 1160 |
NONE => NONE |
1161 |
| SOME (_,_,_,induct_thm) |
|
1162 |
=> SOME (mk_split_free_tac free induct_thm i)) |
|
1163 |
else NONE |
|
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1164 |
end) |
15531 | 1165 |
| split_free_tac _ _ _ = NONE; |
14255 | 1166 |
|
15570 | 1167 |
val split_frees_tacs = List.mapPartial (split_free_tac P i) frees; |
17261 | 1168 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1169 |
val simprocs = if has_rec goal then [record_split_simproc P] else []; |
17261 | 1170 |
|
1171 |
in st |> ((EVERY split_frees_tacs) |
|
17257
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
haftmann
parents:
17190
diff
changeset
|
1172 |
THEN (Simplifier.full_simp_tac (simpset addsimps thms addsimprocs simprocs) i)) |
15570 | 1173 |
end handle Empty => Seq.empty; |
14255 | 1174 |
end; |
1175 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1176 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1177 |
(* record_split_tac *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1178 |
(* splits all records in the goal, which are quantified by ! or !!. *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1179 |
fun record_split_tac i st = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1180 |
let |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1181 |
val sg = Thm.sign_of_thm st; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1182 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1183 |
val has_rec = exists_Const |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1184 |
(fn (s, Type (_, [Type (_, [T, _]), _])) => |
17261 | 1185 |
(s = "all" orelse s = "All") andalso is_recT T |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1186 |
| _ => false); |
17261 | 1187 |
|
1188 |
val goal = List.nth (Thm.prems_of st, i - 1); |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1189 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1190 |
fun is_all t = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1191 |
(case t of (Const (quantifier, _)$_) => |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1192 |
if quantifier = "All" orelse quantifier = "all" then ~1 else 0 |
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
1193 |
| _ => 0); |
17261 | 1194 |
|
1195 |
in if has_rec goal |
|
1196 |
then Simplifier.full_simp_tac |
|
1197 |
(HOL_basic_ss addsimprocs [record_split_simproc is_all]) i st |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1198 |
else Seq.empty |
15570 | 1199 |
end handle Subscript => Seq.empty; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1200 |
|
6358 | 1201 |
(* wrapper *) |
1202 |
||
5707 | 1203 |
val record_split_name = "record_split_tac"; |
1204 |
val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac); |
|
5698 | 1205 |
|
16330 | 1206 |
|
1207 |
||
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1208 |
(** theory extender interface **) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1209 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1210 |
(* prepare arguments *) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1211 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1212 |
fun read_raw_parent sign s = |
16330 | 1213 |
(case Sign.read_typ_abbrev (sign, K NONE) s handle TYPE (msg, _, _) => error msg of |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1214 |
Type (name, Ts) => (Ts, name) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1215 |
| _ => error ("Bad parent record specification: " ^ quote s)); |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1216 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1217 |
fun read_typ sign (env, s) = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
1218 |
let |
5060 | 1219 |
fun def_sort (x, ~1) = assoc (env, x) |
15531 | 1220 |
| def_sort _ = NONE; |
5060 | 1221 |
val T = Type.no_tvars (Sign.read_typ (sign, def_sort) s) handle TYP |