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