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