author | huffman |
Fri, 04 Jan 2008 00:01:02 +0100 | |
changeset 25827 | c2adeb1bae5c |
parent 25815 | c7b1e7b7087b |
child 25879 | 98b93782c3b1 |
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 |
11 |
imports Cfun |
|
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 |
||
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
18 |
instantiation unit :: sq_ord |
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 |
||
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
24 |
instance .. |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
25 |
end |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
26 |
|
16008 | 27 |
instance unit :: po |
28 |
by intro_classes simp_all |
|
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 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
48 |
instantiation "*" :: (sq_ord, sq_ord) sq_ord |
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 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
54 |
instance .. |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
55 |
end |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
56 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
57 |
instance "*" :: (po, po) po |
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 |
|
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 |
subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
76 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
77 |
text {* Pair @{text "(_,_)"} is monotone in both arguments *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
78 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
79 |
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
|
80 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
81 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
82 |
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
|
83 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
84 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
85 |
lemma monofun_pair: |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
86 |
"\<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
|
87 |
by (simp add: less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
88 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
89 |
text {* @{term fst} and @{term snd} are monotone *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
90 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
91 |
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
|
92 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
93 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
94 |
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
|
95 |
by (simp add: monofun_def less_cprod_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
96 |
|
18289 | 97 |
subsection {* Product type is a cpo *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
98 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
99 |
lemma lub_cprod: |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
100 |
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
|
101 |
assumes S: "chain S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
102 |
shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
103 |
apply (rule is_lubI) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
104 |
apply (rule ub_rangeI) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
105 |
apply (rule_tac t = "S i" in surjective_pairing [THEN ssubst]) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
106 |
apply (rule monofun_pair) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
107 |
apply (rule is_ub_thelub) |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
108 |
apply (rule ch2ch_monofun [OF monofun_fst S]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
109 |
apply (rule is_ub_thelub) |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
110 |
apply (rule ch2ch_monofun [OF monofun_snd S]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
111 |
apply (rule_tac t = "u" in surjective_pairing [THEN ssubst]) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
112 |
apply (rule monofun_pair) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
113 |
apply (rule is_lub_thelub) |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
114 |
apply (rule ch2ch_monofun [OF monofun_fst S]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
115 |
apply (erule monofun_fst [THEN ub2ub_monofun]) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
116 |
apply (rule is_lub_thelub) |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
117 |
apply (rule ch2ch_monofun [OF monofun_snd S]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
118 |
apply (erule monofun_snd [THEN ub2ub_monofun]) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
119 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
120 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
121 |
lemma directed_lub_cprod: |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
122 |
fixes S :: "('a::dcpo \<times> 'b::dcpo) set" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
123 |
assumes S: "directed S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
124 |
shows "S <<| (\<Squnion>x\<in>S. fst x, \<Squnion>x\<in>S. snd x)" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
125 |
apply (rule is_lubI) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
126 |
apply (rule is_ubI) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
127 |
apply (rule_tac t=x in surjective_pairing [THEN ssubst]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
128 |
apply (rule monofun_pair) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
129 |
apply (erule is_ub_thelub' [OF dir2dir_monofun [OF monofun_fst S] imageI]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
130 |
apply (erule is_ub_thelub' [OF dir2dir_monofun [OF monofun_snd S] imageI]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
131 |
apply (rule_tac t=u in surjective_pairing [THEN ssubst]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
132 |
apply (rule monofun_pair) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
133 |
apply (rule is_lub_thelub') |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
134 |
apply (rule dir2dir_monofun [OF monofun_fst S]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
135 |
apply (erule ub2ub_monofun' [OF monofun_fst]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
136 |
apply (rule is_lub_thelub') |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
137 |
apply (rule dir2dir_monofun [OF monofun_snd S]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
138 |
apply (erule ub2ub_monofun' [OF monofun_snd]) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
139 |
done |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
140 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
141 |
lemma thelub_cprod: |
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
142 |
"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
|
143 |
\<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
|
144 |
by (rule lub_cprod [THEN thelubI]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
145 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
146 |
instance "*" :: (cpo, cpo) cpo |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
147 |
proof |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
148 |
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
|
149 |
assume "chain S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
150 |
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
|
151 |
by (rule lub_cprod) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
152 |
thus "\<exists>x. range S <<| x" .. |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
153 |
qed |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
154 |
|
25784
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
155 |
instance "*" :: (dcpo, dcpo) dcpo |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
156 |
proof |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
157 |
fix S :: "('a \<times> 'b) set" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
158 |
assume "directed S" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
159 |
hence "S <<| (\<Squnion>x\<in>S. fst x, \<Squnion>x\<in>S. snd x)" |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
160 |
by (rule directed_lub_cprod) |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
161 |
thus "\<exists>x. S <<| x" .. |
71157f7e671e
update instance proofs for sq_ord, po; new instance proofs for dcpo
huffman
parents:
25131
diff
changeset
|
162 |
qed |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
163 |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
164 |
instance "*" :: (finite_po, finite_po) finite_po .. |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
165 |
|
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
166 |
instance "*" :: (chfin, chfin) chfin |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
167 |
proof (intro_classes, clarify) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
168 |
fix Y :: "nat \<Rightarrow> 'a \<times> 'b" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
169 |
assume Y: "chain Y" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
170 |
from Y have "chain (\<lambda>i. fst (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
171 |
by (rule ch2ch_monofun [OF monofun_fst]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
172 |
hence "\<exists>m. max_in_chain m (\<lambda>i. fst (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
173 |
by (rule chfin [rule_format]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
174 |
then obtain m where m: "max_in_chain m (\<lambda>i. fst (Y i))" .. |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
175 |
from Y have "chain (\<lambda>i. snd (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
176 |
by (rule ch2ch_monofun [OF monofun_snd]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
177 |
hence "\<exists>n. max_in_chain n (\<lambda>i. snd (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
178 |
by (rule chfin [rule_format]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
179 |
then obtain n where n: "max_in_chain n (\<lambda>i. snd (Y i))" .. |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
180 |
from m have m': "max_in_chain (max m n) (\<lambda>i. fst (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
181 |
by (rule maxinch_mono [OF _ le_maxI1]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
182 |
from n have n': "max_in_chain (max m n) (\<lambda>i. snd (Y i))" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
183 |
by (rule maxinch_mono [OF _ le_maxI2]) |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
184 |
from m' n' have "max_in_chain (max m n) Y" |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
185 |
unfolding max_in_chain_def Pair_fst_snd_eq by fast |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
186 |
thus "\<exists>n. max_in_chain n Y" .. |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
187 |
qed |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25815
diff
changeset
|
188 |
|
18289 | 189 |
subsection {* Product type is pointed *} |
15593
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 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
|
192 |
by (simp add: less_cprod_def) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
193 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
194 |
lemma least_cprod: "EX x::'a::pcpo * 'b::pcpo. ALL y. x \<sqsubseteq> y" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
195 |
apply (rule_tac x = "(\<bottom>, \<bottom>)" in exI) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
196 |
apply (rule minimal_cprod [THEN allI]) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
197 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
198 |
|
15609 | 199 |
instance "*" :: (pcpo, pcpo) pcpo |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
200 |
by intro_classes (rule least_cprod) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
201 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
202 |
text {* for compatibility with old HOLCF-Version *} |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
203 |
lemma inst_cprod_pcpo: "UU = (UU,UU)" |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
204 |
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
|
205 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
206 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
207 |
subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *} |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
208 |
|
16916 | 209 |
lemma contlub_pair1: "contlub (\<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
|
210 |
apply (rule contlubI) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
211 |
apply (subst thelub_cprod) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
212 |
apply (erule monofun_pair1 [THEN ch2ch_monofun]) |
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
213 |
apply simp |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
214 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
215 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
216 |
lemma contlub_pair2: "contlub (\<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
|
217 |
apply (rule contlubI) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
218 |
apply (subst thelub_cprod) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
219 |
apply (erule monofun_pair2 [THEN ch2ch_monofun]) |
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
220 |
apply simp |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
221 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
222 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
223 |
lemma cont_pair1: "cont (\<lambda>x. (x, y))" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
224 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
225 |
apply (rule monofun_pair1) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
226 |
apply (rule contlub_pair1) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
227 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
228 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
229 |
lemma cont_pair2: "cont (\<lambda>y. (x, y))" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
230 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
231 |
apply (rule monofun_pair2) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
232 |
apply (rule contlub_pair2) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
233 |
done |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
234 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
235 |
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
|
236 |
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
|
237 |
apply (simp add: thelub_cprod) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
238 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
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 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
|
241 |
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
|
242 |
apply (simp add: thelub_cprod) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
243 |
done |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
244 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
245 |
lemma cont_fst: "cont fst" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
246 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
247 |
apply (rule monofun_fst) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
248 |
apply (rule contlub_fst) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
249 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
250 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
251 |
lemma cont_snd: "cont snd" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
252 |
apply (rule monocontlub2cont) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
253 |
apply (rule monofun_snd) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
254 |
apply (rule contlub_snd) |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
255 |
done |
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
256 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
257 |
subsection {* Continuous versions of constants *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
258 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
259 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
260 |
cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)" -- {* continuous pairing *} where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
261 |
"cpair = (\<Lambda> x y. (x, y))" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
262 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
263 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
264 |
cfst :: "('a * 'b) \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
265 |
"cfst = (\<Lambda> p. fst p)" |
17834 | 266 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
267 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
268 |
csnd :: "('a * 'b) \<rightarrow> 'b" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
269 |
"csnd = (\<Lambda> p. snd p)" |
17834 | 270 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
271 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
272 |
csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
273 |
"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
|
274 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
275 |
syntax |
17834 | 276 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1<_,/ _>)") |
277 |
||
278 |
syntax (xsymbols) |
|
279 |
"_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1\<langle>_,/ _\<rangle>)") |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
280 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
281 |
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
|
282 |
"\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
283 |
"\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y" |
17834 | 284 |
|
17816
9942c5ed866a
new syntax translations for continuous lambda abstraction
huffman
parents:
16916
diff
changeset
|
285 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
286 |
"\<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
|
287 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
288 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
289 |
subsection {* Convert all lemmas to the continuous versions *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
290 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
291 |
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
|
292 |
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
|
293 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
294 |
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
|
295 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
296 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
297 |
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
|
298 |
by (simp add: cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
299 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
300 |
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
|
301 |
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
|
302 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
303 |
lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)" |
16916 | 304 |
by (simp add: inst_cprod_pcpo cpair_eq_pair) |
305 |
||
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
|
306 |
lemma cpair_strict: "<\<bottom>, \<bottom>> = \<bottom>" |
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
307 |
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
|
308 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
309 |
lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>" |
16916 | 310 |
by (rule cpair_strict [symmetric]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
311 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
312 |
lemma defined_cpair_rev: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
313 |
"<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
|
314 |
by simp |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
315 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
316 |
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
|
317 |
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
|
318 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
319 |
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
|
320 |
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
|
321 |
|
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
|
322 |
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
|
323 |
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
|
324 |
|
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
|
325 |
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
|
326 |
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
|
327 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
328 |
lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
329 |
by (simp add: inst_cprod_pcpo2) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
330 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
331 |
lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
332 |
by (simp add: inst_cprod_pcpo2) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
333 |
|
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
334 |
lemma surjective_pairing_Cprod2: "<cfst\<cdot>p, csnd\<cdot>p> = p" |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
335 |
apply (unfold cfst_def csnd_def) |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
336 |
apply (simp add: cont_fst cont_snd cpair_eq_pair) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
337 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
338 |
|
16750 | 339 |
lemma less_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)" |
16315 | 340 |
by (simp add: less_cprod_def cfst_def csnd_def cont_fst cont_snd) |
341 |
||
16750 | 342 |
lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)" |
343 |
by (auto simp add: po_eq_conv less_cprod) |
|
344 |
||
17837 | 345 |
lemma compact_cpair [simp]: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>" |
346 |
by (rule compactI, simp add: less_cprod) |
|
347 |
||
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
348 |
lemma lub_cprod2: |
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
349 |
"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
|
350 |
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
|
351 |
apply (erule lub_cprod) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
352 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
353 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
354 |
lemma thelub_cprod2: |
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
355 |
"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
|
356 |
by (rule lub_cprod2 [THEN thelubI]) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
357 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
358 |
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
|
359 |
by (simp add: csplit_def) |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
360 |
|
16081
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
huffman
parents:
16070
diff
changeset
|
361 |
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
|
362 |
by (simp add: csplit_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
363 |
|
16553 | 364 |
lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
365 |
by (simp add: csplit_def surjective_pairing_Cprod2) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
366 |
|
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
|
367 |
lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
368 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
369 |
end |