| author | haftmann | 
| Fri, 06 Mar 2009 11:10:57 +0100 | |
| changeset 30307 | 6c74ef5a349f | 
| parent 29252 | ea97aa6aeba2 | 
| child 31076 | 99fe356cbbc2 | 
| permissions | -rw-r--r-- | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOLCF/Deflation.thy  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
2  | 
Author: Brian Huffman  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
5  | 
header {* Continuous Deflations and Embedding-Projection Pairs *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
7  | 
theory Deflation  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
8  | 
imports Cfun  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
11  | 
defaultsort cpo  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
13  | 
subsection {* Continuous deflations *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
15  | 
locale deflation =  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
16  | 
fixes d :: "'a \<rightarrow> 'a"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
17  | 
assumes idem: "\<And>x. d\<cdot>(d\<cdot>x) = d\<cdot>x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
18  | 
assumes less: "\<And>x. d\<cdot>x \<sqsubseteq> x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
19  | 
begin  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
21  | 
lemma less_ID: "d \<sqsubseteq> ID"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
22  | 
by (rule less_cfun_ext, simp add: less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
24  | 
text {* The set of fixed points is the same as the range. *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
26  | 
lemma fixes_eq_range: "{x. d\<cdot>x = x} = range (\<lambda>x. d\<cdot>x)"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
27  | 
by (auto simp add: eq_sym_conv idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
29  | 
lemma range_eq_fixes: "range (\<lambda>x. d\<cdot>x) = {x. d\<cdot>x = x}"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
30  | 
by (auto simp add: eq_sym_conv idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
31  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
32  | 
text {*
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
33  | 
The pointwise ordering on deflation functions coincides with  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
34  | 
the subset ordering of their sets of fixed-points.  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
35  | 
*}  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
37  | 
lemma lessI:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
38  | 
assumes f: "\<And>x. d\<cdot>x = x \<Longrightarrow> f\<cdot>x = x" shows "d \<sqsubseteq> f"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
39  | 
proof (rule less_cfun_ext)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
40  | 
fix x  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
41  | 
from less have "f\<cdot>(d\<cdot>x) \<sqsubseteq> f\<cdot>x" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
42  | 
also from idem have "f\<cdot>(d\<cdot>x) = d\<cdot>x" by (rule f)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
43  | 
finally show "d\<cdot>x \<sqsubseteq> f\<cdot>x" .  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
44  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
46  | 
lemma lessD: "\<lbrakk>f \<sqsubseteq> d; f\<cdot>x = x\<rbrakk> \<Longrightarrow> d\<cdot>x = x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
47  | 
proof (rule antisym_less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
48  | 
from less show "d\<cdot>x \<sqsubseteq> x" .  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
49  | 
next  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
50  | 
assume "f \<sqsubseteq> d"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
51  | 
hence "f\<cdot>x \<sqsubseteq> d\<cdot>x" by (rule monofun_cfun_fun)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
52  | 
also assume "f\<cdot>x = x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
53  | 
finally show "x \<sqsubseteq> d\<cdot>x" .  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
54  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
56  | 
end  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
58  | 
lemma adm_deflation: "adm (\<lambda>d. deflation d)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
59  | 
by (simp add: deflation_def)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
61  | 
lemma deflation_ID: "deflation ID"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
62  | 
by (simp add: deflation.intro)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
64  | 
lemma deflation_UU: "deflation \<bottom>"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
65  | 
by (simp add: deflation.intro)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
67  | 
lemma deflation_less_iff:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
68  | 
"\<lbrakk>deflation p; deflation q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q \<longleftrightarrow> (\<forall>x. p\<cdot>x = x \<longrightarrow> q\<cdot>x = x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
69  | 
apply safe  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
70  | 
apply (simp add: deflation.lessD)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
71  | 
apply (simp add: deflation.lessI)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
72  | 
done  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
74  | 
text {*
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
75  | 
The composition of two deflations is equal to  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
76  | 
the lesser of the two (if they are comparable).  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
77  | 
*}  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
79  | 
lemma deflation_less_comp1:  | 
| 28611 | 80  | 
assumes "deflation f"  | 
81  | 
assumes "deflation g"  | 
|
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
82  | 
shows "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>(g\<cdot>x) = f\<cdot>x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
83  | 
proof (rule antisym_less)  | 
| 29237 | 84  | 
interpret g: deflation g by fact  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
85  | 
from g.less show "f\<cdot>(g\<cdot>x) \<sqsubseteq> f\<cdot>x" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
86  | 
next  | 
| 29237 | 87  | 
interpret f: deflation f by fact  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
88  | 
assume "f \<sqsubseteq> g" hence "f\<cdot>x \<sqsubseteq> g\<cdot>x" by (rule monofun_cfun_fun)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
89  | 
hence "f\<cdot>(f\<cdot>x) \<sqsubseteq> f\<cdot>(g\<cdot>x)" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
90  | 
also have "f\<cdot>(f\<cdot>x) = f\<cdot>x" by (rule f.idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
91  | 
finally show "f\<cdot>x \<sqsubseteq> f\<cdot>(g\<cdot>x)" .  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
92  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
94  | 
lemma deflation_less_comp2:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
95  | 
"\<lbrakk>deflation f; deflation g; f \<sqsubseteq> g\<rbrakk> \<Longrightarrow> g\<cdot>(f\<cdot>x) = f\<cdot>x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
96  | 
by (simp only: deflation.lessD deflation.idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
99  | 
subsection {* Deflations with finite range *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
101  | 
lemma finite_range_imp_finite_fixes:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
102  | 
  "finite (range f) \<Longrightarrow> finite {x. f x = x}"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
103  | 
proof -  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
104  | 
  have "{x. f x = x} \<subseteq> range f"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
105  | 
by (clarify, erule subst, rule rangeI)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
106  | 
moreover assume "finite (range f)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
107  | 
  ultimately show "finite {x. f x = x}"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
108  | 
by (rule finite_subset)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
109  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
111  | 
locale finite_deflation = deflation +  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
112  | 
  assumes finite_fixes: "finite {x. d\<cdot>x = x}"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
113  | 
begin  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
114  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
115  | 
lemma finite_range: "finite (range (\<lambda>x. d\<cdot>x))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
116  | 
by (simp add: range_eq_fixes finite_fixes)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
118  | 
lemma finite_image: "finite ((\<lambda>x. d\<cdot>x) ` A)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
119  | 
by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
121  | 
lemma compact: "compact (d\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
122  | 
proof (rule compactI2)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
123  | 
fix Y :: "nat \<Rightarrow> 'a"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
124  | 
assume Y: "chain Y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
125  | 
have "finite_chain (\<lambda>i. d\<cdot>(Y i))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
126  | 
proof (rule finite_range_imp_finch)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
127  | 
show "chain (\<lambda>i. d\<cdot>(Y i))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
128  | 
using Y by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
129  | 
have "range (\<lambda>i. d\<cdot>(Y i)) \<subseteq> range (\<lambda>x. d\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
130  | 
by clarsimp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
131  | 
thus "finite (range (\<lambda>i. d\<cdot>(Y i)))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
132  | 
using finite_range by (rule finite_subset)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
133  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
134  | 
hence "\<exists>j. (\<Squnion>i. d\<cdot>(Y i)) = d\<cdot>(Y j)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
135  | 
by (simp add: finite_chain_def maxinch_is_thelub Y)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
136  | 
then obtain j where j: "(\<Squnion>i. d\<cdot>(Y i)) = d\<cdot>(Y j)" ..  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
138  | 
assume "d\<cdot>x \<sqsubseteq> (\<Squnion>i. Y i)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
139  | 
hence "d\<cdot>(d\<cdot>x) \<sqsubseteq> d\<cdot>(\<Squnion>i. Y i)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
140  | 
by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
141  | 
hence "d\<cdot>x \<sqsubseteq> (\<Squnion>i. d\<cdot>(Y i))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
142  | 
by (simp add: contlub_cfun_arg Y idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
143  | 
hence "d\<cdot>x \<sqsubseteq> d\<cdot>(Y j)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
144  | 
using j by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
145  | 
hence "d\<cdot>x \<sqsubseteq> Y j"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
146  | 
using less by (rule trans_less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
147  | 
thus "\<exists>j. d\<cdot>x \<sqsubseteq> Y j" ..  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
148  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
150  | 
end  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
153  | 
subsection {* Continuous embedding-projection pairs *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
155  | 
locale ep_pair =  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
156  | 
fixes e :: "'a \<rightarrow> 'b" and p :: "'b \<rightarrow> 'a"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
157  | 
assumes e_inverse [simp]: "\<And>x. p\<cdot>(e\<cdot>x) = x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
158  | 
and e_p_less: "\<And>y. e\<cdot>(p\<cdot>y) \<sqsubseteq> y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
159  | 
begin  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
161  | 
lemma e_less_iff [simp]: "e\<cdot>x \<sqsubseteq> e\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
162  | 
proof  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
163  | 
assume "e\<cdot>x \<sqsubseteq> e\<cdot>y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
164  | 
hence "p\<cdot>(e\<cdot>x) \<sqsubseteq> p\<cdot>(e\<cdot>y)" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
165  | 
thus "x \<sqsubseteq> y" by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
166  | 
next  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
167  | 
assume "x \<sqsubseteq> y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
168  | 
thus "e\<cdot>x \<sqsubseteq> e\<cdot>y" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
169  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
170  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
171  | 
lemma e_eq_iff [simp]: "e\<cdot>x = e\<cdot>y \<longleftrightarrow> x = y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
172  | 
unfolding po_eq_conv e_less_iff ..  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
174  | 
lemma p_eq_iff:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
175  | 
"\<lbrakk>e\<cdot>(p\<cdot>x) = x; e\<cdot>(p\<cdot>y) = y\<rbrakk> \<Longrightarrow> p\<cdot>x = p\<cdot>y \<longleftrightarrow> x = y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
176  | 
by (safe, erule subst, erule subst, simp)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
177  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
178  | 
lemma p_inverse: "(\<exists>x. y = e\<cdot>x) = (e\<cdot>(p\<cdot>y) = y)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
179  | 
by (auto, rule exI, erule sym)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
180  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
181  | 
lemma e_less_iff_less_p: "e\<cdot>x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> p\<cdot>y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
182  | 
proof  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
183  | 
assume "e\<cdot>x \<sqsubseteq> y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
184  | 
then have "p\<cdot>(e\<cdot>x) \<sqsubseteq> p\<cdot>y" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
185  | 
then show "x \<sqsubseteq> p\<cdot>y" by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
186  | 
next  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
187  | 
assume "x \<sqsubseteq> p\<cdot>y"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
188  | 
then have "e\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>y)" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
189  | 
then show "e\<cdot>x \<sqsubseteq> y" using e_p_less by (rule trans_less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
190  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
192  | 
lemma compact_e_rev: "compact (e\<cdot>x) \<Longrightarrow> compact x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
193  | 
proof -  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
194  | 
assume "compact (e\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
195  | 
hence "adm (\<lambda>y. \<not> e\<cdot>x \<sqsubseteq> y)" by (rule compactD)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
196  | 
hence "adm (\<lambda>y. \<not> e\<cdot>x \<sqsubseteq> e\<cdot>y)" by (rule adm_subst [OF cont_Rep_CFun2])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
197  | 
hence "adm (\<lambda>y. \<not> x \<sqsubseteq> y)" by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
198  | 
thus "compact x" by (rule compactI)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
199  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
200  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
201  | 
lemma compact_e: "compact x \<Longrightarrow> compact (e\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
202  | 
proof -  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
203  | 
assume "compact x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
204  | 
hence "adm (\<lambda>y. \<not> x \<sqsubseteq> y)" by (rule compactD)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
205  | 
hence "adm (\<lambda>y. \<not> x \<sqsubseteq> p\<cdot>y)" by (rule adm_subst [OF cont_Rep_CFun2])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
206  | 
hence "adm (\<lambda>y. \<not> e\<cdot>x \<sqsubseteq> y)" by (simp add: e_less_iff_less_p)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
207  | 
thus "compact (e\<cdot>x)" by (rule compactI)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
208  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
209  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
210  | 
lemma compact_e_iff: "compact (e\<cdot>x) \<longleftrightarrow> compact x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
211  | 
by (rule iffI [OF compact_e_rev compact_e])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
212  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
213  | 
text {* Deflations from ep-pairs *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
215  | 
lemma deflation_e_p: "deflation (e oo p)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
216  | 
by (simp add: deflation.intro e_p_less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
217  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
218  | 
lemma deflation_e_d_p:  | 
| 28611 | 219  | 
assumes "deflation d"  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
220  | 
shows "deflation (e oo d oo p)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
221  | 
proof  | 
| 29237 | 222  | 
interpret deflation d by fact  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
223  | 
fix x :: 'b  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
224  | 
show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
225  | 
by (simp add: idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
226  | 
show "(e oo d oo p)\<cdot>x \<sqsubseteq> x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
227  | 
by (simp add: e_less_iff_less_p less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
228  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
230  | 
lemma finite_deflation_e_d_p:  | 
| 28611 | 231  | 
assumes "finite_deflation d"  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
232  | 
shows "finite_deflation (e oo d oo p)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
233  | 
proof  | 
| 29237 | 234  | 
interpret finite_deflation d by fact  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
235  | 
fix x :: 'b  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
236  | 
show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
237  | 
by (simp add: idem)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
238  | 
show "(e oo d oo p)\<cdot>x \<sqsubseteq> x"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
239  | 
by (simp add: e_less_iff_less_p less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
240  | 
have "finite ((\<lambda>x. e\<cdot>x) ` (\<lambda>x. d\<cdot>x) ` range (\<lambda>x. p\<cdot>x))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
241  | 
by (simp add: finite_image)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
242  | 
hence "finite (range (\<lambda>x. (e oo d oo p)\<cdot>x))"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
243  | 
by (simp add: image_image)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
244  | 
  thus "finite {x. (e oo d oo p)\<cdot>x = x}"
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
245  | 
by (rule finite_range_imp_finite_fixes)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
246  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
248  | 
lemma deflation_p_d_e:  | 
| 28611 | 249  | 
assumes "deflation d"  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
250  | 
assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
251  | 
shows "deflation (p oo d oo e)"  | 
| 28611 | 252  | 
proof -  | 
| 29237 | 253  | 
interpret d: deflation d by fact  | 
| 28613 | 254  | 
  {
 | 
255  | 
fix x  | 
|
256  | 
have "d\<cdot>(e\<cdot>x) \<sqsubseteq> e\<cdot>x"  | 
|
257  | 
by (rule d.less)  | 
|
258  | 
hence "p\<cdot>(d\<cdot>(e\<cdot>x)) \<sqsubseteq> p\<cdot>(e\<cdot>x)"  | 
|
259  | 
by (rule monofun_cfun_arg)  | 
|
260  | 
hence "(p oo d oo e)\<cdot>x \<sqsubseteq> x"  | 
|
261  | 
by simp  | 
|
262  | 
}  | 
|
263  | 
note p_d_e_less = this  | 
|
| 28611 | 264  | 
show ?thesis  | 
| 28613 | 265  | 
proof  | 
266  | 
fix x  | 
|
267  | 
show "(p oo d oo e)\<cdot>x \<sqsubseteq> x"  | 
|
268  | 
by (rule p_d_e_less)  | 
|
269  | 
next  | 
|
270  | 
fix x  | 
|
271  | 
show "(p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x) = (p oo d oo e)\<cdot>x"  | 
|
272  | 
proof (rule antisym_less)  | 
|
273  | 
show "(p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x) \<sqsubseteq> (p oo d oo e)\<cdot>x"  | 
|
274  | 
by (rule p_d_e_less)  | 
|
275  | 
have "p\<cdot>(d\<cdot>(d\<cdot>(d\<cdot>(e\<cdot>x)))) \<sqsubseteq> p\<cdot>(d\<cdot>(e\<cdot>(p\<cdot>(d\<cdot>(e\<cdot>x)))))"  | 
|
276  | 
by (intro monofun_cfun_arg d)  | 
|
277  | 
hence "p\<cdot>(d\<cdot>(e\<cdot>x)) \<sqsubseteq> p\<cdot>(d\<cdot>(e\<cdot>(p\<cdot>(d\<cdot>(e\<cdot>x)))))"  | 
|
278  | 
by (simp only: d.idem)  | 
|
279  | 
thus "(p oo d oo e)\<cdot>x \<sqsubseteq> (p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x)"  | 
|
280  | 
by simp  | 
|
281  | 
qed  | 
|
282  | 
qed  | 
|
| 28611 | 283  | 
qed  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
284  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
285  | 
lemma finite_deflation_p_d_e:  | 
| 28611 | 286  | 
assumes "finite_deflation d"  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
287  | 
assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
288  | 
shows "finite_deflation (p oo d oo e)"  | 
| 28611 | 289  | 
proof -  | 
| 29237 | 290  | 
interpret d: finite_deflation d by fact  | 
| 28611 | 291  | 
show ?thesis  | 
| 28613 | 292  | 
proof (intro_locales)  | 
293  | 
have "deflation d" ..  | 
|
294  | 
thus "deflation (p oo d oo e)"  | 
|
295  | 
using d by (rule deflation_p_d_e)  | 
|
296  | 
next  | 
|
297  | 
show "finite_deflation_axioms (p oo d oo e)"  | 
|
298  | 
proof  | 
|
299  | 
have "finite ((\<lambda>x. d\<cdot>x) ` range (\<lambda>x. e\<cdot>x))"  | 
|
300  | 
by (rule d.finite_image)  | 
|
301  | 
hence "finite ((\<lambda>x. p\<cdot>x) ` (\<lambda>x. d\<cdot>x) ` range (\<lambda>x. e\<cdot>x))"  | 
|
302  | 
by (rule finite_imageI)  | 
|
303  | 
hence "finite (range (\<lambda>x. (p oo d oo e)\<cdot>x))"  | 
|
304  | 
by (simp add: image_image)  | 
|
305  | 
      thus "finite {x. (p oo d oo e)\<cdot>x = x}"
 | 
|
306  | 
by (rule finite_range_imp_finite_fixes)  | 
|
307  | 
qed  | 
|
308  | 
qed  | 
|
| 28611 | 309  | 
qed  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
310  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
311  | 
end  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
312  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
313  | 
subsection {* Uniqueness of ep-pairs *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
314  | 
|
| 28613 | 315  | 
lemma ep_pair_unique_e_lemma:  | 
316  | 
assumes "ep_pair e1 p" and "ep_pair e2 p"  | 
|
317  | 
shows "e1 \<sqsubseteq> e2"  | 
|
318  | 
proof (rule less_cfun_ext)  | 
|
| 29237 | 319  | 
interpret e1: ep_pair e1 p by fact  | 
320  | 
interpret e2: ep_pair e2 p by fact  | 
|
| 28613 | 321  | 
fix x  | 
322  | 
have "e1\<cdot>(p\<cdot>(e2\<cdot>x)) \<sqsubseteq> e2\<cdot>x"  | 
|
323  | 
by (rule e1.e_p_less)  | 
|
324  | 
thus "e1\<cdot>x \<sqsubseteq> e2\<cdot>x"  | 
|
325  | 
by (simp only: e2.e_inverse)  | 
|
326  | 
qed  | 
|
327  | 
||
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
328  | 
lemma ep_pair_unique_e:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
329  | 
"\<lbrakk>ep_pair e1 p; ep_pair e2 p\<rbrakk> \<Longrightarrow> e1 = e2"  | 
| 28613 | 330  | 
by (fast intro: antisym_less elim: ep_pair_unique_e_lemma)  | 
331  | 
||
332  | 
lemma ep_pair_unique_p_lemma:  | 
|
333  | 
assumes "ep_pair e p1" and "ep_pair e p2"  | 
|
334  | 
shows "p1 \<sqsubseteq> p2"  | 
|
335  | 
proof (rule less_cfun_ext)  | 
|
| 29237 | 336  | 
interpret p1: ep_pair e p1 by fact  | 
337  | 
interpret p2: ep_pair e p2 by fact  | 
|
| 28613 | 338  | 
fix x  | 
339  | 
have "e\<cdot>(p1\<cdot>x) \<sqsubseteq> x"  | 
|
340  | 
by (rule p1.e_p_less)  | 
|
341  | 
hence "p2\<cdot>(e\<cdot>(p1\<cdot>x)) \<sqsubseteq> p2\<cdot>x"  | 
|
342  | 
by (rule monofun_cfun_arg)  | 
|
343  | 
thus "p1\<cdot>x \<sqsubseteq> p2\<cdot>x"  | 
|
344  | 
by (simp only: p2.e_inverse)  | 
|
345  | 
qed  | 
|
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
347  | 
lemma ep_pair_unique_p:  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
348  | 
"\<lbrakk>ep_pair e p1; ep_pair e p2\<rbrakk> \<Longrightarrow> p1 = p2"  | 
| 28613 | 349  | 
by (fast intro: antisym_less elim: ep_pair_unique_p_lemma)  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
350  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
351  | 
subsection {* Composing ep-pairs *}
 | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
352  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
353  | 
lemma ep_pair_ID_ID: "ep_pair ID ID"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
354  | 
by default simp_all  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
355  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
356  | 
lemma ep_pair_comp:  | 
| 28613 | 357  | 
assumes "ep_pair e1 p1" and "ep_pair e2 p2"  | 
358  | 
shows "ep_pair (e2 oo e1) (p1 oo p2)"  | 
|
359  | 
proof  | 
|
| 29237 | 360  | 
interpret ep1: ep_pair e1 p1 by fact  | 
361  | 
interpret ep2: ep_pair e2 p2 by fact  | 
|
| 28613 | 362  | 
fix x y  | 
363  | 
show "(p1 oo p2)\<cdot>((e2 oo e1)\<cdot>x) = x"  | 
|
364  | 
by simp  | 
|
365  | 
have "e1\<cdot>(p1\<cdot>(p2\<cdot>y)) \<sqsubseteq> p2\<cdot>y"  | 
|
366  | 
by (rule ep1.e_p_less)  | 
|
367  | 
hence "e2\<cdot>(e1\<cdot>(p1\<cdot>(p2\<cdot>y))) \<sqsubseteq> e2\<cdot>(p2\<cdot>y)"  | 
|
368  | 
by (rule monofun_cfun_arg)  | 
|
369  | 
also have "e2\<cdot>(p2\<cdot>y) \<sqsubseteq> y"  | 
|
370  | 
by (rule ep2.e_p_less)  | 
|
371  | 
finally show "(e2 oo e1)\<cdot>((p1 oo p2)\<cdot>y) \<sqsubseteq> y"  | 
|
372  | 
by simp  | 
|
373  | 
qed  | 
|
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
374  | 
|
| 27681 | 375  | 
locale pcpo_ep_pair = ep_pair +  | 
| 
27401
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
376  | 
constrains e :: "'a::pcpo \<rightarrow> 'b::pcpo"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
377  | 
constrains p :: "'b::pcpo \<rightarrow> 'a::pcpo"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
378  | 
begin  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
379  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
380  | 
lemma e_strict [simp]: "e\<cdot>\<bottom> = \<bottom>"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
381  | 
proof -  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
382  | 
have "\<bottom> \<sqsubseteq> p\<cdot>\<bottom>" by (rule minimal)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
383  | 
hence "e\<cdot>\<bottom> \<sqsubseteq> e\<cdot>(p\<cdot>\<bottom>)" by (rule monofun_cfun_arg)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
384  | 
also have "e\<cdot>(p\<cdot>\<bottom>) \<sqsubseteq> \<bottom>" by (rule e_p_less)  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
385  | 
finally show "e\<cdot>\<bottom> = \<bottom>" by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
386  | 
qed  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
387  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
388  | 
lemma e_defined_iff [simp]: "e\<cdot>x = \<bottom> \<longleftrightarrow> x = \<bottom>"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
389  | 
by (rule e_eq_iff [where y="\<bottom>", unfolded e_strict])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
390  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
391  | 
lemma e_defined: "x \<noteq> \<bottom> \<Longrightarrow> e\<cdot>x \<noteq> \<bottom>"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
392  | 
by simp  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
393  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
394  | 
lemma p_strict [simp]: "p\<cdot>\<bottom> = \<bottom>"  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
395  | 
by (rule e_inverse [where x="\<bottom>", unfolded e_strict])  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
396  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
397  | 
lemmas stricts = e_strict p_strict  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
399  | 
end  | 
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
400  | 
|
| 
 
4edc81f93e35
New theory of deflations and embedding-projection pairs
 
huffman 
parents:  
diff
changeset
 | 
401  | 
end  |