author | huffman |
Fri, 08 May 2009 16:19:51 -0700 | |
changeset 31076 | 99fe356cbbc2 |
parent 29535 | 08824fad8879 |
child 31113 | 15cf300a742f |
permissions | -rw-r--r-- |
15600 | 1 |
(* Title: HOLCF/Cprod.thy |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
2 |
Author: Franz Regensburger |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
3 |
*) |
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 |
header {* The cpo of cartesian products *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
6 |
|
15577 | 7 |
theory Cprod |
25910 | 8 |
imports Bifinite |
15577 | 9 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
10 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
11 |
defaultsort cpo |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
12 |
|
16008 | 13 |
subsection {* Type @{typ unit} is a pcpo *} |
14 |
||
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
15 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
16 |
unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
17 |
"unit_when = (\<Lambda> a _. a)" |
16008 | 18 |
|
18289 | 19 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
20 |
"\<Lambda>(). t" == "CONST unit_when\<cdot>t" |
18289 | 21 |
|
22 |
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a" |
|
23 |
by (simp add: unit_when_def) |
|
24 |
||
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
25 |
subsection {* Continuous versions of constants *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
26 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
27 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
28 |
cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)" -- {* continuous pairing *} where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
29 |
"cpair = (\<Lambda> x y. (x, y))" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
30 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
31 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
32 |
cfst :: "('a * 'b) \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
33 |
"cfst = (\<Lambda> p. fst p)" |
17834 | 34 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
35 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
36 |
csnd :: "('a * 'b) \<rightarrow> 'b" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
37 |
"csnd = (\<Lambda> p. snd p)" |
17834 | 38 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
39 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
40 |
csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
41 |
"csplit = (\<Lambda> f p. f\<cdot>(cfst\<cdot>p)\<cdot>(csnd\<cdot>p))" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
42 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
43 |
syntax |
17834 | 44 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1<_,/ _>)") |
45 |
||
46 |
syntax (xsymbols) |
|
47 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1\<langle>_,/ _\<rangle>)") |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
48 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
49 |
translations |
18078
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents:
18077
diff
changeset
|
50 |
"\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
51 |
"\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y" |
17834 | 52 |
|
17816
9942c5ed866a
new syntax translations for continuous lambda abstraction
huffman
parents:
16916
diff
changeset
|
53 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
54 |
"\<Lambda>(CONST cpair\<cdot>x\<cdot>y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)" |
17816
9942c5ed866a
new syntax translations for continuous lambda abstraction
huffman
parents:
16916
diff
changeset
|
55 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
56 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
57 |
subsection {* Convert all lemmas to the continuous versions *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
58 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
59 |
lemma cpair_eq_pair: "<x, y> = (x, y)" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
60 |
by (simp add: cpair_def cont_pair1 cont_pair2) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
61 |
|
25910 | 62 |
lemma pair_eq_cpair: "(x, y) = <x, y>" |
63 |
by (simp add: cpair_def cont_pair1 cont_pair2) |
|
64 |
||
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
65 |
lemma inject_cpair: "<a,b> = <aa,ba> \<Longrightarrow> a = aa \<and> b = ba" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
66 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
67 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
68 |
lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' \<and> b = b')" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
69 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
70 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
71 |
lemma cpair_below [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')" |
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
72 |
by (simp add: cpair_eq_pair) |
16057
e23a61b3406f
added lemma cpair_eq, made cfst_strict and csnd_strict into simp rules
huffman
parents:
16008
diff
changeset
|
73 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
74 |
lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)" |
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
75 |
by (simp add: cpair_eq_pair) |
16916 | 76 |
|
25913 | 77 |
lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>" |
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
78 |
by simp |
16210
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
huffman
parents:
16093
diff
changeset
|
79 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
80 |
lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>" |
16916 | 81 |
by (rule cpair_strict [symmetric]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
82 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
83 |
lemma defined_cpair_rev: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
84 |
"<a,b> = \<bottom> \<Longrightarrow> a = \<bottom> \<and> b = \<bottom>" |
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
85 |
by simp |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
86 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
87 |
lemma Exh_Cprod2: "\<exists>a b. z = <a, b>" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
88 |
by (simp add: cpair_eq_pair) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
89 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
90 |
lemma cprodE: "\<lbrakk>\<And>x y. p = <x, y> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
91 |
by (cut_tac Exh_Cprod2, auto) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
92 |
|
16210
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
huffman
parents:
16093
diff
changeset
|
93 |
lemma cfst_cpair [simp]: "cfst\<cdot><x, y> = x" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
94 |
by (simp add: cpair_eq_pair cfst_def cont_fst) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
95 |
|
16210
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
huffman
parents:
16093
diff
changeset
|
96 |
lemma csnd_cpair [simp]: "csnd\<cdot><x, y> = y" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
97 |
by (simp add: cpair_eq_pair csnd_def cont_snd) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
98 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
99 |
lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>" |
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
100 |
by (simp add: cfst_def) |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
101 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
102 |
lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>" |
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
103 |
by (simp add: csnd_def) |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
104 |
|
25910 | 105 |
lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p" |
106 |
by (cases p rule: cprodE, simp) |
|
107 |
||
108 |
lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
109 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
110 |
lemma below_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
111 |
by (simp add: below_prod_def cfst_def csnd_def cont_fst cont_snd) |
16315 | 112 |
|
16750 | 113 |
lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)" |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
114 |
by (auto simp add: po_eq_conv below_cprod) |
16750 | 115 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
116 |
lemma cfst_below_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
117 |
by (simp add: below_cprod) |
25879 | 118 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
119 |
lemma csnd_below_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
120 |
by (simp add: below_cprod) |
25879 | 121 |
|
122 |
lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
123 |
by (rule compactI, simp add: cfst_below_iff) |
25879 | 124 |
|
125 |
lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29535
diff
changeset
|
126 |
by (rule compactI, simp add: csnd_below_iff) |
25879 | 127 |
|
128 |
lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>" |
|
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
129 |
by (simp add: cpair_eq_pair) |
17837 | 130 |
|
25879 | 131 |
lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)" |
29535
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents:
29530
diff
changeset
|
132 |
by (simp add: cpair_eq_pair) |
25905 | 133 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
134 |
lemma lub_cprod2: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
135 |
"chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
136 |
apply (simp add: cpair_eq_pair cfst_def csnd_def cont_fst cont_snd) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
137 |
apply (erule lub_cprod) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
138 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
139 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
140 |
lemma thelub_cprod2: |
27413 | 141 |
"chain S \<Longrightarrow> (\<Squnion>i. S i) = <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
142 |
by (rule lub_cprod2 [THEN thelubI]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
143 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
144 |
lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>" |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
145 |
by (simp add: csplit_def) |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
146 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
147 |
lemma csplit2 [simp]: "csplit\<cdot>f\<cdot><x,y> = f\<cdot>x\<cdot>y" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
148 |
by (simp add: csplit_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
149 |
|
16553 | 150 |
lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z" |
25910 | 151 |
by (simp add: csplit_def cpair_cfst_csnd) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
152 |
|
16210
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
huffman
parents:
16093
diff
changeset
|
153 |
lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
154 |
|
25910 | 155 |
subsection {* Product type is a bifinite domain *} |
156 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
157 |
instantiation "*" :: (profinite, profinite) profinite |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
158 |
begin |
25910 | 159 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
160 |
definition |
25910 | 161 |
approx_cprod_def: |
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
162 |
"approx = (\<lambda>n. \<Lambda>\<langle>x, y\<rangle>. \<langle>approx n\<cdot>x, approx n\<cdot>y\<rangle>)" |
25910 | 163 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
164 |
instance proof |
25910 | 165 |
fix i :: nat and x :: "'a \<times> 'b" |
27310 | 166 |
show "chain (approx :: nat \<Rightarrow> 'a \<times> 'b \<rightarrow> 'a \<times> 'b)" |
25910 | 167 |
unfolding approx_cprod_def by simp |
168 |
show "(\<Squnion>i. approx i\<cdot>x) = x" |
|
169 |
unfolding approx_cprod_def |
|
170 |
by (simp add: lub_distribs eta_cfun) |
|
171 |
show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
|
172 |
unfolding approx_cprod_def csplit_def by simp |
|
173 |
have "{x::'a \<times> 'b. approx i\<cdot>x = x} \<subseteq> |
|
174 |
{x::'a. approx i\<cdot>x = x} \<times> {x::'b. approx i\<cdot>x = x}" |
|
175 |
unfolding approx_cprod_def |
|
176 |
by (clarsimp simp add: pair_eq_cpair) |
|
177 |
thus "finite {x::'a \<times> 'b. approx i\<cdot>x = x}" |
|
178 |
by (rule finite_subset, |
|
179 |
intro finite_cartesian_product finite_fixes_approx) |
|
180 |
qed |
|
181 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
182 |
end |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
183 |
|
25910 | 184 |
instance "*" :: (bifinite, bifinite) bifinite .. |
185 |
||
186 |
lemma approx_cpair [simp]: |
|
187 |
"approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>" |
|
188 |
unfolding approx_cprod_def by simp |
|
189 |
||
190 |
lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)" |
|
191 |
by (cases p rule: cprodE, simp) |
|
192 |
||
193 |
lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)" |
|
194 |
by (cases p rule: cprodE, simp) |
|
195 |
||
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
196 |
end |