author | wenzelm |
Tue, 29 Sep 2009 21:34:59 +0200 | |
changeset 32761 | 54fee94b2b29 |
parent 32760 | ea6672bff5dd |
child 32763 | ebfaf9e3c03a |
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 |
32745
192d58483fdf
Simplification of various aspects of the IsTuple component
Thomas Sewell <tsewell@nicta.com.au>
parents:
32744
diff
changeset
|
2 |
Authors: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen |
192d58483fdf
Simplification of various aspects of the IsTuple component
Thomas Sewell <tsewell@nicta.com.au>
parents:
32744
diff
changeset
|
3 |
Thomas Sewell, NICTA |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
4 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
5 |
Extensible records with structural subtyping in HOL. |
5698 | 6 |
*) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
7 |
|
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31136
diff
changeset
|
8 |
signature BASIC_RECORD = |
5698 | 9 |
sig |
7178 | 10 |
val record_simproc: simproc |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
11 |
val record_eq_simproc: simproc |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
12 |
val record_upd_simproc: simproc |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
13 |
val record_split_simproc: (term -> int) -> simproc |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
14 |
val record_ex_sel_eq_simproc: simproc |
5698 | 15 |
val record_split_tac: int -> tactic |
15273
771af451a062
* extended interface of record_split_simp_tac and record_split_simproc
schirmer
parents:
15258
diff
changeset
|
16 |
val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic |
5713 | 17 |
val record_split_name: string |
5698 | 18 |
val record_split_wrapper: string * wrapper |
32740 | 19 |
val print_record_type_abbr: bool Unsynchronized.ref |
20 |
val print_record_type_as_fields: bool Unsynchronized.ref |
|
5698 | 21 |
end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
22 |
|
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31136
diff
changeset
|
23 |
signature RECORD = |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
24 |
sig |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31136
diff
changeset
|
25 |
include BASIC_RECORD |
32740 | 26 |
val timing: bool Unsynchronized.ref |
27 |
val record_quick_and_dirty_sensitive: bool Unsynchronized.ref |
|
8574 | 28 |
val updateN: string |
21363 | 29 |
val updN: string |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
30 |
val ext_typeN: string |
21363 | 31 |
val extN: string |
32 |
val makeN: string |
|
33 |
val moreN: string |
|
34 |
val ext_dest: string |
|
35 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
36 |
val last_extT: typ -> (string * typ list) option |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
37 |
val dest_recTs : typ -> (string * typ list) list |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
38 |
val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
39 |
val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ) |
26088 | 40 |
val get_parent: theory -> string -> (typ list * string) option |
41 |
val get_extension: theory -> string -> (string * typ list) option |
|
16458 | 42 |
val get_extinjects: theory -> thm list |
43 |
val get_simpset: theory -> simpset |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
44 |
val print_records: theory -> unit |
27278
129574589734
export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents:
27239
diff
changeset
|
45 |
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
|
46 |
val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list |
26477
ecf06644f6cb
eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents:
26423
diff
changeset
|
47 |
val add_record: bool -> string list * string -> string option -> (string * string * mixfix) list |
16458 | 48 |
-> theory -> theory |
26477
ecf06644f6cb
eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents:
26423
diff
changeset
|
49 |
val add_record_i: bool -> string list * string -> (typ list * string) option |
16458 | 50 |
-> (string * typ * mixfix) list -> theory -> theory |
18708 | 51 |
val setup: theory -> theory |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
52 |
end; |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
53 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
54 |
|
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
55 |
signature ISTUPLE_SUPPORT = |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
56 |
sig |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
57 |
val add_istuple_type: bstring * string list -> (typ * typ) -> theory -> term * term * theory |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
58 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
59 |
val mk_cons_tuple: term * term -> term |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
60 |
val dest_cons_tuple: term -> term * term |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
61 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
62 |
val istuple_intros_tac: theory -> int -> tactic |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
63 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
64 |
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
|
65 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
66 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
67 |
structure IsTupleSupport: ISTUPLE_SUPPORT = |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
68 |
struct |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
69 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
70 |
val isomN = "_TupleIsom"; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
71 |
val defN = "_def"; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
72 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
73 |
val istuple_UNIV_I = @{thm "istuple_UNIV_I"}; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
74 |
val istuple_True_simp = @{thm "istuple_True_simp"}; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
75 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
76 |
val istuple_intro = @{thm "isomorphic_tuple_intro"}; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
77 |
val istuple_intros = build_net (@{thms "isomorphic_tuple.intros"}); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
78 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
79 |
val constname = fst o dest_Const; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
80 |
val tuple_istuple = (constname @{term tuple_istuple}, @{thm tuple_istuple}); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
81 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
82 |
val istuple_constN = constname @{term isomorphic_tuple}; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
83 |
val istuple_consN = constname @{term istuple_cons}; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
84 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
85 |
val tup_isom_typeN = fst (dest_Type @{typ "('a, 'b, 'c) tuple_isomorphism"}); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
86 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
87 |
fun named_cterm_instantiate values thm = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
88 |
let |
32758
cd47afaf0d78
Replace OldTerm.term_vars with Term.add_vars in named_cterm_instantiate.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32757
diff
changeset
|
89 |
fun match name ((name', _), _) = name = name' |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
90 |
| match name _ = false; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
91 |
fun getvar name = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
92 |
(case find_first (match name) (Term.add_vars (prop_of thm) []) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
93 |
SOME var => cterm_of (theory_of_thm thm) (Var var) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
94 |
| 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
|
95 |
in |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
96 |
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
|
97 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
98 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
99 |
structure IsTupleThms = TheoryDataFun |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
100 |
( |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
101 |
type T = thm Symtab.table; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
102 |
val empty = Symtab.make [tuple_istuple]; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
103 |
val copy = I; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
104 |
val extend = I; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
105 |
val merge = K (Symtab.merge Thm.eq_thm_prop); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
106 |
); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
107 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
108 |
fun do_typedef name repT alphas thy = |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
109 |
let |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
110 |
fun get_thms thy name = |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
111 |
let |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
112 |
val SOME { Rep_inject=rep_inject, Abs_name=absN, abs_type=absT, |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
113 |
Abs_inverse=abs_inverse, ...} = Typedef.get_info thy name; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
114 |
val rewrite_rule = MetaSimplifier.rewrite_rule [istuple_UNIV_I, istuple_True_simp]; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
115 |
in (map rewrite_rule [rep_inject, abs_inverse], |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
116 |
Const (absN, repT --> absT), absT) end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
117 |
in |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
118 |
thy |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
119 |
|> Typecopy.typecopy (Binding.name name, alphas) repT NONE |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
120 |
|-> (fn (name, _) => `(fn thy => get_thms thy name)) |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
121 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
122 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
123 |
fun mk_cons_tuple (left, right) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
124 |
let |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
125 |
val (leftT, rightT) = (fastype_of left, fastype_of right); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
126 |
val prodT = HOLogic.mk_prodT (leftT, rightT); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
127 |
val isomT = Type (tup_isom_typeN, [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
|
128 |
in |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
129 |
Const (istuple_consN, isomT --> leftT --> rightT --> prodT) $ |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
130 |
Const (fst tuple_istuple, 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
|
131 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
132 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
133 |
fun dest_cons_tuple (v as Const (ic, _) $ Const _ $ left $ right) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
134 |
if ic = istuple_consN then (left, right) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
135 |
else raise TERM ("dest_cons_tuple", [v]) |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
136 |
| dest_cons_tuple v = raise TERM ("dest_cons_tuple", [v]); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
137 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
138 |
fun add_istuple_type (name, alphas) (leftT, rightT) thy = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
139 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
140 |
val repT = HOLogic.mk_prodT (leftT, rightT); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
141 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
142 |
val (([rep_inject, abs_inverse], absC, absT), typ_thy) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
143 |
thy |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
144 |
|> do_typedef name repT alphas |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
145 |
||> Sign.add_path name; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
146 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
147 |
(*construct a type and body for the isomorphism constant by |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
148 |
instantiating the theorem to which the definition will be applied*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
149 |
val intro_inst = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
150 |
rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] istuple_intro; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
151 |
val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
152 |
val isomT = fastype_of body; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
153 |
val isom_bind = Binding.name (name ^ isomN); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
154 |
val isom = Const (Sign.full_name typ_thy isom_bind, isomT); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
155 |
val isom_spec = (name ^ isomN ^ defN, Logic.mk_equals (isom, body)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
156 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
157 |
val ([isom_def], cdef_thy) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
158 |
typ_thy |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
159 |
|> Sign.add_consts_i [Syntax.no_syn (isom_bind, isomT)] |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
160 |
|> PureThy.add_defs false [Thm.no_attributes (apfst Binding.name isom_spec)]; |
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 istuple = isom_def RS (abs_inverse RS (rep_inject RS istuple_intro)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
163 |
val cons = Const (istuple_consN, isomT --> leftT --> rightT --> absT); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
164 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
165 |
val thm_thy = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
166 |
cdef_thy |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
167 |
|> IsTupleThms.map (Symtab.insert Thm.eq_thm_prop (constname isom, istuple)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
168 |
|> Sign.parent_path; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
169 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
170 |
(isom, cons $ isom, thm_thy) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
171 |
end; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
172 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
173 |
fun istuple_intros_tac thy = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
174 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
175 |
val isthms = IsTupleThms.get thy; |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
176 |
fun err s t = raise TERM ("istuple_intros_tac: " ^ s, [t]); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
177 |
val use_istuple_thm_tac = SUBGOAL (fn (goal, n) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
178 |
let |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
179 |
val goal' = Envir.beta_eta_contract goal; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
180 |
val isom = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
181 |
(case goal' of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
182 |
Const tp $ (Const pr $ Const is) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
183 |
if fst tp = "Trueprop" andalso fst pr = istuple_constN |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
184 |
then Const is |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
185 |
else err "unexpected goal predicate" goal' |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
186 |
| _ => err "unexpected goal format" goal'); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
187 |
val isthm = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
188 |
(case Symtab.lookup isthms (constname isom) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
189 |
SOME isthm => isthm |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
190 |
| NONE => err "no thm found for constant" isom); |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
191 |
in rtac isthm n end); |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
192 |
in |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
193 |
fn n => resolve_from_net_tac istuple_intros n THEN use_istuple_thm_tac n |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
194 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
195 |
|
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
196 |
end; |
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
197 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
198 |
|
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31136
diff
changeset
|
199 |
structure Record: RECORD = |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
200 |
struct |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
201 |
|
21546 | 202 |
val eq_reflection = thm "eq_reflection"; |
27104
791607529f6d
rep_datatype command now takes list of constructors as input arguments
haftmann
parents:
26957
diff
changeset
|
203 |
val Pair_eq = thm "Product_Type.prod.inject"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
204 |
val atomize_all = thm "HOL.atomize_all"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
205 |
val atomize_imp = thm "HOL.atomize_imp"; |
17960 | 206 |
val meta_allE = thm "Pure.meta_allE"; |
207 |
val prop_subst = thm "prop_subst"; |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
208 |
val Pair_sel_convs = [fst_conv, snd_conv]; |
26359 | 209 |
val K_record_comp = @{thm "K_record_comp"}; |
210 |
val K_comp_convs = [@{thm o_apply}, K_record_comp] |
|
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
211 |
val transitive_thm = thm "transitive"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
212 |
val o_assoc = @{thm "o_assoc"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
213 |
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
|
214 |
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
|
215 |
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
|
216 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
217 |
val refl_conj_eq = thm "refl_conj_eq"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
218 |
val meta_all_sameI = thm "meta_all_sameI"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
219 |
val meta_iffD2 = thm "meta_iffD2"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
220 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
221 |
val surject_assistI = @{thm "istuple_surjective_proof_assistI"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
222 |
val surject_assist_idE = @{thm "istuple_surjective_proof_assist_idE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
223 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
224 |
val updacc_accessor_eqE = @{thm "update_accessor_accessor_eqE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
225 |
val updacc_updator_eqE = @{thm "update_accessor_updator_eqE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
226 |
val updacc_eq_idI = @{thm "istuple_update_accessor_eq_assist_idI"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
227 |
val updacc_eq_triv = @{thm "istuple_update_accessor_eq_assist_triv"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
228 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
229 |
val updacc_foldE = @{thm "update_accessor_congruence_foldE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
230 |
val updacc_unfoldE = @{thm "update_accessor_congruence_unfoldE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
231 |
val updacc_noopE = @{thm "update_accessor_noopE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
232 |
val updacc_noop_compE = @{thm "update_accessor_noop_compE"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
233 |
val updacc_cong_idI = @{thm "update_accessor_cong_assist_idI"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
234 |
val updacc_cong_triv = @{thm "update_accessor_cong_assist_triv"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
235 |
val updacc_cong_from_eq = @{thm "istuple_update_accessor_cong_from_eq"}; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
236 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
237 |
val o_eq_dest = thm "o_eq_dest"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
238 |
val o_eq_id_dest = thm "o_eq_id_dest"; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
239 |
val o_eq_dest_lhs = thm "o_eq_dest_lhs"; |
11832 | 240 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
241 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
242 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
243 |
(** name components **) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
244 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
245 |
val rN = "r"; |
15215 | 246 |
val wN = "w"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
247 |
val moreN = "more"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
248 |
val schemeN = "_scheme"; |
17261 | 249 |
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
|
250 |
val inner_typeN = "_inner_type"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
251 |
val extN ="_ext"; |
15215 | 252 |
val casesN = "_cases"; |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
253 |
val ext_dest = "_sel"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
254 |
val updateN = "_update"; |
15215 | 255 |
val updN = "_upd"; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
256 |
val makeN = "make"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
257 |
val fields_selN = "fields"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
258 |
val extendN = "extend"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
259 |
val truncateN = "truncate"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
260 |
|
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31136
diff
changeset
|
261 |
(*see typedef.ML*) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
262 |
val RepN = "Rep_"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
263 |
val AbsN = "Abs_"; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
264 |
|
32335 | 265 |
|
266 |
||
4894 | 267 |
(*** utilities ***) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
268 |
|
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
269 |
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
|
270 |
|
19748 | 271 |
fun varifyT midx = |
272 |
let fun varify (a, S) = TVar ((a, midx + 1), S); |
|
273 |
in map_type_tfree varify end; |
|
274 |
||
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
275 |
fun domain_type' T = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
276 |
domain_type T handle Match => T; |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
277 |
|
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
278 |
fun range_type' T = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
279 |
range_type T handle Match => T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
280 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
281 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
282 |
(* messages *) (* FIXME proper context *) |
5698 | 283 |
|
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
284 |
fun trace_thm str thm = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
285 |
tracing (str ^ Pretty.string_of (Display.pretty_thm_without_context thm)); |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
286 |
|
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
287 |
fun trace_thms str thms = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
288 |
(tracing str; map (trace_thm "") thms); |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
289 |
|
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
290 |
fun trace_term str t = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
291 |
tracing (str ^ Syntax.string_of_term_global Pure.thy t); |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
292 |
|
32335 | 293 |
|
15012
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents:
14981
diff
changeset
|
294 |
(* timing *) |
28fa57b57209
Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents:
14981
diff
changeset
|
295 |
|
32740 | 296 |
val timing = Unsynchronized.ref false; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
297 |
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
|
298 |
fun timing_msg s = if ! timing then warning s else (); |
17261 | 299 |
|
32335 | 300 |
|
12255 | 301 |
(* syntax *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
302 |
|
12247 | 303 |
fun prune n xs = Library.drop (n, xs); |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30345
diff
changeset
|
304 |
fun prefix_base s = Long_Name.map_base_name (fn bname => s ^ bname); |
11832 | 305 |
|
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
306 |
val Trueprop = HOLogic.mk_Trueprop; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
307 |
fun All xs t = Term.list_all_free (xs, t); |
4894 | 308 |
|
11934
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
309 |
infix 9 $$; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
310 |
infix 0 :== ===; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
311 |
infixr 0 ==>; |
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
312 |
|
6c1bf72430b6
derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents:
11927
diff
changeset
|
313 |
val (op $$) = Term.list_comb; |
24255 | 314 |
val (op :==) = PrimitiveDefs.mk_defpair; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
315 |
val (op ===) = Trueprop o HOLogic.mk_eq; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
316 |
val (op ==>) = Logic.mk_implies; |
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
317 |
|
32335 | 318 |
|
11832 | 319 |
(* morphisms *) |
320 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
321 |
fun mk_RepN name = suffix ext_typeN (prefix_base RepN name); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
322 |
fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name); |
11832 | 323 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
324 |
fun mk_Rep name repT absT = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
325 |
Const (suffix ext_typeN (prefix_base RepN name), absT --> repT); |
11832 | 326 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
327 |
fun mk_Abs name repT absT = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
328 |
Const (mk_AbsN name, repT --> absT); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
329 |
|
32335 | 330 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
331 |
(* constructor *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
332 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
333 |
fun mk_extC (name, T) Ts = (suffix extN name, Ts ---> T); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
334 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
335 |
fun mk_ext (name, T) ts = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
336 |
let val Ts = map fastype_of ts |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
337 |
in list_comb (Const (mk_extC (name, T) Ts), ts) end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
338 |
|
32335 | 339 |
|
15215 | 340 |
(* cases *) |
341 |
||
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
342 |
fun mk_casesC (name, T, vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
343 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
344 |
fun mk_cases (name, T, vT) f = |
17261 | 345 |
let val Ts = binder_types (fastype_of f) |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
346 |
in Const (mk_casesC (name, T, vT) Ts) $ f end; |
17261 | 347 |
|
32335 | 348 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
349 |
(* selector *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
350 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
351 |
fun mk_selC sT (c, T) = (c, sT --> T); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
352 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
353 |
fun mk_sel s (c, T) = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
354 |
let val sT = fastype_of s |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
355 |
in Const (mk_selC sT (c, T)) $ s end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
356 |
|
32335 | 357 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
358 |
(* updates *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
359 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
360 |
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
|
361 |
|
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
362 |
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
|
363 |
let val vT = domain_type (fastype_of v); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
364 |
in Const (mk_updC sfx sT (c, vT)) $ v end; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
365 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
366 |
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
|
367 |
|
32335 | 368 |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
369 |
(* types *) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
370 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
371 |
fun dest_recT (typ as Type (c_ext_type, Ts as (T :: _))) = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
372 |
(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
|
373 |
NONE => raise TYPE ("Record.dest_recT", [typ], []) |
15570 | 374 |
| 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
|
375 |
| dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []); |
5197 | 376 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
377 |
fun is_recT T = |
17261 | 378 |
(case try dest_recT T of NONE => false | SOME _ => true); |
11833 | 379 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
380 |
fun dest_recTs T = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
381 |
let val ((c, Ts), U) = dest_recT T |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
382 |
in (c, Ts) :: dest_recTs U |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
383 |
end handle TYPE _ => []; |
14255 | 384 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
385 |
fun last_extT T = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
386 |
let val ((c, Ts), U) = dest_recT T in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
387 |
(case last_extT U of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
388 |
NONE => SOME (c, Ts) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
389 |
| SOME l => SOME l) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
390 |
end handle TYPE _ => NONE; |
14255 | 391 |
|
17261 | 392 |
fun rec_id i T = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
393 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
394 |
val rTs = dest_recTs T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
395 |
val rTs' = if i < 0 then rTs else Library.take (i, rTs); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
396 |
in Library.foldl (fn (s, (c, T)) => s ^ c) ("", rTs') end; (* FIXME ? *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
397 |
|
32335 | 398 |
|
399 |
||
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
400 |
(*** extend theory by record definition ***) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
401 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
402 |
(** record info **) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
403 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
404 |
(* type record_info and parent_info *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
405 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
406 |
type record_info = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
407 |
{args: (string * sort) list, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
408 |
parent: (typ list * string) option, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
409 |
fields: (string * typ) list, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
410 |
extension: (string * typ list), |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
411 |
induct: thm, |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
412 |
extdef: thm}; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
413 |
|
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
414 |
fun make_record_info args parent fields extension induct extdef = |
17261 | 415 |
{args = args, parent = parent, fields = fields, extension = extension, |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
416 |
induct = induct, extdef = extdef}: record_info; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
417 |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
418 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
419 |
type parent_info = |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
420 |
{name: string, |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
421 |
fields: (string * typ) list, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
422 |
extension: (string * typ list), |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
423 |
induct: thm, |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
424 |
extdef: thm}; |
11927
96f267adc029
provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents:
11923
diff
changeset
|
425 |
|
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
426 |
fun make_parent_info name fields extension induct extdef = |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
427 |
{name = name, fields = fields, extension = extension, |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
428 |
induct = induct, extdef = extdef}: parent_info; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
429 |
|
22846 | 430 |
|
431 |
(* theory data *) |
|
5001 | 432 |
|
7178 | 433 |
type record_data = |
434 |
{records: record_info Symtab.table, |
|
435 |
sel_upd: |
|
32744
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
436 |
{selectors: (int * bool) Symtab.table, |
7178 | 437 |
updates: string Symtab.table, |
32744
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
438 |
simpset: Simplifier.simpset, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
439 |
defset: Simplifier.simpset, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
440 |
foldcong: Simplifier.simpset, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
441 |
unfoldcong: Simplifier.simpset}, |
14255 | 442 |
equalities: thm Symtab.table, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
443 |
extinjects: thm list, |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
444 |
extsplit: thm Symtab.table, (* maps extension name to split rule *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
445 |
splits: (thm*thm*thm*thm) Symtab.table, (* !!, !, EX - split-equalities, induct rule *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
446 |
extfields: (string*typ) list Symtab.table, (* maps extension to its fields *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
447 |
fieldext: (string*typ list) Symtab.table}; (* maps field to its extension *) |
7178 | 448 |
|
17261 | 449 |
fun make_record_data |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
450 |
records sel_upd equalities extinjects extsplit splits extfields fieldext = |
17261 | 451 |
{records = records, sel_upd = sel_upd, |
452 |
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
|
453 |
extfields = extfields, fieldext = fieldext }: record_data; |
7178 | 454 |
|
16458 | 455 |
structure RecordsData = TheoryDataFun |
22846 | 456 |
( |
7178 | 457 |
type T = record_data; |
458 |
val empty = |
|
459 |
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
|
460 |
{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
|
461 |
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
|
462 |
foldcong = HOL_basic_ss, unfoldcong = HOL_basic_ss} |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
463 |
Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; |
7178 | 464 |
|
6556 | 465 |
val copy = I; |
16458 | 466 |
val extend = I; |
467 |
fun merge _ |
|
7178 | 468 |
({records = recs1, |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
469 |
sel_upd = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
470 |
{selectors = sels1, updates = upds1, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
471 |
simpset = ss1, defset = ds1, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
472 |
foldcong = fc1, unfoldcong = uc1}, |
14255 | 473 |
equalities = equalities1, |
17261 | 474 |
extinjects = extinjects1, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
475 |
extsplit = extsplit1, |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
476 |
splits = splits1, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
477 |
extfields = extfields1, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
478 |
fieldext = fieldext1}, |
7178 | 479 |
{records = recs2, |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
480 |
sel_upd = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
481 |
{selectors = sels2, updates = upds2, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
482 |
simpset = ss2, defset = ds2, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
483 |
foldcong = fc2, unfoldcong = uc2}, |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
484 |
equalities = equalities2, |
17261 | 485 |
extinjects = extinjects2, |
486 |
extsplit = extsplit2, |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
487 |
splits = splits2, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
488 |
extfields = extfields2, |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
489 |
fieldext = fieldext2}) = |
17261 | 490 |
make_record_data |
7178 | 491 |
(Symtab.merge (K true) (recs1, recs2)) |
492 |
{selectors = Symtab.merge (K true) (sels1, sels2), |
|
493 |
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
|
494 |
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
|
495 |
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
|
496 |
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
|
497 |
unfoldcong = Simplifier.merge_ss (uc1, uc2)} |
22634 | 498 |
(Symtab.merge Thm.eq_thm_prop (equalities1, equalities2)) |
499 |
(Library.merge Thm.eq_thm_prop (extinjects1, extinjects2)) |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
500 |
(Symtab.merge Thm.eq_thm_prop (extsplit1, extsplit2)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
501 |
(Symtab.merge (fn ((a, b, c, d), (w, x, y, z)) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
502 |
Thm.eq_thm (a, w) andalso Thm.eq_thm (b, x) andalso |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
503 |
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
|
504 |
(Symtab.merge (K true) (extfields1, extfields2)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
505 |
(Symtab.merge (K true) (fieldext1, fieldext2)); |
22846 | 506 |
); |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
507 |
|
22846 | 508 |
fun print_records thy = |
509 |
let |
|
510 |
val {records = recs, ...} = RecordsData.get thy; |
|
26943
aec0d97a01c4
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26626
diff
changeset
|
511 |
val prt_typ = Syntax.pretty_typ_global thy; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
512 |
|
22846 | 513 |
fun pretty_parent NONE = [] |
514 |
| pretty_parent (SOME (Ts, name)) = |
|
515 |
[Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]]; |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
516 |
|
22846 | 517 |
fun pretty_field (c, T) = Pretty.block |
518 |
[Pretty.str (Sign.extern_const thy c), Pretty.str " ::", |
|
519 |
Pretty.brk 1, Pretty.quote (prt_typ T)]; |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
520 |
|
22846 | 521 |
fun pretty_record (name, {args, parent, fields, ...}: record_info) = |
522 |
Pretty.block (Pretty.fbreaks (Pretty.block |
|
523 |
[prt_typ (Type (name, map TFree args)), Pretty.str " = "] :: |
|
524 |
pretty_parent parent @ map pretty_field fields)); |
|
525 |
in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end; |
|
5006 | 526 |
|
16458 | 527 |
|
7178 | 528 |
(* access 'records' *) |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
529 |
|
17412 | 530 |
val get_record = Symtab.lookup o #records o RecordsData.get; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
531 |
|
4890 | 532 |
fun put_record name info thy = |
7178 | 533 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
534 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
535 |
RecordsData.get thy; |
17412 | 536 |
val data = make_record_data (Symtab.update (name, info) records) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
537 |
sel_upd equalities extinjects extsplit splits extfields fieldext; |
7178 | 538 |
in RecordsData.put data thy end; |
539 |
||
22846 | 540 |
|
7178 | 541 |
(* access 'sel_upd' *) |
542 |
||
16458 | 543 |
val get_sel_upd = #sel_upd o RecordsData.get; |
7178 | 544 |
|
17510 | 545 |
val is_selector = Symtab.defined o #selectors o get_sel_upd; |
17412 | 546 |
val get_updates = Symtab.lookup o #updates o get_sel_upd; |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
547 |
fun get_ss_with_context getss thy = |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
548 |
Simplifier.theory_context thy (getss (get_sel_upd thy)); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
549 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
550 |
val get_simpset = get_ss_with_context (#simpset); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
551 |
val get_sel_upd_defs = get_ss_with_context (#defset); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
552 |
val get_foldcong_ss = get_ss_with_context (#foldcong); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
553 |
val get_unfoldcong_ss = get_ss_with_context (#unfoldcong); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
554 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
555 |
fun get_update_details u thy = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
556 |
let val sel_upd = get_sel_upd thy in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
557 |
(case Symtab.lookup (#updates sel_upd) u of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
558 |
SOME s => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
559 |
let val SOME (dep, ismore) = Symtab.lookup (#selectors sel_upd) s |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
560 |
in SOME (s, dep, ismore) end |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
561 |
| NONE => NONE) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
562 |
end; |
7178 | 563 |
|
32744
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
564 |
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
|
565 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
566 |
val all = names @ [more]; |
32744
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
567 |
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
|
568 |
val upds = map (suffix updateN) all ~~ all; |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
569 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
570 |
val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong}, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
571 |
equalities, extinjects, extsplit, splits, extfields, fieldext} = RecordsData.get thy; |
32744
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
572 |
val data = make_record_data records |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
573 |
{selectors = fold Symtab.update_new sels selectors, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
574 |
updates = fold Symtab.update_new upds updates, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
575 |
simpset = Simplifier.addsimps (simpset, simps), |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
576 |
defset = Simplifier.addsimps (defset, defs), |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
577 |
foldcong = foldcong addcongs folds, |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
578 |
unfoldcong = unfoldcong addcongs unfolds} |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
579 |
equalities extinjects extsplit splits extfields fieldext; |
50406c4951d9
Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents:
32743
diff
changeset
|
580 |
in RecordsData.put data thy end; |
22846 | 581 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
582 |
|
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
583 |
(* access 'equalities' *) |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
584 |
|
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
585 |
fun add_record_equalities name thm thy = |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
586 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
587 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
588 |
RecordsData.get thy; |
17261 | 589 |
val data = make_record_data records sel_upd |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
590 |
(Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext; |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
591 |
in RecordsData.put data thy end; |
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
592 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
593 |
val get_equalities = Symtab.lookup o #equalities o RecordsData.get; |
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
594 |
|
22846 | 595 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
596 |
(* access 'extinjects' *) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
597 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
598 |
fun add_extinjects thm thy = |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
599 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
600 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
601 |
RecordsData.get thy; |
22846 | 602 |
val data = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
603 |
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
|
604 |
extsplit splits extfields fieldext; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
605 |
in RecordsData.put data thy end; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
606 |
|
22634 | 607 |
val get_extinjects = rev o #extinjects o RecordsData.get; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
608 |
|
22846 | 609 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
610 |
(* access 'extsplit' *) |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
611 |
|
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
612 |
fun add_extsplit name thm thy = |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
613 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
614 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
615 |
RecordsData.get thy; |
17261 | 616 |
val data = make_record_data records sel_upd |
17412 | 617 |
equalities extinjects (Symtab.update_new (name, thm) extsplit) splits |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
618 |
extfields fieldext; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
619 |
in RecordsData.put data thy end; |
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
620 |
|
17412 | 621 |
val get_extsplit = Symtab.lookup o #extsplit o RecordsData.get; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
622 |
|
26088 | 623 |
|
14255 | 624 |
(* access 'splits' *) |
625 |
||
626 |
fun add_record_splits name thmP thy = |
|
627 |
let |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
628 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
629 |
RecordsData.get thy; |
17261 | 630 |
val data = make_record_data records sel_upd |
17412 | 631 |
equalities extinjects extsplit (Symtab.update_new (name, thmP) splits) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
632 |
extfields fieldext; |
14255 | 633 |
in RecordsData.put data thy end; |
634 |
||
17412 | 635 |
val get_splits = Symtab.lookup o #splits o RecordsData.get; |
14255 | 636 |
|
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
637 |
|
26088 | 638 |
(* parent/extension of named record *) |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
639 |
|
26088 | 640 |
val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o RecordsData.get); |
641 |
val get_extension = Option.map #extension oo (Symtab.lookup o #records o RecordsData.get); |
|
17261 | 642 |
|
14079
1c22e5499eeb
- record_split_tac now also works for object-level universal quantifier
berghofe
parents:
13904
diff
changeset
|
643 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
644 |
(* access 'extfields' *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
645 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
646 |
fun add_extfields name fields thy = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
647 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
648 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
649 |
RecordsData.get thy; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
650 |
val data = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
651 |
make_record_data records sel_upd |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
652 |
equalities extinjects extsplit splits |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
653 |
(Symtab.update_new (name, fields) extfields) fieldext; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
654 |
in RecordsData.put data thy end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
655 |
|
17412 | 656 |
val get_extfields = Symtab.lookup o #extfields o RecordsData.get; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
657 |
|
18858 | 658 |
fun get_extT_fields thy T = |
15059 | 659 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
660 |
val ((name, Ts), moreT) = dest_recT T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
661 |
val recname = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
662 |
let val (nm :: recn :: rst) = rev (Long_Name.explode name) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
663 |
in Long_Name.implode (rev (nm :: rst)) end; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
664 |
val midx = maxidx_of_typs (moreT :: Ts); |
19748 | 665 |
val varifyT = varifyT midx; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
666 |
val {records, extfields, ...} = RecordsData.get thy; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
667 |
val (flds, (more, _)) = split_last (Symtab.lookup_list extfields name); |
17412 | 668 |
val args = map varifyT (snd (#extension (the (Symtab.lookup records recname)))); |
15058 | 669 |
|
19748 | 670 |
val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) (Vartab.empty); |
15059 | 671 |
val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
672 |
in (flds', (more, moreT)) end; |
15058 | 673 |
|
18858 | 674 |
fun get_recT_fields thy T = |
17261 | 675 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
676 |
val (root_flds, (root_more, root_moreT)) = get_extT_fields thy T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
677 |
val (rest_flds, rest_more) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
678 |
if is_recT root_moreT then get_recT_fields thy root_moreT |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
679 |
else ([], (root_more, root_moreT)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
680 |
in (root_flds @ rest_flds, rest_more) end; |
15059 | 681 |
|
15058 | 682 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
683 |
(* access 'fieldext' *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
684 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
685 |
fun add_fieldext extname_types fields thy = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
686 |
let |
17261 | 687 |
val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
688 |
RecordsData.get thy; |
17261 | 689 |
val fieldext' = |
17412 | 690 |
fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
691 |
val data = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
692 |
make_record_data records sel_upd equalities extinjects |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
693 |
extsplit splits extfields fieldext'; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
694 |
in RecordsData.put data thy end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
695 |
|
17412 | 696 |
val get_fieldext = Symtab.lookup o #fieldext o RecordsData.get; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
697 |
|
21962 | 698 |
|
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
699 |
(* parent records *) |
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
700 |
|
15531 | 701 |
fun add_parents thy NONE parents = parents |
702 |
| add_parents thy (SOME (types, name)) parents = |
|
12247 | 703 |
let |
704 |
fun err msg = error (msg ^ " parent record " ^ quote name); |
|
12255 | 705 |
|
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
706 |
val {args, parent, fields, extension, induct, extdef} = |
15531 | 707 |
(case get_record thy name of SOME info => info | NONE => err "Unknown"); |
12247 | 708 |
val _ = if length types <> length args then err "Bad number of arguments for" else (); |
12255 | 709 |
|
12247 | 710 |
fun bad_inst ((x, S), T) = |
22578 | 711 |
if Sign.of_sort thy (T, S) then NONE else SOME x |
15570 | 712 |
val bads = List.mapPartial bad_inst (args ~~ types); |
21962 | 713 |
val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in"); |
12255 | 714 |
|
12247 | 715 |
val inst = map fst args ~~ types; |
17377 | 716 |
val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); |
15570 | 717 |
val parent' = Option.map (apfst (map subst)) parent; |
12247 | 718 |
val fields' = map (apsnd subst) fields; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
719 |
val extension' = apsnd (map subst) extension; |
12247 | 720 |
in |
12255 | 721 |
add_parents thy parent' |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
722 |
(make_parent_info name fields' extension' induct extdef :: parents) |
12247 | 723 |
end; |
4867
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
724 |
|
9be2bf0ce909
package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff
changeset
|
725 |
|
21962 | 726 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
727 |
(** concrete syntax for records **) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
728 |
|
22693 | 729 |
(* decode type *) |
730 |
||
731 |
fun decode_type thy t = |
|
732 |
let |
|
23578 | 733 |
fun get_sort xs n = AList.lookup (op =) xs (n: indexname) |> the_default (Sign.defaultS thy); |
22693 | 734 |
val map_sort = Sign.intern_sort thy; |
735 |
in |
|
736 |
Syntax.typ_of_term (get_sort (Syntax.term_sorts map_sort t)) map_sort t |
|
737 |
|> Sign.intern_tycons thy |
|
738 |
end; |
|
739 |
||
740 |
||
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
741 |
(* parse translations *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
742 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
743 |
fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
744 |
if c = mark then Syntax.const (suffix sfx name) $ Abs ("_", dummyT, arg) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
745 |
else raise TERM ("gen_field_tr: " ^ mark, [t]) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
746 |
| gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
747 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
748 |
fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
749 |
if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
750 |
else [gen_field_tr mark sfx tm] |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
751 |
| gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm]; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
752 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
753 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
754 |
fun record_update_tr [t, u] = |
21078 | 755 |
Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
756 |
| record_update_tr ts = raise TERM ("record_update_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
757 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
758 |
fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
759 |
| update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
760 |
| update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
761 |
(c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
762 |
| update_name_tr ts = raise TERM ("update_name_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
763 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
764 |
fun dest_ext_field mark (t as (Const (c, _) $ Const (name, _) $ arg)) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
765 |
if c = mark then (name, arg) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
766 |
else raise TERM ("dest_ext_field: " ^ mark, [t]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
767 |
| dest_ext_field _ t = raise TERM ("dest_ext_field", [t]); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
768 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
769 |
fun dest_ext_fields sep mark (trm as (Const (c, _) $ t $ u)) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
770 |
if c = sep then dest_ext_field mark t :: dest_ext_fields sep mark u |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
771 |
else [dest_ext_field mark trm] |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
772 |
| dest_ext_fields _ mark t = [dest_ext_field mark t]; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
773 |
|
21772 | 774 |
fun gen_ext_fields_tr sep mark sfx more ctxt t = |
17261 | 775 |
let |
21772 | 776 |
val thy = ProofContext.theory_of ctxt; |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
777 |
val msg = "error in record input: "; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
778 |
|
17261 | 779 |
val fieldargs = dest_ext_fields sep mark t; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
780 |
fun splitargs (field :: fields) ((name, arg) :: fargs) = |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
781 |
if can (unsuffix name) field |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
782 |
then |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
783 |
let val (args, rest) = splitargs fields fargs |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
784 |
in (arg :: args, rest) end |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
785 |
else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
786 |
| splitargs [] (fargs as (_ :: _)) = ([], fargs) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
787 |
| splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
788 |
| splitargs _ _ = ([], []); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
789 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
790 |
fun mk_ext (fargs as (name, arg) :: _) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
791 |
(case get_fieldext thy (Sign.intern_const thy name) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
792 |
SOME (ext, _) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
793 |
(case get_extfields thy ext of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
794 |
SOME flds => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
795 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
796 |
val (args, rest) = splitargs (map fst (but_last flds)) fargs; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
797 |
val more' = mk_ext rest; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
798 |
in list_comb (Syntax.const (suffix sfx ext), args @ [more']) end |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
799 |
| NONE => raise TERM (msg ^ "no fields defined for " ^ ext, [t])) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
800 |
| NONE => raise TERM (msg ^ name ^" is no proper field", [t])) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
801 |
| mk_ext [] = more; |
17261 | 802 |
in mk_ext fieldargs end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
803 |
|
21772 | 804 |
fun gen_ext_type_tr sep mark sfx more ctxt t = |
17261 | 805 |
let |
21772 | 806 |
val thy = ProofContext.theory_of ctxt; |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
807 |
val msg = "error in record-type input: "; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
808 |
|
17261 | 809 |
val fieldargs = dest_ext_fields sep mark t; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
810 |
fun splitargs (field :: fields) ((name, arg) :: fargs) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
811 |
if can (unsuffix name) field then |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
812 |
let val (args, rest) = splitargs fields fargs |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
813 |
in (arg :: args, rest) end |
14709
d01983034ded
tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents:
14702
diff
changeset
|
814 |
else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t]) |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
815 |
| splitargs [] (fargs as (_ :: _)) = ([], fargs) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
816 |
| splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
817 |
| splitargs _ _ = ([], []); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
818 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
819 |
fun mk_ext (fargs as (name, arg) :: _) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
820 |
(case get_fieldext thy (Sign.intern_const thy name) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
821 |
SOME (ext, alphas) => |
18858 | 822 |
(case get_extfields thy ext of |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
823 |
SOME flds => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
824 |
(let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
825 |
val flds' = but_last flds; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
826 |
val types = map snd flds'; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
827 |
val (args, rest) = splitargs (map fst flds') fargs; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
828 |
val argtypes = map (Sign.certify_typ thy o decode_type thy) args; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
829 |
val midx = fold (fn T => fn i => Int.max (maxidx_of_typ T, i)) argtypes 0; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
830 |
val varifyT = varifyT midx; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
831 |
val vartypes = map varifyT types; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
832 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
833 |
val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes) Vartab.empty; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
834 |
val alphas' = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
835 |
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
|
836 |
(but_last alphas); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
837 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
838 |
val more' = mk_ext rest; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
839 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
840 |
list_comb (Syntax.const (suffix sfx ext), alphas' @ [more']) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
841 |
end handle TYPE_MATCH => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
842 |
raise TERM (msg ^ "type is no proper record (extension)", [t])) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
843 |
| NONE => raise TERM (msg ^ "no fields defined for " ^ ext, [t])) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
844 |
| NONE => raise TERM (msg ^ name ^" is no proper field", [t])) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
845 |
| mk_ext [] = more; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
846 |
|
17261 | 847 |
in mk_ext fieldargs end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
848 |
|
21772 | 849 |
fun gen_adv_record_tr sep mark sfx unit ctxt [t] = |
850 |
gen_ext_fields_tr sep mark sfx unit ctxt t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
851 |
| gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
852 |
|
21772 | 853 |
fun gen_adv_record_scheme_tr sep mark sfx ctxt [t, more] = |
854 |
gen_ext_fields_tr sep mark sfx more ctxt t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
855 |
| gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
856 |
|
21772 | 857 |
fun gen_adv_record_type_tr sep mark sfx unit ctxt [t] = |
858 |
gen_ext_type_tr sep mark sfx unit ctxt t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
859 |
| gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
860 |
|
21772 | 861 |
fun gen_adv_record_type_scheme_tr sep mark sfx ctxt [t, more] = |
862 |
gen_ext_type_tr sep mark sfx more ctxt t |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
863 |
| gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
864 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
865 |
val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
866 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
867 |
val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
868 |
|
17261 | 869 |
val adv_record_type_tr = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
870 |
gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
871 |
(Syntax.term_of_typ false (HOLogic.unitT)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
872 |
|
17261 | 873 |
val adv_record_type_scheme_tr = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
874 |
gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
875 |
|
15215 | 876 |
|
24867 | 877 |
val parse_translation = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
878 |
[("_record_update", record_update_tr), |
17261 | 879 |
("_update_name", update_name_tr)]; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
880 |
|
24867 | 881 |
val adv_parse_translation = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
882 |
[("_record", adv_record_tr), |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
883 |
("_record_scheme", adv_record_scheme_tr), |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
884 |
("_record_type", adv_record_type_tr), |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
885 |
("_record_type_scheme", adv_record_type_scheme_tr)]; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
886 |
|
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
887 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
888 |
(* print translations *) |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
889 |
|
32740 | 890 |
val print_record_type_abbr = Unsynchronized.ref true; |
891 |
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
|
892 |
|
25705 | 893 |
fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ k $ u) = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
894 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
895 |
val t = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
896 |
(case k of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
897 |
Abs (_, _, Abs (_, _, t) $ Bound 0) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
898 |
if null (loose_bnos t) then t else raise Match |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
899 |
| Abs (x, _, t) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
900 |
if null (loose_bnos t) then t else raise Match |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
901 |
| _ => raise Match); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
902 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
903 |
(* FIXME ? *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
904 |
(* (case k of (Const ("K_record", _) $ t) => t |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
905 |
| Abs (x, _, Const ("K_record", _) $ t $ Bound 0) => t |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
906 |
| _ => raise Match)*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
907 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
908 |
(case try (unsuffix sfx) name_field of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
909 |
SOME name => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
910 |
apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
911 |
| NONE => ([], tm)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
912 |
end |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
913 |
| gen_field_upds_tr' _ _ tm = ([], tm); |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
914 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
915 |
fun record_update_tr' tm = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
916 |
let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
917 |
if null ts then raise Match |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
918 |
else |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
919 |
Syntax.const "_record_update" $ u $ |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
920 |
foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts) |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
921 |
end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
922 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
923 |
fun gen_field_tr' sfx tr' name = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
924 |
let val name_sfx = suffix sfx name |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
925 |
in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
926 |
|
21772 | 927 |
fun record_tr' sep mark record record_scheme unit ctxt t = |
17261 | 928 |
let |
21772 | 929 |
val thy = ProofContext.theory_of ctxt; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
930 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
931 |
fun field_lst t = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
932 |
(case strip_comb t of |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
933 |
(Const (ext, _), args as (_ :: _)) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
934 |
(case try (unsuffix extN) (Sign.intern_const thy ext) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
935 |
SOME ext' => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
936 |
(case get_extfields thy ext' of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
937 |
SOME flds => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
938 |
(let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
939 |
val f :: fs = but_last (map fst flds); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
940 |
val flds' = Sign.extern_const thy f :: map Long_Name.base_name fs; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
941 |
val (args', more) = split_last args; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
942 |
in (flds' ~~ args') @ field_lst more end |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
943 |
handle Library.UnequalLengths => [("", t)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
944 |
| NONE => [("", t)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
945 |
| NONE => [("", t)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
946 |
| _ => [("", t)]); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
947 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
948 |
val (flds, (_, more)) = split_last (field_lst t); |
21226
a607ae87ee81
field-update in records is generalised to take a function on the field
schirmer
parents:
21109
diff
changeset
|
949 |
val _ = if null flds then raise Match else (); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
950 |
val flds' = map (fn (n, t) => Syntax.const mark $ Syntax.const n $ t) flds; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
951 |
val flds'' = foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds'; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
952 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
953 |
if unit more |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
954 |
then Syntax.const record $ flds'' |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
955 |
else Syntax.const record_scheme $ flds'' $ more |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
956 |
end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
957 |
|
17261 | 958 |
fun gen_record_tr' name = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
959 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
960 |
val name_sfx = suffix extN name; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
961 |
val unit = (fn Const (@{const_syntax "Product_Type.Unity"}, _) => true | _ => false); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
962 |
fun tr' ctxt ts = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
963 |
record_tr' "_fields" "_field" "_record" "_record_scheme" unit ctxt |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
964 |
(list_comb (Syntax.const name_sfx, ts)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
965 |
in (name_sfx, tr') end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
966 |
|
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
967 |
fun print_translation names = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
968 |
map (gen_field_tr' updateN record_update_tr') names; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
969 |
|
19748 | 970 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
971 |
(* record_type_abbr_tr' *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
972 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
973 |
(*try to reconstruct the record name type abbreviation from |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
974 |
the (nested) extension types*) |
21772 | 975 |
fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt tm = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
976 |
let |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
977 |
val thy = ProofContext.theory_of ctxt; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
978 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
979 |
(*tm is term representation of a (nested) field type. We first reconstruct the |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
980 |
type from tm so that we can continue on the type level rather then the term level*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
981 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
982 |
(*WORKAROUND: |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
983 |
If a record type occurs in an error message of type inference there |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
984 |
may be some internal frees donoted by ??: |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
985 |
(Const "_tfree",_) $ Free ("??'a", _). |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
986 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
987 |
This will unfortunately be translated to Type ("??'a", []) instead of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
988 |
TFree ("??'a", _) by typ_of_term, which will confuse unify below. |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
989 |
fixT works around.*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
990 |
fun fixT (T as Type (x, [])) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
991 |
if String.isPrefix "??'" x then TFree (x, Sign.defaultS thy) else T |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
992 |
| fixT (Type (x, xs)) = Type (x, map fixT xs) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
993 |
| fixT T = T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
994 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
995 |
val T = fixT (decode_type thy tm); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
996 |
val midx = maxidx_of_typ T; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
997 |
val varifyT = varifyT midx; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
998 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
999 |
fun mk_type_abbr subst name alphas = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1000 |
let val abbrT = Type (name, map (fn a => varifyT (TFree (a, Sign.defaultS thy))) alphas) in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1001 |
Syntax.term_of_typ (! Syntax.show_sorts) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1002 |
(Sign.extern_typ thy (Envir.norm_type subst abbrT)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1003 |
end; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1004 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1005 |
fun match rT T = Sign.typ_match thy (varifyT rT, T) Vartab.empty; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1006 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1007 |
if ! print_record_type_abbr then |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1008 |
(case last_extT T of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1009 |
SOME (name, _) => |
32335 | 1010 |
if name = lastExt then |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1011 |
(let val subst = match schemeT T in |
32335 | 1012 |
if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree (zeta, Sign.defaultS thy)))) |
1013 |
then mk_type_abbr subst abbr alphas |
|
1014 |
else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta]) |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1015 |
end handle TYPE_MATCH => default_tr' ctxt tm) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1016 |
else raise Match (*give print translation of specialised record a chance*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1017 |
| _ => raise Match) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1018 |
else default_tr' ctxt tm |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1019 |
end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1020 |
|
21772 | 1021 |
fun record_type_tr' sep mark record record_scheme ctxt t = |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1022 |
let |
21772 | 1023 |
val thy = ProofContext.theory_of ctxt; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1024 |
|
22693 | 1025 |
val T = decode_type thy t; |
1026 |
val varifyT = varifyT (Term.maxidx_of_typ T); |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1027 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1028 |
fun term_of_type T = Syntax.term_of_typ (! Syntax.show_sorts) (Sign.extern_typ thy T); |
17261 | 1029 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1030 |
fun field_lst T = |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1031 |
(case T of |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1032 |
Type (ext, args) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1033 |
(case try (unsuffix ext_typeN) ext of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1034 |
SOME ext' => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1035 |
(case get_extfields thy ext' of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1036 |
SOME flds => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1037 |
(case get_fieldext thy (fst (hd flds)) of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1038 |
SOME (_, alphas) => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1039 |
(let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1040 |
val f :: fs = but_last flds; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1041 |
val flds' = apfst (Sign.extern_const thy) f :: |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1042 |
map (apfst Long_Name.base_name) fs; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1043 |
val (args', more) = split_last args; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1044 |
val alphavars = map varifyT (but_last alphas); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1045 |
val subst = fold2 (curry (Sign.typ_match thy)) alphavars args' Vartab.empty; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1046 |
val flds'' = (map o apsnd) (Envir.norm_type subst o varifyT) flds'; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1047 |
in flds'' @ field_lst more end |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1048 |
handle TYPE_MATCH => [("", T)] |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1049 |
| Library.UnequalLengths => [("", T)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1050 |
| NONE => [("", T)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1051 |
| NONE => [("", T)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1052 |
| NONE => [("", T)]) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1053 |
| _ => [("", T)]); |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1054 |
|
22219
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
haftmann
parents:
21962
diff
changeset
|
1055 |
val (flds, (_, moreT)) = split_last (field_lst T); |
61b5bab471ce
print translation for record types with empty-sorted type variables raise Match instead of producing an error
haftmann
parents:
21962
diff
changeset
|
1056 |
val flds' = map (fn (n, T) => Syntax.const mark $ Syntax.const n $ term_of_type T) flds; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1057 |
val flds'' = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1058 |
foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds' |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1059 |
handle Empty => raise Match; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1060 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1061 |
if not (! print_record_type_as_fields) orelse null flds then raise Match |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1062 |
else if moreT = HOLogic.unitT then Syntax.const record $ flds'' |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1063 |
else Syntax.const record_scheme $ flds'' $ term_of_type moreT |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1064 |
end; |
17261 | 1065 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1066 |
|
17261 | 1067 |
fun gen_record_type_tr' name = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1068 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1069 |
val name_sfx = suffix ext_typeN name; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1070 |
fun tr' ctxt ts = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1071 |
record_type_tr' "_field_types" "_field_type" "_record_type" "_record_type_scheme" |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1072 |
ctxt (list_comb (Syntax.const name_sfx, ts)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1073 |
in (name_sfx, tr') end; |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1074 |
|
17261 | 1075 |
|
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1076 |
fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1077 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1078 |
val name_sfx = suffix ext_typeN name; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1079 |
val default_tr' = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1080 |
record_type_tr' "_field_types" "_field_type" "_record_type" "_record_type_scheme"; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1081 |
fun tr' ctxt ts = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1082 |
record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1083 |
(list_comb (Syntax.const name_sfx, ts)); |
14700
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1084 |
in (name_sfx, tr') end; |
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
schirmer
parents:
14643
diff
changeset
|
1085 |
|
32335 | 1086 |
|
1087 |
||
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1088 |
(** record simprocs **) |
14358 | 1089 |
|
32740 | 1090 |
val record_quick_and_dirty_sensitive = Unsynchronized.ref false; |
15015
c5768e8c4da4
* record_upd_simproc also simplifies trivial updates:
schirmer
parents:
15012
diff
changeset
|
1091 |
|
15215 | 1092 |
|
18858 | 1093 |
fun quick_and_dirty_prove stndrd thy asms prop tac = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1094 |
if ! record_quick_and_dirty_sensitive andalso ! quick_and_dirty then |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1095 |
Goal.prove (ProofContext.init thy) [] [] |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1096 |
(Logic.list_implies (map Logic.varify asms, Logic.varify prop)) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1097 |
(K (SkipProof.cheat_tac @{theory HOL})) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1098 |
(*Drule.standard can take quite a while for large records, thats why |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1099 |
we varify the proposition manually here.*) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1100 |
else |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1101 |
let val prf = Goal.prove (ProofContext.init thy) [] asms prop tac |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1102 |
in if stndrd then standard prf else prf end; |
15215 | 1103 |
|
17261 | 1104 |
fun quick_and_dirty_prf noopt opt () = |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1105 |
if ! record_quick_and_dirty_sensitive andalso ! quick_and_dirty |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1106 |
then noopt () |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1107 |
else opt (); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1108 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1109 |
fun is_sel_upd_pair thy (Const (s, t)) (Const (u, t')) = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1110 |
(case get_updates thy u of |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1111 |
SOME u_name => u_name = s |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1112 |
| NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')])); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1113 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1114 |
fun mk_comp f g = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1115 |
let |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1116 |
val x = fastype_of g; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1117 |
val a = domain_type x; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1118 |
val b = range_type x; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1119 |
val c = range_type (fastype_of f); |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1120 |
val T = (b --> c) --> ((a --> b) --> (a --> c)) |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1121 |
in Const ("Fun.comp", T) $ f $ g end; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1122 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1123 |
fun mk_comp_id f = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1124 |
let val T = range_type (fastype_of f) |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1125 |
in mk_comp (Const ("Fun.id", T --> T)) f end; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1126 |
|
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
1127 |
fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1128 |
| get_upd_funs _ = []; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1129 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1130 |
fun get_accupd_simps thy term defset intros_tac = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1131 |
let |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1132 |
val (acc, [body]) = strip_comb term; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1133 |
val recT = domain_type (fastype_of acc); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1134 |
val upd_funs = sort_distinct TermOrd.fast_term_ord (get_upd_funs body); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1135 |
fun get_simp upd = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1136 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1137 |
val T = domain_type (fastype_of upd); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1138 |
val lhs = mk_comp acc (upd $ Free ("f", T)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1139 |
val rhs = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1140 |
if is_sel_upd_pair thy acc upd |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1141 |
then mk_comp (Free ("f", T)) acc |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1142 |
else mk_comp_id acc; |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1143 |
val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1144 |
val othm = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1145 |
Goal.prove (ProofContext.init thy) [] [] prop |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1146 |
(fn prems => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1147 |
EVERY |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1148 |
[simp_tac defset 1, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1149 |
REPEAT_DETERM (intros_tac 1), |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1150 |
TRY (simp_tac (HOL_ss addsimps id_o_apps) 1)]); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1151 |
val dest = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1152 |
if is_sel_upd_pair thy acc upd |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1153 |
then o_eq_dest |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1154 |
else o_eq_id_dest; |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1155 |
in standard (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
|
1156 |
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
|
1157 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1158 |
(* FIXME dup? *) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1159 |
structure SymSymTab = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1160 |
Table(type key = string * string val ord = prod_ord fast_string_ord fast_string_ord); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1161 |
|
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1162 |
fun get_updupd_simp thy defset intros_tac u u' comp = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1163 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1164 |
val f = Free ("f", domain_type (fastype_of u)); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1165 |
val f' = Free ("f'", domain_type (fastype_of u')); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1166 |
val lhs = mk_comp (u $ f) (u' $ f'); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1167 |
val rhs = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1168 |
if comp |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1169 |
then u $ mk_comp f f' |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1170 |
else mk_comp (u' $ f') (u $ f); |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1171 |
val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1172 |
val othm = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1173 |
Goal.prove (ProofContext.init thy) [] [] prop |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1174 |
(fn prems => |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1175 |
EVERY |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1176 |
[simp_tac defset 1, |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1177 |
REPEAT_DETERM (intros_tac 1), |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1178 |
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
|
1179 |
val dest = if comp then o_eq_dest_lhs else o_eq_dest; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1180 |
in standard (othm RS dest) end; |
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1181 |
|
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1182 |
fun get_updupd_simps thy term defset intros_tac = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1183 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1184 |
val recT = fastype_of term; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1185 |
val upd_funs = get_upd_funs term; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1186 |
val cname = fst o dest_Const; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1187 |
fun getswap u u' = get_updupd_simp thy defset intros_tac u u' (cname u = cname u'); |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
1188 |
fun build_swaps_to_eq upd [] swaps = swaps |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1189 |
| build_swaps_to_eq upd (u :: us) swaps = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1190 |
let |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1191 |
val key = (cname u, cname upd); |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1192 |
val newswaps = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1193 |
if SymSymTab.defined swaps key then swaps |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1194 |
else SymSymTab.insert (K true) (key, getswap u upd) swaps; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1195 |
in |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1196 |
if cname u = cname upd then newswaps |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1197 |
else build_swaps_to_eq upd us newswaps |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1198 |
end; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1199 |
fun swaps_needed [] prev seen swaps = map snd (SymSymTab.dest swaps) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1200 |
| swaps_needed (u :: us) prev seen swaps = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1201 |
if Symtab.defined seen (cname u) |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1202 |
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
|
1203 |
else swaps_needed us (u :: prev) (Symtab.insert (K true) (cname u, ()) seen) swaps; |
32752
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents:
32749
diff
changeset
|
1204 |
in swaps_needed upd_funs [] Symtab.empty SymSymTab.empty end; |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1205 |
|
32745
192d58483fdf
Simplification of various aspects of the IsTuple component
Thomas Sewell <tsewell@nicta.com.au>
parents:
32744
diff
changeset
|
1206 |
val named_cterm_instantiate = IsTupleSupport.named_cterm_instantiate; |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1207 |
|
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1208 |
fun prove_unfold_defs thy ss T 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
|
1209 |
let |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1210 |
val defset = get_sel_upd_defs thy; |
32745
192d58483fdf
Simplification of various aspects of the IsTuple component
Thomas Sewell <tsewell@nicta.com.au>
parents:
32744
diff
changeset
|
1211 |
val in_tac = IsTupleSupport.istuple_intros_tac thy; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1212 |
val prop' = Envir.beta_eta_contract prop; |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1213 |
val (lhs, rhs) = Logic.dest_equals (Logic.strip_assums_concl prop'); |
32743
c4e9a48bc50e
Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents:
32335
diff
changeset
|
1214 |
val (head, args) = strip_comb lhs; |
32761
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1215 |
val simps = |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1216 |
if length args = 1 |
54fee94b2b29
tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents:
32760
diff
changeset
|
1217 |
then get_accupd_simps thy lhs defset in_tac |
54fee94b2b29
tuned whitespace -- recover basic |