author | wenzelm |
Sat, 30 Dec 2006 16:08:06 +0100 | |
changeset 21966 | edab0ecfbd7c |
parent 21547 | 9c9fdf4c2949 |
child 22276 | 96a4db55a0b3 |
permissions | -rw-r--r-- |
17006 | 1 |
(* Title: HOL/FixedPoint.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
4 |
Author: Stefan Berghofer, TU Muenchen |
17006 | 5 |
Copyright 1992 University of Cambridge |
6 |
*) |
|
7 |
||
8 |
header{* Fixed Points and the Knaster-Tarski Theorem*} |
|
9 |
||
10 |
theory FixedPoint |
|
21326 | 11 |
imports Product_Type LOrder |
17006 | 12 |
begin |
13 |
||
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
14 |
subsection {* Complete lattices *} |
21312 | 15 |
(*FIXME Meet \<rightarrow> Inf *) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
16 |
consts |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
17 |
Meet :: "'a::order set \<Rightarrow> 'a" |
21312 | 18 |
Sup :: "'a::order set \<Rightarrow> 'a" |
19 |
||
20 |
defs Sup_def: "Sup A == Meet {b. \<forall>a \<in> A. a <= b}" |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
21 |
|
21312 | 22 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21326
diff
changeset
|
23 |
SUP :: "('a \<Rightarrow> 'b::order) \<Rightarrow> 'b" (binder "SUP " 10) where |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21326
diff
changeset
|
24 |
"SUP x. f x == Sup (f ` UNIV)" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21326
diff
changeset
|
25 |
|
21312 | 26 |
(* |
27 |
abbreviation |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21326
diff
changeset
|
28 |
bot :: "'a::order" where |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21326
diff
changeset
|
29 |
"bot == Sup {}" |
21312 | 30 |
*) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
31 |
axclass comp_lat < order |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
32 |
Meet_lower: "x \<in> A \<Longrightarrow> Meet A <= x" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
33 |
Meet_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z <= x) \<Longrightarrow> z <= Meet A" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
34 |
|
21312 | 35 |
theorem Sup_upper: "(x::'a::comp_lat) \<in> A \<Longrightarrow> x <= Sup A" |
36 |
by (auto simp: Sup_def intro: Meet_greatest) |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
37 |
|
21312 | 38 |
theorem Sup_least: "(\<And>x::'a::comp_lat. x \<in> A \<Longrightarrow> x <= z) \<Longrightarrow> Sup A <= z" |
39 |
by (auto simp: Sup_def intro: Meet_lower) |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
40 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
41 |
text {* A complete lattice is a lattice *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
42 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
43 |
lemma is_meet_Meet: "is_meet (\<lambda>(x::'a::comp_lat) y. Meet {x, y})" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
44 |
by (auto simp: is_meet_def intro: Meet_lower Meet_greatest) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
45 |
|
21312 | 46 |
lemma is_join_Sup: "is_join (\<lambda>(x::'a::comp_lat) y. Sup {x, y})" |
47 |
by (auto simp: is_join_def intro: Sup_upper Sup_least) |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
48 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
49 |
instance comp_lat < lorder |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
50 |
proof |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
51 |
from is_meet_Meet show "\<exists>m::'a\<Rightarrow>'a\<Rightarrow>'a. is_meet m" by iprover |
21312 | 52 |
from is_join_Sup show "\<exists>j::'a\<Rightarrow>'a\<Rightarrow>'a. is_join j" by iprover |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
53 |
qed |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
54 |
|
21312 | 55 |
subsubsection {* Properties *} |
56 |
||
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
57 |
lemma mono_join: "mono f \<Longrightarrow> join (f A) (f B) <= f (join A B)" |
21312 | 58 |
by (auto simp add: mono_def) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
59 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
60 |
lemma mono_meet: "mono f \<Longrightarrow> f (meet A B) <= meet (f A) (f B)" |
21312 | 61 |
by (auto simp add: mono_def) |
62 |
||
63 |
lemma Sup_insert[simp]: "Sup (insert (a::'a::comp_lat) A) = join a (Sup A)" |
|
64 |
apply(simp add:Sup_def) |
|
65 |
apply(rule order_antisym) |
|
66 |
apply(rule Meet_lower) |
|
67 |
apply(clarsimp) |
|
68 |
apply(rule le_joinI2) |
|
69 |
apply(rule Meet_greatest) |
|
70 |
apply blast |
|
71 |
apply simp |
|
72 |
apply rule |
|
73 |
apply(rule Meet_greatest)apply blast |
|
74 |
apply(rule Meet_greatest) |
|
75 |
apply(rule Meet_lower) |
|
76 |
apply blast |
|
77 |
done |
|
78 |
||
79 |
lemma bot_least[simp]: "Sup{} \<le> (x::'a::comp_lat)" |
|
80 |
apply(simp add: Sup_def) |
|
81 |
apply(rule Meet_lower) |
|
82 |
apply blast |
|
83 |
done |
|
84 |
(* |
|
85 |
lemma Meet_singleton[simp]: "Meet{a} = (a::'a::comp_lat)" |
|
86 |
apply(rule order_antisym) |
|
87 |
apply(simp add: Meet_lower) |
|
88 |
apply(rule Meet_greatest) |
|
89 |
apply(simp) |
|
90 |
done |
|
91 |
*) |
|
92 |
lemma le_SupI: "(l::'a::comp_lat) : M \<Longrightarrow> l \<le> Sup M" |
|
93 |
apply(simp add:Sup_def) |
|
94 |
apply(rule Meet_greatest) |
|
95 |
apply(simp) |
|
96 |
done |
|
97 |
||
98 |
lemma le_SUPI: "(l::'a::comp_lat) = M i \<Longrightarrow> l \<le> (SUP i. M i)" |
|
99 |
apply(simp add:SUP_def) |
|
100 |
apply(blast intro:le_SupI) |
|
101 |
done |
|
102 |
||
103 |
lemma Sup_leI: "(!!x. x:M \<Longrightarrow> x \<le> u) \<Longrightarrow> Sup M \<le> (u::'a::comp_lat)" |
|
104 |
apply(simp add:Sup_def) |
|
105 |
apply(rule Meet_lower) |
|
106 |
apply(blast) |
|
107 |
done |
|
108 |
||
109 |
||
110 |
lemma SUP_leI: "(!!i. M i \<le> u) \<Longrightarrow> (SUP i. M i) \<le> (u::'a::comp_lat)" |
|
111 |
apply(simp add:SUP_def) |
|
112 |
apply(blast intro!:Sup_leI) |
|
113 |
done |
|
114 |
||
115 |
lemma SUP_const[simp]: "(SUP i. M) = (M::'a::comp_lat)" |
|
21316 | 116 |
by(simp add:SUP_def image_constant_conv join_absorp1) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
117 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
118 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
119 |
subsection {* Some instances of the type class of complete lattices *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
120 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
121 |
subsubsection {* Booleans *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
122 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
123 |
defs Meet_bool_def: "Meet A == ALL x:A. x" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
124 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
125 |
instance bool :: comp_lat |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
126 |
apply intro_classes |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
127 |
apply (unfold Meet_bool_def) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
128 |
apply (iprover intro!: le_boolI elim: ballE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
129 |
apply (iprover intro!: ballI le_boolI elim: ballE le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
130 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
131 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
132 |
theorem meet_bool_eq: "meet P Q = (P \<and> Q)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
133 |
apply (rule order_antisym) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
134 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
135 |
apply (rule conjI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
136 |
apply (rule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
137 |
apply (rule meet_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
138 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
139 |
apply (rule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
140 |
apply (rule meet_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
141 |
apply assumption+ |
21312 | 142 |
apply (rule le_meetI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
143 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
144 |
apply (erule conjunct1) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
145 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
146 |
apply (erule conjunct2) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
147 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
148 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
149 |
theorem join_bool_eq: "join P Q = (P \<or> Q)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
150 |
apply (rule order_antisym) |
21312 | 151 |
apply (rule join_leI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
152 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
153 |
apply (erule disjI1) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
154 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
155 |
apply (erule disjI2) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
156 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
157 |
apply (erule disjE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
158 |
apply (rule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
159 |
apply (rule join_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
160 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
161 |
apply (rule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
162 |
apply (rule join_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
163 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
164 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
165 |
|
21312 | 166 |
theorem Sup_bool_eq: "Sup A = (EX x:A. x)" |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
167 |
apply (rule order_antisym) |
21312 | 168 |
apply (rule Sup_least) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
169 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
170 |
apply (erule bexI, assumption) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
171 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
172 |
apply (erule bexE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
173 |
apply (rule le_boolE) |
21312 | 174 |
apply (rule Sup_upper) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
175 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
176 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
177 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
178 |
subsubsection {* Functions *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
179 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
180 |
text {* |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
181 |
Handy introduction and elimination rules for @{text "\<le>"} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
182 |
on unary and binary predicates |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
183 |
*} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
184 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
185 |
lemma predicate1I [intro]: |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
186 |
assumes PQ: "\<And>x. P x \<Longrightarrow> Q x" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
187 |
shows "P \<le> Q" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
188 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
189 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
190 |
apply (rule PQ) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
191 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
192 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
193 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
194 |
lemma predicate1D [elim]: "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
195 |
apply (erule le_funE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
196 |
apply (erule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
197 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
198 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
199 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
200 |
lemma predicate2I [intro]: |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
201 |
assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
202 |
shows "P \<le> Q" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
203 |
apply (rule le_funI)+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
204 |
apply (rule le_boolI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
205 |
apply (rule PQ) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
206 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
207 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
208 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
209 |
lemma predicate2D [elim]: "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
210 |
apply (erule le_funE)+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
211 |
apply (erule le_boolE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
212 |
apply assumption+ |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
213 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
214 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
215 |
defs Meet_fun_def: "Meet A == (\<lambda>x. Meet {y. EX f:A. y = f x})" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
216 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
217 |
instance "fun" :: (type, comp_lat) comp_lat |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
218 |
apply intro_classes |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
219 |
apply (unfold Meet_fun_def) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
220 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
221 |
apply (rule Meet_lower) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
222 |
apply (rule CollectI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
223 |
apply (rule bexI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
224 |
apply (rule refl) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
225 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
226 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
227 |
apply (rule Meet_greatest) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
228 |
apply (erule CollectE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
229 |
apply (erule bexE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
230 |
apply (iprover elim: le_funE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
231 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
232 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
233 |
theorem meet_fun_eq: "meet f g = (\<lambda>x. meet (f x) (g x))" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
234 |
apply (rule order_antisym) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
235 |
apply (rule le_funI) |
21312 | 236 |
apply (rule le_meetI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
237 |
apply (rule le_funD [OF meet_left_le]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
238 |
apply (rule le_funD [OF meet_right_le]) |
21312 | 239 |
apply (rule le_meetI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
240 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
241 |
apply (rule meet_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
242 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
243 |
apply (rule meet_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
244 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
245 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
246 |
theorem join_fun_eq: "join f g = (\<lambda>x. join (f x) (g x))" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
247 |
apply (rule order_antisym) |
21312 | 248 |
apply (rule join_leI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
249 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
250 |
apply (rule join_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
251 |
apply (rule le_funI) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
252 |
apply (rule join_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
253 |
apply (rule le_funI) |
21312 | 254 |
apply (rule join_leI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
255 |
apply (rule le_funD [OF join_left_le]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
256 |
apply (rule le_funD [OF join_right_le]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
257 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
258 |
|
21312 | 259 |
theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y::'a::comp_lat. EX f:A. y = f x})" |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
260 |
apply (rule order_antisym) |
21312 | 261 |
apply (rule Sup_least) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
262 |
apply (rule le_funI) |
21312 | 263 |
apply (rule Sup_upper) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
264 |
apply fast |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
265 |
apply (rule le_funI) |
21312 | 266 |
apply (rule Sup_least) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
267 |
apply (erule CollectE) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
268 |
apply (erule bexE) |
21312 | 269 |
apply (drule le_funD [OF Sup_upper]) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
270 |
apply simp |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
271 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
272 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
273 |
subsubsection {* Sets *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
274 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
275 |
defs Meet_set_def: "Meet S == \<Inter>S" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
276 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
277 |
instance set :: (type) comp_lat |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
278 |
by intro_classes (auto simp add: Meet_set_def) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
279 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
280 |
theorem meet_set_eq: "meet A B = A \<inter> B" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
281 |
apply (rule subset_antisym) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
282 |
apply (rule Int_greatest) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
283 |
apply (rule meet_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
284 |
apply (rule meet_right_le) |
21312 | 285 |
apply (rule le_meetI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
286 |
apply (rule Int_lower1) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
287 |
apply (rule Int_lower2) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
288 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
289 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
290 |
theorem join_set_eq: "join A B = A \<union> B" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
291 |
apply (rule subset_antisym) |
21312 | 292 |
apply (rule join_leI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
293 |
apply (rule Un_upper1) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
294 |
apply (rule Un_upper2) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
295 |
apply (rule Un_least) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
296 |
apply (rule join_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
297 |
apply (rule join_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
298 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
299 |
|
21312 | 300 |
theorem Sup_set_eq: "Sup S = \<Union>S" |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
301 |
apply (rule subset_antisym) |
21312 | 302 |
apply (rule Sup_least) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
303 |
apply (erule Union_upper) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
304 |
apply (rule Union_least) |
21312 | 305 |
apply (erule Sup_upper) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
306 |
done |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
307 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
308 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
309 |
subsection {* Least and greatest fixed points *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
310 |
|
17006 | 311 |
constdefs |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
312 |
lfp :: "(('a::comp_lat) => 'a) => 'a" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
313 |
"lfp f == Meet {u. f u <= u}" --{*least fixed point*} |
17006 | 314 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
315 |
gfp :: "(('a::comp_lat) => 'a) => 'a" |
21312 | 316 |
"gfp f == Sup {u. u <= f u}" --{*greatest fixed point*} |
17006 | 317 |
|
318 |
||
319 |
subsection{*Proof of Knaster-Tarski Theorem using @{term lfp}*} |
|
320 |
||
321 |
||
322 |
text{*@{term "lfp f"} is the least upper bound of |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
323 |
the set @{term "{u. f(u) \<le> u}"} *} |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
324 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
325 |
lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
326 |
by (auto simp add: lfp_def intro: Meet_lower) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
327 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
328 |
lemma lfp_greatest: "(!!u. f u \<le> u ==> A \<le> u) ==> A \<le> lfp f" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
329 |
by (auto simp add: lfp_def intro: Meet_greatest) |
17006 | 330 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
331 |
lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
332 |
by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) |
17006 | 333 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
334 |
lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
335 |
by (iprover intro: lfp_lemma2 monoD lfp_lowerbound) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
336 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
337 |
lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
338 |
by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) |
17006 | 339 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
340 |
subsection{*General induction rules for least fixed points*} |
17006 | 341 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
342 |
theorem lfp_induct: |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
343 |
assumes mono: "mono f" and ind: "f (meet (lfp f) P) <= P" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
344 |
shows "lfp f <= P" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
345 |
proof - |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
346 |
have "meet (lfp f) P <= lfp f" by (rule meet_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
347 |
with mono have "f (meet (lfp f) P) <= f (lfp f)" .. |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
348 |
also from mono have "f (lfp f) = lfp f" by (rule lfp_unfold [symmetric]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
349 |
finally have "f (meet (lfp f) P) <= lfp f" . |
21312 | 350 |
from this and ind have "f (meet (lfp f) P) <= meet (lfp f) P" by (rule le_meetI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
351 |
hence "lfp f <= meet (lfp f) P" by (rule lfp_lowerbound) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
352 |
also have "meet (lfp f) P <= P" by (rule meet_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
353 |
finally show ?thesis . |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
354 |
qed |
17006 | 355 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
356 |
lemma lfp_induct_set: |
17006 | 357 |
assumes lfp: "a: lfp(f)" |
358 |
and mono: "mono(f)" |
|
359 |
and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)" |
|
360 |
shows "P(a)" |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
361 |
by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
362 |
(auto simp: meet_set_eq intro: indhyp) |
17006 | 363 |
|
364 |
||
365 |
text{*Version of induction for binary relations*} |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
366 |
lemmas lfp_induct2 = lfp_induct_set [of "(a,b)", split_format (complete)] |
17006 | 367 |
|
368 |
||
369 |
lemma lfp_ordinal_induct: |
|
370 |
assumes mono: "mono f" |
|
371 |
shows "[| !!S. P S ==> P(f S); !!M. !S:M. P S ==> P(Union M) |] |
|
372 |
==> P(lfp f)" |
|
373 |
apply(subgoal_tac "lfp f = Union{S. S \<subseteq> lfp f & P S}") |
|
374 |
apply (erule ssubst, simp) |
|
375 |
apply(subgoal_tac "Union{S. S \<subseteq> lfp f & P S} \<subseteq> lfp f") |
|
376 |
prefer 2 apply blast |
|
377 |
apply(rule equalityI) |
|
378 |
prefer 2 apply assumption |
|
379 |
apply(drule mono [THEN monoD]) |
|
380 |
apply (cut_tac mono [THEN lfp_unfold], simp) |
|
381 |
apply (rule lfp_lowerbound, auto) |
|
382 |
done |
|
383 |
||
384 |
||
385 |
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, |
|
386 |
to control unfolding*} |
|
387 |
||
388 |
lemma def_lfp_unfold: "[| h==lfp(f); mono(f) |] ==> h = f(h)" |
|
389 |
by (auto intro!: lfp_unfold) |
|
390 |
||
391 |
lemma def_lfp_induct: |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
392 |
"[| A == lfp(f); mono(f); |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
393 |
f (meet A P) \<le> P |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
394 |
|] ==> A \<le> P" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
395 |
by (blast intro: lfp_induct) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
396 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
397 |
lemma def_lfp_induct_set: |
17006 | 398 |
"[| A == lfp(f); mono(f); a:A; |
399 |
!!x. [| x: f(A Int {x. P(x)}) |] ==> P(x) |
|
400 |
|] ==> P(a)" |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
401 |
by (blast intro: lfp_induct_set) |
17006 | 402 |
|
403 |
(*Monotonicity of lfp!*) |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
404 |
lemma lfp_mono: "(!!Z. f Z \<le> g Z) ==> lfp f \<le> lfp g" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
405 |
by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans) |
17006 | 406 |
|
407 |
||
408 |
subsection{*Proof of Knaster-Tarski Theorem using @{term gfp}*} |
|
409 |
||
410 |
||
411 |
text{*@{term "gfp f"} is the greatest lower bound of |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
412 |
the set @{term "{u. u \<le> f(u)}"} *} |
17006 | 413 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
414 |
lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f" |
21312 | 415 |
by (auto simp add: gfp_def intro: Sup_upper) |
17006 | 416 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
417 |
lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X" |
21312 | 418 |
by (auto simp add: gfp_def intro: Sup_least) |
17006 | 419 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
420 |
lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
421 |
by (iprover intro: gfp_least order_trans monoD gfp_upperbound) |
17006 | 422 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
423 |
lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
424 |
by (iprover intro: gfp_lemma2 monoD gfp_upperbound) |
17006 | 425 |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
426 |
lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
427 |
by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) |
17006 | 428 |
|
429 |
subsection{*Coinduction rules for greatest fixed points*} |
|
430 |
||
431 |
text{*weak version*} |
|
432 |
lemma weak_coinduct: "[| a: X; X \<subseteq> f(X) |] ==> a : gfp(f)" |
|
433 |
by (rule gfp_upperbound [THEN subsetD], auto) |
|
434 |
||
435 |
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f" |
|
436 |
apply (erule gfp_upperbound [THEN subsetD]) |
|
437 |
apply (erule imageI) |
|
438 |
done |
|
439 |
||
440 |
lemma coinduct_lemma: |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
441 |
"[| X \<le> f (join X (gfp f)); mono f |] ==> join X (gfp f) \<le> f (join X (gfp f))" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
442 |
apply (frule gfp_lemma2) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
443 |
apply (drule mono_join) |
21312 | 444 |
apply (rule join_leI) |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
445 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
446 |
apply (rule order_trans) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
447 |
apply (rule order_trans) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
448 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
449 |
apply (rule join_right_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
450 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
451 |
done |
17006 | 452 |
|
453 |
text{*strong version, thanks to Coen and Frost*} |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
454 |
lemma coinduct_set: "[| mono(f); a: X; X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
455 |
by (blast intro: weak_coinduct [OF _ coinduct_lemma, simplified join_set_eq]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
456 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
457 |
lemma coinduct: "[| mono(f); X \<le> f (join X (gfp f)) |] ==> X \<le> gfp(f)" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
458 |
apply (rule order_trans) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
459 |
apply (rule join_left_le) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
460 |
apply (erule gfp_upperbound [OF coinduct_lemma]) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
461 |
apply assumption |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
462 |
done |
17006 | 463 |
|
464 |
lemma gfp_fun_UnI2: "[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))" |
|
465 |
by (blast dest: gfp_lemma2 mono_Un) |
|
466 |
||
467 |
subsection{*Even Stronger Coinduction Rule, by Martin Coen*} |
|
468 |
||
469 |
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both |
|
470 |
@{term lfp} and @{term gfp}*} |
|
471 |
||
472 |
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)" |
|
17589 | 473 |
by (iprover intro: subset_refl monoI Un_mono monoD) |
17006 | 474 |
|
475 |
lemma coinduct3_lemma: |
|
476 |
"[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))); mono(f) |] |
|
477 |
==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))" |
|
478 |
apply (rule subset_trans) |
|
479 |
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) |
|
480 |
apply (rule Un_least [THEN Un_least]) |
|
481 |
apply (rule subset_refl, assumption) |
|
482 |
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) |
|
483 |
apply (rule monoD, assumption) |
|
484 |
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) |
|
485 |
done |
|
486 |
||
487 |
lemma coinduct3: |
|
488 |
"[| mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)" |
|
489 |
apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) |
|
490 |
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto) |
|
491 |
done |
|
492 |
||
493 |
||
494 |
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, |
|
495 |
to control unfolding*} |
|
496 |
||
497 |
lemma def_gfp_unfold: "[| A==gfp(f); mono(f) |] ==> A = f(A)" |
|
498 |
by (auto intro!: gfp_unfold) |
|
499 |
||
500 |
lemma def_coinduct: |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
501 |
"[| A==gfp(f); mono(f); X \<le> f(join X A) |] ==> X \<le> A" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
502 |
by (iprover intro!: coinduct) |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
503 |
|
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
504 |
lemma def_coinduct_set: |
17006 | 505 |
"[| A==gfp(f); mono(f); a:X; X \<subseteq> f(X Un A) |] ==> a: A" |
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
506 |
by (auto intro!: coinduct_set) |
17006 | 507 |
|
508 |
(*The version used in the induction/coinduction package*) |
|
509 |
lemma def_Collect_coinduct: |
|
510 |
"[| A == gfp(%w. Collect(P(w))); mono(%w. Collect(P(w))); |
|
511 |
a: X; !!z. z: X ==> P (X Un A) z |] ==> |
|
512 |
a : A" |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
513 |
apply (erule def_coinduct_set, auto) |
17006 | 514 |
done |
515 |
||
516 |
lemma def_coinduct3: |
|
517 |
"[| A==gfp(f); mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A" |
|
518 |
by (auto intro!: coinduct3) |
|
519 |
||
520 |
text{*Monotonicity of @{term gfp}!*} |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
521 |
lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g" |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
522 |
by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans) |
17006 | 523 |
|
524 |
||
21312 | 525 |
|
17006 | 526 |
ML |
527 |
{* |
|
528 |
val lfp_def = thm "lfp_def"; |
|
529 |
val lfp_lowerbound = thm "lfp_lowerbound"; |
|
530 |
val lfp_greatest = thm "lfp_greatest"; |
|
531 |
val lfp_unfold = thm "lfp_unfold"; |
|
532 |
val lfp_induct = thm "lfp_induct"; |
|
533 |
val lfp_induct2 = thm "lfp_induct2"; |
|
534 |
val lfp_ordinal_induct = thm "lfp_ordinal_induct"; |
|
535 |
val def_lfp_unfold = thm "def_lfp_unfold"; |
|
536 |
val def_lfp_induct = thm "def_lfp_induct"; |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
537 |
val def_lfp_induct_set = thm "def_lfp_induct_set"; |
17006 | 538 |
val lfp_mono = thm "lfp_mono"; |
539 |
val gfp_def = thm "gfp_def"; |
|
540 |
val gfp_upperbound = thm "gfp_upperbound"; |
|
541 |
val gfp_least = thm "gfp_least"; |
|
542 |
val gfp_unfold = thm "gfp_unfold"; |
|
543 |
val weak_coinduct = thm "weak_coinduct"; |
|
544 |
val weak_coinduct_image = thm "weak_coinduct_image"; |
|
545 |
val coinduct = thm "coinduct"; |
|
546 |
val gfp_fun_UnI2 = thm "gfp_fun_UnI2"; |
|
547 |
val coinduct3 = thm "coinduct3"; |
|
548 |
val def_gfp_unfold = thm "def_gfp_unfold"; |
|
549 |
val def_coinduct = thm "def_coinduct"; |
|
550 |
val def_Collect_coinduct = thm "def_Collect_coinduct"; |
|
551 |
val def_coinduct3 = thm "def_coinduct3"; |
|
552 |
val gfp_mono = thm "gfp_mono"; |
|
21017
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
553 |
val le_funI = thm "le_funI"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
554 |
val le_boolI = thm "le_boolI"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
555 |
val le_boolI' = thm "le_boolI'"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
556 |
val meet_fun_eq = thm "meet_fun_eq"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
557 |
val meet_bool_eq = thm "meet_bool_eq"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
558 |
val le_funE = thm "le_funE"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
559 |
val le_boolE = thm "le_boolE"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
560 |
val le_boolD = thm "le_boolD"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
561 |
val le_bool_def = thm "le_bool_def"; |
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents:
17589
diff
changeset
|
562 |
val le_fun_def = thm "le_fun_def"; |
17006 | 563 |
*} |
564 |
||
565 |
end |