author | wenzelm |
Thu, 28 Aug 2008 19:29:56 +0200 | |
changeset 28043 | 4d05f04cc671 |
parent 27310 | d0229bc6c461 |
child 29063 | 7619f0561cd7 |
permissions | -rw-r--r-- |
15600 | 1 |
(* Title: HOLCF/Sprod.thy |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
2 |
ID: $Id$ |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
3 |
Author: Franz Regensburger and Brian Huffman |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
4 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
5 |
Strict product with typedef. |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
6 |
*) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
7 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
8 |
header {* The type of strict products *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
9 |
|
15577 | 10 |
theory Sprod |
16699 | 11 |
imports Cprod |
15577 | 12 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
13 |
|
16082 | 14 |
defaultsort pcpo |
15 |
||
15591
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
16 |
subsection {* Definition of strict product type *} |
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
17 |
|
17817 | 18 |
pcpodef (Sprod) ('a, 'b) "**" (infixr "**" 20) = |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
19 |
"{p::'a \<times> 'b. p = \<bottom> \<or> (cfst\<cdot>p \<noteq> \<bottom> \<and> csnd\<cdot>p \<noteq> \<bottom>)}" |
16699 | 20 |
by simp |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
21 |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
22 |
instance "**" :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
23 |
by (rule typedef_finite_po [OF type_definition_Sprod]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
24 |
|
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
25 |
instance "**" :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
26 |
by (rule typedef_chfin [OF type_definition_Sprod less_Sprod_def]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
27 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
28 |
syntax (xsymbols) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
29 |
"**" :: "[type, type] => type" ("(_ \<otimes>/ _)" [21,20] 20) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
30 |
syntax (HTML output) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
31 |
"**" :: "[type, type] => type" ("(_ \<otimes>/ _)" [21,20] 20) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
32 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
33 |
lemma spair_lemma: |
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
34 |
"<strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a> \<in> Sprod" |
25914 | 35 |
by (simp add: Sprod_def strictify_conv_if) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
36 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
37 |
subsection {* Definitions of constants *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
38 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
39 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
40 |
sfst :: "('a ** 'b) \<rightarrow> 'a" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
41 |
"sfst = (\<Lambda> p. cfst\<cdot>(Rep_Sprod p))" |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
42 |
|
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
43 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
44 |
ssnd :: "('a ** 'b) \<rightarrow> 'b" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
45 |
"ssnd = (\<Lambda> p. csnd\<cdot>(Rep_Sprod p))" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
46 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
47 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
48 |
spair :: "'a \<rightarrow> 'b \<rightarrow> ('a ** 'b)" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
49 |
"spair = (\<Lambda> a b. Abs_Sprod |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
50 |
<strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a>)" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
51 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
52 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
53 |
ssplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a ** 'b) \<rightarrow> 'c" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
54 |
"ssplit = (\<Lambda> f. strictify\<cdot>(\<Lambda> p. f\<cdot>(sfst\<cdot>p)\<cdot>(ssnd\<cdot>p)))" |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
55 |
|
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
25131
diff
changeset
|
56 |
syntax |
18078
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents:
17837
diff
changeset
|
57 |
"@stuple" :: "['a, args] => 'a ** 'b" ("(1'(:_,/ _:'))") |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
58 |
translations |
18078
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents:
17837
diff
changeset
|
59 |
"(:x, y, z:)" == "(:x, (:y, z:):)" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18078
diff
changeset
|
60 |
"(:x, y:)" == "CONST spair\<cdot>x\<cdot>y" |
18078
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents:
17837
diff
changeset
|
61 |
|
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents:
17837
diff
changeset
|
62 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18078
diff
changeset
|
63 |
"\<Lambda>(CONST spair\<cdot>x\<cdot>y). t" == "CONST ssplit\<cdot>(\<Lambda> x y. t)" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
64 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
65 |
subsection {* Case analysis *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
66 |
|
25914 | 67 |
lemma Rep_Sprod_spair: |
68 |
"Rep_Sprod (:a, b:) = <strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a>" |
|
69 |
unfolding spair_def |
|
70 |
by (simp add: cont_Abs_Sprod Abs_Sprod_inverse spair_lemma) |
|
71 |
||
72 |
lemmas Rep_Sprod_simps = |
|
73 |
Rep_Sprod_inject [symmetric] less_Sprod_def |
|
74 |
Rep_Sprod_strict Rep_Sprod_spair |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
75 |
|
27310 | 76 |
lemma Exh_Sprod: |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
77 |
"z = \<bottom> \<or> (\<exists>a b. z = (:a, b:) \<and> a \<noteq> \<bottom> \<and> b \<noteq> \<bottom>)" |
25914 | 78 |
apply (insert Rep_Sprod [of z]) |
79 |
apply (simp add: Rep_Sprod_simps eq_cprod) |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
80 |
apply (simp add: Sprod_def) |
25914 | 81 |
apply (erule disjE, simp) |
82 |
apply (simp add: strictify_conv_if) |
|
83 |
apply fast |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
84 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
85 |
|
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
86 |
lemma sprodE [cases type: **]: |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
87 |
"\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x y. \<lbrakk>p = (:x, y:); x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
27310 | 88 |
by (cut_tac z=p in Exh_Sprod, auto) |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
89 |
|
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
90 |
lemma sprod_induct [induct type: **]: |
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
91 |
"\<lbrakk>P \<bottom>; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> P (:x, y:)\<rbrakk> \<Longrightarrow> P x" |
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
92 |
by (cases x, simp_all) |
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
93 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
94 |
subsection {* Properties of @{term spair} *} |
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
95 |
|
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
96 |
lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>" |
25914 | 97 |
by (simp add: Rep_Sprod_simps strictify_conv_if) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
98 |
|
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
99 |
lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>" |
25914 | 100 |
by (simp add: Rep_Sprod_simps strictify_conv_if) |
101 |
||
102 |
lemma spair_strict_iff [simp]: "((:x, y:) = \<bottom>) = (x = \<bottom> \<or> y = \<bottom>)" |
|
103 |
by (simp add: Rep_Sprod_simps strictify_conv_if) |
|
104 |
||
105 |
lemma spair_less_iff: |
|
106 |
"((:a, b:) \<sqsubseteq> (:c, d:)) = (a = \<bottom> \<or> b = \<bottom> \<or> (a \<sqsubseteq> c \<and> b \<sqsubseteq> d))" |
|
107 |
by (simp add: Rep_Sprod_simps strictify_conv_if) |
|
108 |
||
109 |
lemma spair_eq_iff: |
|
110 |
"((:a, b:) = (:c, d:)) = |
|
111 |
(a = c \<and> b = d \<or> (a = \<bottom> \<or> b = \<bottom>) \<and> (c = \<bottom> \<or> d = \<bottom>))" |
|
112 |
by (simp add: Rep_Sprod_simps strictify_conv_if) |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
113 |
|
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
114 |
lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>" |
25914 | 115 |
by simp |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
116 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
117 |
lemma spair_strict_rev: "(:x, y:) \<noteq> \<bottom> \<Longrightarrow> x \<noteq> \<bottom> \<and> y \<noteq> \<bottom>" |
25914 | 118 |
by simp |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
119 |
|
25914 | 120 |
lemma spair_defined: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<noteq> \<bottom>" |
121 |
by simp |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
122 |
|
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
123 |
lemma spair_defined_rev: "(:x, y:) = \<bottom> \<Longrightarrow> x = \<bottom> \<or> y = \<bottom>" |
25914 | 124 |
by simp |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
125 |
|
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
126 |
lemma spair_eq: |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
127 |
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ((:x, y:) = (:a, b:)) = (x = a \<and> y = b)" |
25914 | 128 |
by (simp add: spair_eq_iff) |
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
129 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
130 |
lemma spair_inject: |
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
131 |
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; (:x, y:) = (:a, b:)\<rbrakk> \<Longrightarrow> x = a \<and> y = b" |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
132 |
by (rule spair_eq [THEN iffD1]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
133 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
134 |
lemma inst_sprod_pcpo2: "UU = (:UU,UU:)" |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
135 |
by simp |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
136 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
137 |
subsection {* Properties of @{term sfst} and @{term ssnd} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
138 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
139 |
lemma sfst_strict [simp]: "sfst\<cdot>\<bottom> = \<bottom>" |
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
140 |
by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
141 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
142 |
lemma ssnd_strict [simp]: "ssnd\<cdot>\<bottom> = \<bottom>" |
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
143 |
by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_strict) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
144 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
145 |
lemma sfst_spair [simp]: "y \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>(:x, y:) = x" |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
146 |
by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_spair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
147 |
|
16212
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
huffman
parents:
16082
diff
changeset
|
148 |
lemma ssnd_spair [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>(:x, y:) = y" |
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
149 |
by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_spair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
150 |
|
16777
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
151 |
lemma sfst_defined_iff [simp]: "(sfst\<cdot>p = \<bottom>) = (p = \<bottom>)" |
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
152 |
by (cases p, simp_all) |
16777
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
153 |
|
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
154 |
lemma ssnd_defined_iff [simp]: "(ssnd\<cdot>p = \<bottom>) = (p = \<bottom>)" |
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
155 |
by (cases p, simp_all) |
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
156 |
|
16777
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
157 |
lemma sfst_defined: "p \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>p \<noteq> \<bottom>" |
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
158 |
by simp |
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
159 |
|
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
160 |
lemma ssnd_defined: "p \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>p \<noteq> \<bottom>" |
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
161 |
by simp |
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
huffman
parents:
16751
diff
changeset
|
162 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
163 |
lemma surjective_pairing_Sprod2: "(:sfst\<cdot>p, ssnd\<cdot>p:) = p" |
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
164 |
by (cases p, simp_all) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
165 |
|
16751 | 166 |
lemma less_sprod: "x \<sqsubseteq> y = (sfst\<cdot>x \<sqsubseteq> sfst\<cdot>y \<and> ssnd\<cdot>x \<sqsubseteq> ssnd\<cdot>y)" |
16699 | 167 |
apply (simp add: less_Sprod_def sfst_def ssnd_def cont_Rep_Sprod) |
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
168 |
apply (rule less_cprod) |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
169 |
done |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
170 |
|
16751 | 171 |
lemma eq_sprod: "(x = y) = (sfst\<cdot>x = sfst\<cdot>y \<and> ssnd\<cdot>x = ssnd\<cdot>y)" |
172 |
by (auto simp add: po_eq_conv less_sprod) |
|
173 |
||
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
174 |
lemma spair_less: |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
175 |
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a \<and> y \<sqsubseteq> b)" |
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
176 |
apply (cases "a = \<bottom>", simp) |
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
177 |
apply (cases "b = \<bottom>", simp) |
16317
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
178 |
apply (simp add: less_sprod) |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
179 |
done |
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
huffman
parents:
16212
diff
changeset
|
180 |
|
25881 | 181 |
lemma sfst_less_iff: "sfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> (:y, ssnd\<cdot>x:)" |
182 |
apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp) |
|
183 |
apply (simp add: less_sprod) |
|
184 |
done |
|
185 |
||
186 |
lemma ssnd_less_iff: "ssnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> (:sfst\<cdot>x, y:)" |
|
187 |
apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp) |
|
188 |
apply (simp add: less_sprod) |
|
189 |
done |
|
190 |
||
191 |
subsection {* Compactness *} |
|
192 |
||
193 |
lemma compact_sfst: "compact x \<Longrightarrow> compact (sfst\<cdot>x)" |
|
194 |
by (rule compactI, simp add: sfst_less_iff) |
|
195 |
||
196 |
lemma compact_ssnd: "compact x \<Longrightarrow> compact (ssnd\<cdot>x)" |
|
197 |
by (rule compactI, simp add: ssnd_less_iff) |
|
198 |
||
199 |
lemma compact_spair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (:x, y:)" |
|
200 |
by (rule compact_Sprod, simp add: Rep_Sprod_spair strictify_conv_if) |
|
201 |
||
202 |
lemma compact_spair_iff: |
|
203 |
"compact (:x, y:) = (x = \<bottom> \<or> y = \<bottom> \<or> (compact x \<and> compact y))" |
|
204 |
apply (safe elim!: compact_spair) |
|
205 |
apply (drule compact_sfst, simp) |
|
206 |
apply (drule compact_ssnd, simp) |
|
207 |
apply simp |
|
208 |
apply simp |
|
209 |
done |
|
210 |
||
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
211 |
subsection {* Properties of @{term ssplit} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
212 |
|
16059
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
huffman
parents:
15930
diff
changeset
|
213 |
lemma ssplit1 [simp]: "ssplit\<cdot>f\<cdot>\<bottom> = \<bottom>" |
15591
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
214 |
by (simp add: ssplit_def) |
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
215 |
|
16920 | 216 |
lemma ssplit2 [simp]: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ssplit\<cdot>f\<cdot>(:x, y:) = f\<cdot>x\<cdot>y" |
15591
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
217 |
by (simp add: ssplit_def) |
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
218 |
|
16553 | 219 |
lemma ssplit3 [simp]: "ssplit\<cdot>spair\<cdot>z = z" |
25757
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
huffman
parents:
25135
diff
changeset
|
220 |
by (cases z, simp_all) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
221 |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
222 |
subsection {* Strict product preserves flatness *} |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
223 |
|
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
224 |
instance "**" :: (flat, flat) flat |
27310 | 225 |
proof |
226 |
fix x y :: "'a \<otimes> 'b" |
|
227 |
assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y" |
|
228 |
apply (induct x, simp) |
|
229 |
apply (induct y, simp) |
|
230 |
apply (simp add: spair_less_iff flat_less_iff) |
|
231 |
done |
|
232 |
qed |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25757
diff
changeset
|
233 |
|
25914 | 234 |
subsection {* Strict product is a bifinite domain *} |
235 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
236 |
instantiation "**" :: (bifinite, bifinite) bifinite |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
237 |
begin |
25914 | 238 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
239 |
definition |
25914 | 240 |
approx_sprod_def: |
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
241 |
"approx = (\<lambda>n. \<Lambda>(:x, y:). (:approx n\<cdot>x, approx n\<cdot>y:))" |
25914 | 242 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
243 |
instance proof |
25914 | 244 |
fix i :: nat and x :: "'a \<otimes> 'b" |
27310 | 245 |
show "chain (approx :: nat \<Rightarrow> 'a \<otimes> 'b \<rightarrow> 'a \<otimes> 'b)" |
25914 | 246 |
unfolding approx_sprod_def by simp |
247 |
show "(\<Squnion>i. approx i\<cdot>x) = x" |
|
248 |
unfolding approx_sprod_def |
|
249 |
by (simp add: lub_distribs eta_cfun) |
|
250 |
show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
|
251 |
unfolding approx_sprod_def |
|
252 |
by (simp add: ssplit_def strictify_conv_if) |
|
253 |
have "Rep_Sprod ` {x::'a \<otimes> 'b. approx i\<cdot>x = x} \<subseteq> {x. approx i\<cdot>x = x}" |
|
254 |
unfolding approx_sprod_def |
|
27310 | 255 |
apply (clarify, case_tac x) |
25914 | 256 |
apply (simp add: Rep_Sprod_strict) |
257 |
apply (simp add: Rep_Sprod_spair spair_eq_iff) |
|
258 |
done |
|
259 |
hence "finite (Rep_Sprod ` {x::'a \<otimes> 'b. approx i\<cdot>x = x})" |
|
260 |
using finite_fixes_approx by (rule finite_subset) |
|
261 |
thus "finite {x::'a \<otimes> 'b. approx i\<cdot>x = x}" |
|
262 |
by (rule finite_imageD, simp add: inj_on_def Rep_Sprod_inject) |
|
263 |
qed |
|
264 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
265 |
end |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
25914
diff
changeset
|
266 |
|
25914 | 267 |
lemma approx_spair [simp]: |
268 |
"approx i\<cdot>(:x, y:) = (:approx i\<cdot>x, approx i\<cdot>y:)" |
|
269 |
unfolding approx_sprod_def |
|
270 |
by (simp add: ssplit_def strictify_conv_if) |
|
271 |
||
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
272 |
end |