author | wenzelm |
Sat, 29 Mar 2008 19:14:06 +0100 | |
changeset 26484 | 28dd7c7a9a2e |
parent 26407 | 562a1d615336 |
child 26962 | c8b20f615d6c |
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 |
ID: $Id$ |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
3 |
Author: Franz Regensburger |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
4 |
|
16070
4a83dd540b88
removed LICENCE note -- everything is subject to Isabelle licence as
wenzelm
parents:
16057
diff
changeset
|
5 |
Partial ordering for cartesian product of HOL products. |
15576
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 cpo of cartesian products *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
9 |
|
15577 | 10 |
theory Cprod |
25910 | 11 |
imports Bifinite |
15577 | 12 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
13 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
14 |
defaultsort cpo |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
15 |
|
16008 | 16 |
subsection {* Type @{typ unit} is a pcpo *} |
17 |
||
26025 | 18 |
instantiation unit :: sq_ord |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
19 |
begin |
16008 | 20 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
21 |
definition |
16008 | 22 |
less_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<equiv> True" |
23 |
||
26025 | 24 |
instance .. |
25 |
end |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
26 |
|
26025 | 27 |
instance unit :: discrete_cpo |
28 |
by intro_classes simp |
|
16008 | 29 |
|
25815 | 30 |
instance unit :: finite_po .. |
16008 | 31 |
|
32 |
instance unit :: pcpo |
|
33 |
by intro_classes simp |
|
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 |
unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
37 |
"unit_when = (\<Lambda> a _. a)" |
16008 | 38 |
|
18289 | 39 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
40 |
"\<Lambda>(). t" == "CONST unit_when\<cdot>t" |
18289 | 41 |
|
42 |
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a" |
|
43 |
by (simp add: unit_when_def) |
|
44 |
||
45 |
||
46 |
subsection {* Product type is a partial order *} |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
47 |
|
26035 | 48 |
instantiation "*" :: (sq_ord, sq_ord) sq_ord |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
49 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
50 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
51 |
definition |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
52 |
less_cprod_def: "(op \<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
53 |
|
26035 | 54 |
instance .. |
55 |
end |
|
56 |
||
57 |
instance "*" :: (po, po) po |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
58 |
proof |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
59 |
fix x :: "'a \<times> 'b" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
60 |
show "x \<sqsubseteq> x" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
61 |
unfolding less_cprod_def by simp |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
62 |
next |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
63 |
fix x y :: "'a \<times> 'b" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
64 |
assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
65 |
unfolding less_cprod_def Pair_fst_snd_eq |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
66 |
by (fast intro: antisym_less) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
67 |
next |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
68 |
fix x y z :: "'a \<times> 'b" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
69 |
assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
70 |
unfolding less_cprod_def |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
71 |
by (fast intro: trans_less) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
72 |
qed |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
73 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
74 |
subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
75 |
|
26029 | 76 |
lemma prod_lessI: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q" |
77 |
unfolding less_cprod_def by simp |
|
78 |
||
79 |
lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) = (a \<sqsubseteq> c \<and> b \<sqsubseteq> d)" |
|
80 |
unfolding less_cprod_def by simp |
|
81 |
||
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
82 |
text {* Pair @{text "(_,_)"} is monotone in both arguments *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
83 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
84 |
lemma monofun_pair1: "monofun (\<lambda>x. (x, y))" |
26029 | 85 |
by (simp add: monofun_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
86 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
87 |
lemma monofun_pair2: "monofun (\<lambda>y. (x, y))" |
26029 | 88 |
by (simp add: monofun_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
89 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
90 |
lemma monofun_pair: |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
91 |
"\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)" |
26029 | 92 |
by simp |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
93 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
94 |
text {* @{term fst} and @{term snd} are monotone *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
95 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
96 |
lemma monofun_fst: "monofun fst" |
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
|
97 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
98 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
99 |
lemma monofun_snd: "monofun snd" |
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
|
100 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
101 |
|
18289 | 102 |
subsection {* Product type is a cpo *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
103 |
|
26018 | 104 |
lemma is_lub_Pair: |
105 |
"\<lbrakk>range X <<| x; range Y <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>i. (X i, Y i)) <<| (x, y)" |
|
106 |
apply (rule is_lubI [OF ub_rangeI]) |
|
107 |
apply (simp add: less_cprod_def is_ub_lub) |
|
108 |
apply (frule ub2ub_monofun [OF monofun_fst]) |
|
109 |
apply (drule ub2ub_monofun [OF monofun_snd]) |
|
110 |
apply (simp add: less_cprod_def is_lub_lub) |
|
111 |
done |
|
112 |
||
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
113 |
lemma lub_cprod: |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
114 |
fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
115 |
assumes S: "chain S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
116 |
shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
26018 | 117 |
proof - |
118 |
have "chain (\<lambda>i. fst (S i))" |
|
119 |
using monofun_fst S by (rule ch2ch_monofun) |
|
120 |
hence 1: "range (\<lambda>i. fst (S i)) <<| (\<Squnion>i. fst (S i))" |
|
26027 | 121 |
by (rule cpo_lubI) |
26018 | 122 |
have "chain (\<lambda>i. snd (S i))" |
123 |
using monofun_snd S by (rule ch2ch_monofun) |
|
124 |
hence 2: "range (\<lambda>i. snd (S i)) <<| (\<Squnion>i. snd (S i))" |
|
26027 | 125 |
by (rule cpo_lubI) |
26018 | 126 |
show "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
127 |
using is_lub_Pair [OF 1 2] by simp |
|
128 |
qed |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
129 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
130 |
lemma thelub_cprod: |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
131 |
"chain (S::nat \<Rightarrow> 'a::cpo \<times> 'b::cpo) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
132 |
\<Longrightarrow> lub (range S) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
133 |
by (rule lub_cprod [THEN thelubI]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
134 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
135 |
instance "*" :: (cpo, cpo) cpo |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
136 |
proof |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
137 |
fix S :: "nat \<Rightarrow> ('a \<times> 'b)" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
138 |
assume "chain S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
139 |
hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
140 |
by (rule lub_cprod) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
141 |
thus "\<exists>x. range S <<| x" .. |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
142 |
qed |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
143 |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
144 |
instance "*" :: (finite_po, finite_po) finite_po .. |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
145 |
|
26025 | 146 |
instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo |
147 |
proof |
|
148 |
fix x y :: "'a \<times> 'b" |
|
149 |
show "x \<sqsubseteq> y \<longleftrightarrow> x = y" |
|
150 |
unfolding less_cprod_def Pair_fst_snd_eq |
|
151 |
by simp |
|
152 |
qed |
|
153 |
||
18289 | 154 |
subsection {* Product type is pointed *} |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
155 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
156 |
lemma minimal_cprod: "(\<bottom>, \<bottom>) \<sqsubseteq> p" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
157 |
by (simp add: less_cprod_def) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
158 |
|
25908 | 159 |
instance "*" :: (pcpo, pcpo) pcpo |
160 |
by intro_classes (fast intro: minimal_cprod) |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
161 |
|
25908 | 162 |
lemma inst_cprod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
163 |
by (rule minimal_cprod [THEN UU_I, symmetric]) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
164 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
165 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
166 |
subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *} |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
167 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
168 |
lemma cont_pair1: "cont (\<lambda>x. (x, y))" |
26018 | 169 |
apply (rule contI) |
170 |
apply (rule is_lub_Pair) |
|
26027 | 171 |
apply (erule cpo_lubI) |
26018 | 172 |
apply (rule lub_const) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
173 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
174 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
175 |
lemma cont_pair2: "cont (\<lambda>y. (x, y))" |
26018 | 176 |
apply (rule contI) |
177 |
apply (rule is_lub_Pair) |
|
178 |
apply (rule lub_const) |
|
26027 | 179 |
apply (erule cpo_lubI) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
180 |
done |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
181 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
182 |
lemma contlub_fst: "contlub fst" |
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
|
183 |
apply (rule contlubI) |
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
|
184 |
apply (simp add: thelub_cprod) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
185 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
186 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
187 |
lemma contlub_snd: "contlub snd" |
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
|
188 |
apply (rule contlubI) |
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
|
189 |
apply (simp add: thelub_cprod) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
190 |
done |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
191 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
192 |
lemma cont_fst: "cont fst" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
193 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
194 |
apply (rule monofun_fst) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
195 |
apply (rule contlub_fst) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
196 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
197 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
198 |
lemma cont_snd: "cont snd" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
199 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
200 |
apply (rule monofun_snd) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
201 |
apply (rule contlub_snd) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
202 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
203 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
204 |
subsection {* Continuous versions of constants *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
205 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
206 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
207 |
cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)" -- {* continuous pairing *} where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
208 |
"cpair = (\<Lambda> x y. (x, y))" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
209 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
210 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
211 |
cfst :: "('a * 'b) \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
212 |
"cfst = (\<Lambda> p. fst p)" |
17834 | 213 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
214 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
215 |
csnd :: "('a * 'b) \<rightarrow> 'b" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
216 |
"csnd = (\<Lambda> p. snd p)" |
17834 | 217 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
218 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
219 |
csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
220 |
"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
|
221 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
222 |
syntax |
17834 | 223 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1<_,/ _>)") |
224 |
||
225 |
syntax (xsymbols) |
|
226 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1\<langle>_,/ _\<rangle>)") |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
227 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
228 |
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
|
229 |
"\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
230 |
"\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y" |
17834 | 231 |
|
17816
9942c5ed866a
new syntax translations for continuous lambda abstraction
huffman
parents:
16916
diff
changeset
|
232 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
233 |
"\<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
|
234 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
235 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
236 |
subsection {* Convert all lemmas to the continuous versions *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
237 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
238 |
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
|
239 |
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
|
240 |
|
25910 | 241 |
lemma pair_eq_cpair: "(x, y) = <x, y>" |
242 |
by (simp add: cpair_def cont_pair1 cont_pair2) |
|
243 |
||
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
244 |
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
|
245 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
246 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
247 |
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
|
248 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
249 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
250 |
lemma cpair_less [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
251 |
by (simp add: cpair_eq_pair less_cprod_def) |
16057
e23a61b3406f
added lemma cpair_eq, made cfst_strict and csnd_strict into simp rules
huffman
parents:
16008
diff
changeset
|
252 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
253 |
lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)" |
16916 | 254 |
by (simp add: inst_cprod_pcpo cpair_eq_pair) |
255 |
||
25913 | 256 |
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
|
257 |
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
|
258 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
259 |
lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>" |
16916 | 260 |
by (rule cpair_strict [symmetric]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
261 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
262 |
lemma defined_cpair_rev: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
263 |
"<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
|
264 |
by simp |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
265 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
266 |
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
|
267 |
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
|
268 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
269 |
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
|
270 |
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
|
271 |
|
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
|
272 |
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
|
273 |
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
|
274 |
|
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
|
275 |
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
|
276 |
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
|
277 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
278 |
lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>" |
25913 | 279 |
unfolding inst_cprod_pcpo2 by (rule cfst_cpair) |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
280 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
281 |
lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>" |
25913 | 282 |
unfolding inst_cprod_pcpo2 by (rule csnd_cpair) |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
283 |
|
25910 | 284 |
lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p" |
285 |
by (cases p rule: cprodE, simp) |
|
286 |
||
287 |
lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
288 |
|
16750 | 289 |
lemma less_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)" |
16315 | 290 |
by (simp add: less_cprod_def cfst_def csnd_def cont_fst cont_snd) |
291 |
||
16750 | 292 |
lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)" |
293 |
by (auto simp add: po_eq_conv less_cprod) |
|
294 |
||
25879 | 295 |
lemma cfst_less_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>" |
296 |
by (simp add: less_cprod) |
|
297 |
||
298 |
lemma csnd_less_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>" |
|
299 |
by (simp add: less_cprod) |
|
300 |
||
301 |
lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)" |
|
302 |
by (rule compactI, simp add: cfst_less_iff) |
|
303 |
||
304 |
lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)" |
|
305 |
by (rule compactI, simp add: csnd_less_iff) |
|
306 |
||
307 |
lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>" |
|
17837 | 308 |
by (rule compactI, simp add: less_cprod) |
309 |
||
25879 | 310 |
lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)" |
311 |
apply (safe intro!: compact_cpair) |
|
312 |
apply (drule compact_cfst, simp) |
|
313 |
apply (drule compact_csnd, simp) |
|
314 |
done |
|
315 |
||
25905 | 316 |
instance "*" :: (chfin, chfin) chfin |
25921 | 317 |
apply intro_classes |
25905 | 318 |
apply (erule compact_imp_max_in_chain) |
319 |
apply (rule_tac p="\<Squnion>i. Y i" in cprodE, simp) |
|
320 |
done |
|
321 |
||
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
322 |
lemma lub_cprod2: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
323 |
"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
|
324 |
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
|
325 |
apply (erule lub_cprod) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
326 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
327 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
328 |
lemma thelub_cprod2: |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
329 |
"chain S \<Longrightarrow> lub (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
|
330 |
by (rule lub_cprod2 [THEN thelubI]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
331 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
332 |
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
|
333 |
by (simp add: csplit_def) |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
334 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
335 |
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
|
336 |
by (simp add: csplit_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
337 |
|
16553 | 338 |
lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z" |
25910 | 339 |
by (simp add: csplit_def cpair_cfst_csnd) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
340 |
|
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
|
341 |
lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
342 |
|
25910 | 343 |
subsection {* Product type is a bifinite domain *} |
344 |
||
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
26035
diff
changeset
|
345 |
instance "*" :: (profinite, profinite) approx .. |
25910 | 346 |
|
347 |
defs (overloaded) |
|
348 |
approx_cprod_def: |
|
349 |
"approx \<equiv> \<lambda>n. \<Lambda>\<langle>x, y\<rangle>. \<langle>approx n\<cdot>x, approx n\<cdot>y\<rangle>" |
|
350 |
||
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
26035
diff
changeset
|
351 |
instance "*" :: (profinite, profinite) profinite |
25910 | 352 |
proof |
353 |
fix i :: nat and x :: "'a \<times> 'b" |
|
354 |
show "chain (\<lambda>i. approx i\<cdot>x)" |
|
355 |
unfolding approx_cprod_def by simp |
|
356 |
show "(\<Squnion>i. approx i\<cdot>x) = x" |
|
357 |
unfolding approx_cprod_def |
|
358 |
by (simp add: lub_distribs eta_cfun) |
|
359 |
show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
|
360 |
unfolding approx_cprod_def csplit_def by simp |
|
361 |
have "{x::'a \<times> 'b. approx i\<cdot>x = x} \<subseteq> |
|
362 |
{x::'a. approx i\<cdot>x = x} \<times> {x::'b. approx i\<cdot>x = x}" |
|
363 |
unfolding approx_cprod_def |
|
364 |
by (clarsimp simp add: pair_eq_cpair) |
|
365 |
thus "finite {x::'a \<times> 'b. approx i\<cdot>x = x}" |
|
366 |
by (rule finite_subset, |
|
367 |
intro finite_cartesian_product finite_fixes_approx) |
|
368 |
qed |
|
369 |
||
370 |
instance "*" :: (bifinite, bifinite) bifinite .. |
|
371 |
||
372 |
lemma approx_cpair [simp]: |
|
373 |
"approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>" |
|
374 |
unfolding approx_cprod_def by simp |
|
375 |
||
376 |
lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)" |
|
377 |
by (cases p rule: cprodE, simp) |
|
378 |
||
379 |
lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)" |
|
380 |
by (cases p rule: cprodE, simp) |
|
381 |
||
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
382 |
end |