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