author | wenzelm |
Wed, 24 Jul 2002 22:15:55 +0200 | |
changeset 13421 | 8fcdf4a26468 |
parent 13412 | 666137b488a4 |
child 14080 | 9a50427d7165 |
permissions | -rw-r--r-- |
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Record.thy |
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
3 |
Author: Wolfgang Naraschewski and Markus Wenzel, TU Muenchen |
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
4 |
*) |
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
5 |
|
11821 | 6 |
header {* Extensible records with structural subtyping *} |
7 |
||
11826 | 8 |
theory Record = Product_Type |
11821 | 9 |
files ("Tools/record_package.ML"): |
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
10 |
|
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
11 |
|
11821 | 12 |
subsection {* Abstract product types *} |
13 |
||
13412 | 14 |
locale product_type = |
15 |
fixes Rep and Abs and pair and dest1 and dest2 |
|
16 |
assumes "typedef": "type_definition Rep Abs UNIV" |
|
17 |
and pair: "pair == (\<lambda>a b. Abs (a, b))" |
|
18 |
and dest1: "dest1 == (\<lambda>p. fst (Rep p))" |
|
19 |
and dest2: "dest2 == (\<lambda>p. snd (Rep p))" |
|
11821 | 20 |
|
13412 | 21 |
lemma (in product_type) |
22 |
"inject": "(pair x y = pair x' y') = (x = x' \<and> y = y')" |
|
23 |
by (simp add: pair type_definition.Abs_inject [OF "typedef"]) |
|
11821 | 24 |
|
13412 | 25 |
lemma (in product_type) conv1: "dest1 (pair x y) = x" |
26 |
by (simp add: pair dest1 type_definition.Abs_inverse [OF "typedef"]) |
|
11826 | 27 |
|
13412 | 28 |
lemma (in product_type) conv2: "dest2 (pair x y) = y" |
29 |
by (simp add: pair dest2 type_definition.Abs_inverse [OF "typedef"]) |
|
11821 | 30 |
|
13412 | 31 |
lemma (in product_type) induct [induct type]: |
32 |
assumes hyp: "!!x y. P (pair x y)" |
|
33 |
shows "P p" |
|
34 |
proof (rule type_definition.Abs_induct [OF "typedef"]) |
|
35 |
fix q show "P (Abs q)" |
|
36 |
proof (induct q) |
|
37 |
fix x y have "P (pair x y)" by (rule hyp) |
|
38 |
also have "pair x y = Abs (x, y)" by (simp only: pair) |
|
39 |
finally show "P (Abs (x, y))" . |
|
11821 | 40 |
qed |
41 |
qed |
|
42 |
||
13412 | 43 |
lemma (in product_type) cases [cases type]: |
44 |
"(!!x y. p = pair x y ==> C) ==> C" |
|
45 |
by (induct p) (auto simp add: "inject") |
|
11826 | 46 |
|
13412 | 47 |
lemma (in product_type) surjective_pairing: |
48 |
"p = pair (dest1 p) (dest2 p)" |
|
49 |
by (induct p) (simp only: conv1 conv2) |
|
11826 | 50 |
|
13412 | 51 |
lemma (in product_type) split_paired_all: |
52 |
"(!!x. PROP P x) == (!!a b. PROP P (pair a b))" |
|
11826 | 53 |
proof |
54 |
fix a b |
|
55 |
assume "!!x. PROP P x" |
|
56 |
thus "PROP P (pair a b)" . |
|
57 |
next |
|
58 |
fix x |
|
59 |
assume "!!a b. PROP P (pair a b)" |
|
60 |
hence "PROP P (pair (dest1 x) (dest2 x))" . |
|
13412 | 61 |
thus "PROP P x" by (simp only: surjective_pairing [symmetric]) |
11826 | 62 |
qed |
63 |
||
11821 | 64 |
|
11833 | 65 |
subsection {* Concrete record syntax *} |
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
66 |
|
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
67 |
nonterminals |
5198 | 68 |
ident field_type field_types field fields update updates |
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
69 |
|
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
70 |
syntax |
11821 | 71 |
"_constify" :: "id => ident" ("_") |
72 |
"_constify" :: "longid => ident" ("_") |
|
5198 | 73 |
|
11821 | 74 |
"_field_type" :: "[ident, type] => field_type" ("(2_ ::/ _)") |
75 |
"" :: "field_type => field_types" ("_") |
|
76 |
"_field_types" :: "[field_type, field_types] => field_types" ("_,/ _") |
|
77 |
"_record_type" :: "field_types => type" ("(3'(| _ |'))") |
|
10093 | 78 |
"_record_type_scheme" :: "[field_types, type] => type" ("(3'(| _,/ (2... ::/ _) |'))") |
5198 | 79 |
|
11821 | 80 |
"_field" :: "[ident, 'a] => field" ("(2_ =/ _)") |
81 |
"" :: "field => fields" ("_") |
|
82 |
"_fields" :: "[field, fields] => fields" ("_,/ _") |
|
83 |
"_record" :: "fields => 'a" ("(3'(| _ |'))") |
|
10093 | 84 |
"_record_scheme" :: "[fields, 'a] => 'a" ("(3'(| _,/ (2... =/ _) |'))") |
5198 | 85 |
|
10641 | 86 |
"_update_name" :: idt |
11821 | 87 |
"_update" :: "[ident, 'a] => update" ("(2_ :=/ _)") |
88 |
"" :: "update => updates" ("_") |
|
89 |
"_updates" :: "[update, updates] => updates" ("_,/ _") |
|
10093 | 90 |
"_record_update" :: "['a, updates] => 'b" ("_/(3'(| _ |'))" [900,0] 900) |
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
91 |
|
10331 | 92 |
syntax (xsymbols) |
11821 | 93 |
"_record_type" :: "field_types => type" ("(3\<lparr>_\<rparr>)") |
10093 | 94 |
"_record_type_scheme" :: "[field_types, type] => type" ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)") |
95 |
"_record" :: "fields => 'a" ("(3\<lparr>_\<rparr>)") |
|
96 |
"_record_scheme" :: "[fields, 'a] => 'a" ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)") |
|
97 |
"_record_update" :: "['a, updates] => 'b" ("_/(3\<lparr>_\<rparr>)" [900,0] 900) |
|
9729 | 98 |
|
11833 | 99 |
|
100 |
subsection {* Package setup *} |
|
101 |
||
102 |
use "Tools/record_package.ML" |
|
103 |
setup RecordPackage.setup |
|
10641 | 104 |
|
4870
cc36acb5b114
Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff
changeset
|
105 |
end |