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