| author | desharna | 
| Sun, 17 Mar 2024 19:30:34 +0100 | |
| changeset 79923 | 6fc9c4344df4 | 
| parent 75464 | 84e6f9b542e2 | 
| permissions | -rw-r--r-- | 
| 63612 | 1 | (* Title: HOL/Complete_Partial_Order.thy | 
| 2 | Author: Brian Huffman, Portland State University | |
| 3 | Author: Alexander Krauss, TU Muenchen | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 4 | *) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 5 | |
| 60758 | 6 | section \<open>Chain-complete partial orders and their fixpoints\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 7 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 8 | theory Complete_Partial_Order | 
| 63612 | 9 | imports Product_Type | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 10 | begin | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 11 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 12 | |
| 60758 | 13 | subsection \<open>Chains\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 14 | |
| 63612 | 15 | text \<open> | 
| 16 | A chain is a totally-ordered set. Chains are parameterized over | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 17 | the order for maximal flexibility, since type classes are not enough. | 
| 60758 | 18 | \<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 19 | |
| 63612 | 20 | definition chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 21 | where "chain ord S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. ord x y \<or> ord y x)" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 22 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 23 | lemma chainI: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 24 | assumes "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> ord x y \<or> ord y x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 25 | shows "chain ord S" | 
| 63612 | 26 | using assms unfolding chain_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 27 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 28 | lemma chainD: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 29 | assumes "chain ord S" and "x \<in> S" and "y \<in> S" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 30 | shows "ord x y \<or> ord y x" | 
| 63612 | 31 | using assms unfolding chain_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 32 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 33 | lemma chainE: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 34 | assumes "chain ord S" and "x \<in> S" and "y \<in> S" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 35 | obtains "ord x y" | "ord y x" | 
| 63612 | 36 | using assms unfolding chain_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 37 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 38 | lemma chain_empty: "chain ord {}"
 | 
| 63612 | 39 | by (simp add: chain_def) | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 40 | |
| 67399 | 41 | lemma chain_equality: "chain (=) A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. x = y)" | 
| 63612 | 42 | by (auto simp add: chain_def) | 
| 43 | ||
| 44 | lemma chain_subset: "chain ord A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> chain ord B" | |
| 45 | by (rule chainI) (blast dest: chainD) | |
| 60057 | 46 | |
| 63612 | 47 | lemma chain_imageI: | 
| 48 | assumes chain: "chain le_a Y" | |
| 49 | and mono: "\<And>x y. x \<in> Y \<Longrightarrow> y \<in> Y \<Longrightarrow> le_a x y \<Longrightarrow> le_b (f x) (f y)" | |
| 50 | shows "chain le_b (f ` Y)" | |
| 51 | by (blast intro: chainI dest: chainD[OF chain] mono) | |
| 60061 | 52 | |
| 53 | ||
| 60758 | 54 | subsection \<open>Chain-complete partial orders\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 55 | |
| 60758 | 56 | text \<open> | 
| 63612 | 57 | A \<open>ccpo\<close> has a least upper bound for any chain. In particular, the | 
| 58 | empty set is a chain, so every \<open>ccpo\<close> must have a bottom element. | |
| 60758 | 59 | \<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 60 | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 61 | class ccpo = order + Sup + | 
| 67399 | 62 | assumes ccpo_Sup_upper: "chain (\<le>) A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> Sup A" | 
| 63 | assumes ccpo_Sup_least: "chain (\<le>) A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup A \<le> z" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 64 | begin | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 65 | |
| 67399 | 66 | lemma chain_singleton: "Complete_Partial_Order.chain (\<le>) {x}"
 | 
| 63612 | 67 | by (rule chainI) simp | 
| 60061 | 68 | |
| 69 | lemma ccpo_Sup_singleton [simp]: "\<Squnion>{x} = x"
 | |
| 73411 | 70 | by (rule order.antisym) (auto intro: ccpo_Sup_least ccpo_Sup_upper simp add: chain_singleton) | 
| 63612 | 71 | |
| 60061 | 72 | |
| 60758 | 73 | subsection \<open>Transfinite iteration of a function\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 74 | |
| 63612 | 75 | context notes [[inductive_internals]] | 
| 76 | begin | |
| 61689 | 77 | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 78 | inductive_set iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set"
 | 
| 63612 | 79 | for f :: "'a \<Rightarrow> 'a" | 
| 80 | where | |
| 81 | step: "x \<in> iterates f \<Longrightarrow> f x \<in> iterates f" | |
| 67399 | 82 | | Sup: "chain (\<le>) M \<Longrightarrow> \<forall>x\<in>M. x \<in> iterates f \<Longrightarrow> Sup M \<in> iterates f" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 83 | |
| 61689 | 84 | end | 
| 85 | ||
| 67399 | 86 | lemma iterates_le_f: "x \<in> iterates f \<Longrightarrow> monotone (\<le>) (\<le>) f \<Longrightarrow> x \<le> f x" | 
| 63612 | 87 | by (induct x rule: iterates.induct) | 
| 88 | (force dest: monotoneD intro!: ccpo_Sup_upper ccpo_Sup_least)+ | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 89 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 90 | lemma chain_iterates: | 
| 67399 | 91 | assumes f: "monotone (\<le>) (\<le>) f" | 
| 92 | shows "chain (\<le>) (iterates f)" (is "chain _ ?C") | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 93 | proof (rule chainI) | 
| 63612 | 94 | fix x y | 
| 95 | assume "x \<in> ?C" "y \<in> ?C" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 96 | then show "x \<le> y \<or> y \<le> x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 97 | proof (induct x arbitrary: y rule: iterates.induct) | 
| 63612 | 98 | fix x y | 
| 99 | assume y: "y \<in> ?C" | |
| 100 | and IH: "\<And>z. z \<in> ?C \<Longrightarrow> x \<le> z \<or> z \<le> x" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 101 | from y show "f x \<le> y \<or> y \<le> f x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 102 | proof (induct y rule: iterates.induct) | 
| 63612 | 103 | case (step y) | 
| 104 | with IH f show ?case by (auto dest: monotoneD) | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 105 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 106 | case (Sup M) | 
| 67399 | 107 | then have chM: "chain (\<le>) M" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 108 | and IH': "\<And>z. z \<in> M \<Longrightarrow> f x \<le> z \<or> z \<le> f x" by auto | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 109 | show "f x \<le> Sup M \<or> Sup M \<le> f x" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 110 | proof (cases "\<exists>z\<in>M. f x \<le> z") | 
| 63612 | 111 | case True | 
| 112 | then have "f x \<le> Sup M" | |
| 73252 | 113 | by (blast intro: ccpo_Sup_upper[OF chM] order_trans) | 
| 63612 | 114 | then show ?thesis .. | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 115 | next | 
| 63612 | 116 | case False | 
| 117 | with IH' show ?thesis | |
| 118 | by (auto intro: ccpo_Sup_least[OF chM]) | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 119 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 120 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 121 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 122 | case (Sup M y) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 123 | show ?case | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 124 | proof (cases "\<exists>x\<in>M. y \<le> x") | 
| 63612 | 125 | case True | 
| 126 | then have "y \<le> Sup M" | |
| 73252 | 127 | by (blast intro: ccpo_Sup_upper[OF Sup(1)] order_trans) | 
| 63612 | 128 | then show ?thesis .. | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 129 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 130 | case False with Sup | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 131 | show ?thesis by (auto intro: ccpo_Sup_least) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 132 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 133 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 134 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 135 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 136 | lemma bot_in_iterates: "Sup {} \<in> iterates f"
 | 
| 63612 | 137 | by (auto intro: iterates.Sup simp add: chain_empty) | 
| 138 | ||
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 139 | |
| 60758 | 140 | subsection \<open>Fixpoint combinator\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 141 | |
| 63612 | 142 | definition fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
 | 
| 143 | where "fixp f = Sup (iterates f)" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 144 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 145 | lemma iterates_fixp: | 
| 67399 | 146 | assumes f: "monotone (\<le>) (\<le>) f" | 
| 63612 | 147 | shows "fixp f \<in> iterates f" | 
| 148 | unfolding fixp_def | |
| 149 | by (simp add: iterates.Sup chain_iterates f) | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 150 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 151 | lemma fixp_unfold: | 
| 67399 | 152 | assumes f: "monotone (\<le>) (\<le>) f" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 153 | shows "fixp f = f (fixp f)" | 
| 73411 | 154 | proof (rule order.antisym) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 155 | show "fixp f \<le> f (fixp f)" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 156 | by (intro iterates_le_f iterates_fixp f) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 157 | have "f (fixp f) \<le> Sup (iterates f)" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 158 | by (intro ccpo_Sup_upper chain_iterates f iterates.step iterates_fixp) | 
| 63612 | 159 | then show "f (fixp f) \<le> fixp f" | 
| 160 | by (simp only: fixp_def) | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 161 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 162 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 163 | lemma fixp_lowerbound: | 
| 67399 | 164 | assumes f: "monotone (\<le>) (\<le>) f" | 
| 63612 | 165 | and z: "f z \<le> z" | 
| 166 | shows "fixp f \<le> z" | |
| 167 | unfolding fixp_def | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 168 | proof (rule ccpo_Sup_least[OF chain_iterates[OF f]]) | 
| 63612 | 169 | fix x | 
| 170 | assume "x \<in> iterates f" | |
| 171 | then show "x \<le> z" | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 172 | proof (induct x rule: iterates.induct) | 
| 63612 | 173 | case (step x) | 
| 174 | from f \<open>x \<le> z\<close> have "f x \<le> f z" by (rule monotoneD) | |
| 175 | also note z | |
| 176 | finally show "f x \<le> z" . | |
| 177 | next | |
| 178 | case (Sup M) | |
| 179 | then show ?case | |
| 180 | by (auto intro: ccpo_Sup_least) | |
| 181 | qed | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 182 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 183 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 184 | end | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 185 | |
| 63612 | 186 | |
| 60758 | 187 | subsection \<open>Fixpoint induction\<close> | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 188 | |
| 60758 | 189 | setup \<open>Sign.map_naming (Name_Space.mandatory_path "ccpo")\<close> | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 190 | |
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 191 | definition admissible :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 63612 | 192 |   where "admissible lub ord P \<longleftrightarrow> (\<forall>A. chain ord A \<longrightarrow> A \<noteq> {} \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (lub A))"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 193 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 194 | lemma admissibleI: | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 195 |   assumes "\<And>A. chain ord A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> \<forall>x\<in>A. P x \<Longrightarrow> P (lub A)"
 | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 196 | shows "ccpo.admissible lub ord P" | 
| 63612 | 197 | using assms unfolding ccpo.admissible_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 198 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 199 | lemma admissibleD: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 200 | assumes "ccpo.admissible lub ord P" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 201 | assumes "chain ord A" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 202 |   assumes "A \<noteq> {}"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 203 | assumes "\<And>x. x \<in> A \<Longrightarrow> P x" | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 204 | shows "P (lub A)" | 
| 63612 | 205 | using assms by (auto simp: ccpo.admissible_def) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 206 | |
| 60758 | 207 | setup \<open>Sign.map_naming Name_Space.parent_path\<close> | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 208 | |
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 209 | lemma (in ccpo) fixp_induct: | 
| 67399 | 210 | assumes adm: "ccpo.admissible Sup (\<le>) P" | 
| 211 | assumes mono: "monotone (\<le>) (\<le>) f" | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 212 |   assumes bot: "P (Sup {})"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 213 | assumes step: "\<And>x. P x \<Longrightarrow> P (f x)" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 214 | shows "P (fixp f)" | 
| 63612 | 215 | unfolding fixp_def | 
| 216 | using adm chain_iterates[OF mono] | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 217 | proof (rule ccpo.admissibleD) | 
| 63612 | 218 |   show "iterates f \<noteq> {}"
 | 
| 219 | using bot_in_iterates by auto | |
| 220 | next | |
| 221 | fix x | |
| 222 | assume "x \<in> iterates f" | |
| 223 | then show "P x" | |
| 224 | proof (induct rule: iterates.induct) | |
| 225 | case prems: (step x) | |
| 226 | from this(2) show ?case by (rule step) | |
| 227 | next | |
| 228 | case (Sup M) | |
| 229 |     then show ?case by (cases "M = {}") (auto intro: step bot ccpo.admissibleD adm)
 | |
| 230 | qed | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 231 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 232 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 233 | lemma admissible_True: "ccpo.admissible lub ord (\<lambda>x. True)" | 
| 63612 | 234 | unfolding ccpo.admissible_def by simp | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 235 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 236 | (*lemma admissible_False: "\<not> ccpo.admissible lub ord (\<lambda>x. False)" | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 237 | unfolding ccpo.admissible_def chain_def by simp | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 238 | *) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 239 | lemma admissible_const: "ccpo.admissible lub ord (\<lambda>x. t)" | 
| 63612 | 240 | by (auto intro: ccpo.admissibleI) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 241 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 242 | lemma admissible_conj: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 243 | assumes "ccpo.admissible lub ord (\<lambda>x. P x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 244 | assumes "ccpo.admissible lub ord (\<lambda>x. Q x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 245 | shows "ccpo.admissible lub ord (\<lambda>x. P x \<and> Q x)" | 
| 63612 | 246 | using assms unfolding ccpo.admissible_def by simp | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 247 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 248 | lemma admissible_all: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 249 | assumes "\<And>y. ccpo.admissible lub ord (\<lambda>x. P x y)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 250 | shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y. P x y)" | 
| 63612 | 251 | using assms unfolding ccpo.admissible_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 252 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 253 | lemma admissible_ball: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 254 | assumes "\<And>y. y \<in> A \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x y)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 255 | shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y\<in>A. P x y)" | 
| 63612 | 256 | using assms unfolding ccpo.admissible_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 257 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 258 | lemma chain_compr: "chain ord A \<Longrightarrow> chain ord {x \<in> A. P x}"
 | 
| 63612 | 259 | unfolding chain_def by fast | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 260 | |
| 63612 | 261 | context ccpo | 
| 262 | begin | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 263 | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 264 | lemma admissible_disj: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 265 | fixes P Q :: "'a \<Rightarrow> bool" | 
| 67399 | 266 | assumes P: "ccpo.admissible Sup (\<le>) (\<lambda>x. P x)" | 
| 267 | assumes Q: "ccpo.admissible Sup (\<le>) (\<lambda>x. Q x)" | |
| 268 | shows "ccpo.admissible Sup (\<le>) (\<lambda>x. P x \<or> Q x)" | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 269 | proof (rule ccpo.admissibleI) | 
| 63612 | 270 | fix A :: "'a set" | 
| 67399 | 271 | assume chain: "chain (\<le>) A" | 
| 63810 | 272 |   assume A: "A \<noteq> {}" and P_Q: "\<forall>x\<in>A. P x \<or> Q x"
 | 
| 273 | have "(\<exists>x\<in>A. P x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y) \<or> (\<exists>x\<in>A. Q x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> Q y)" | |
| 274 | (is "?P \<or> ?Q" is "?P1 \<and> ?P2 \<or> _") | |
| 275 | proof (rule disjCI) | |
| 276 | assume "\<not> ?Q" | |
| 277 | then consider "\<forall>x\<in>A. \<not> Q x" | a where "a \<in> A" "\<forall>y\<in>A. a \<le> y \<longrightarrow> \<not> Q y" | |
| 278 | by blast | |
| 279 | then show ?P | |
| 280 | proof cases | |
| 281 | case 1 | |
| 282 | with P_Q have "\<forall>x\<in>A. P x" by blast | |
| 283 | with A show ?P by blast | |
| 284 | next | |
| 285 | case 2 | |
| 286 | note a = \<open>a \<in> A\<close> | |
| 287 | show ?P | |
| 288 | proof | |
| 289 | from P_Q 2 have *: "\<forall>y\<in>A. a \<le> y \<longrightarrow> P y" by blast | |
| 290 | with a have "P a" by blast | |
| 291 | with a show ?P1 by blast | |
| 292 | show ?P2 | |
| 293 | proof | |
| 294 | fix x | |
| 295 | assume x: "x \<in> A" | |
| 296 | with chain a show "\<exists>y\<in>A. x \<le> y \<and> P y" | |
| 297 | proof (rule chainE) | |
| 298 | assume le: "a \<le> x" | |
| 299 | with * a x have "P x" by blast | |
| 300 | with x le show ?thesis by blast | |
| 301 | next | |
| 302 | assume "a \<ge> x" | |
| 303 | with a \<open>P a\<close> show ?thesis by blast | |
| 304 | qed | |
| 305 | qed | |
| 306 | qed | |
| 307 | qed | |
| 308 | qed | |
| 309 | moreover | |
| 73252 | 310 |   have "Sup A = Sup {x \<in> A. P x}" if "\<And>x. x\<in>A \<Longrightarrow> \<exists>y\<in>A. x \<le> y \<and> P y" for P
 | 
| 73411 | 311 | proof (rule order.antisym) | 
| 67399 | 312 |     have chain_P: "chain (\<le>) {x \<in> A. P x}"
 | 
| 63810 | 313 | by (rule chain_compr [OF chain]) | 
| 314 |     show "Sup A \<le> Sup {x \<in> A. P x}"
 | |
| 73252 | 315 | proof (rule ccpo_Sup_least [OF chain]) | 
| 316 |       show "\<And>x. x \<in> A \<Longrightarrow> x \<le> \<Squnion> {x \<in> A. P x}"
 | |
| 317 | by (blast intro: ccpo_Sup_upper[OF chain_P] order_trans dest: that) | |
| 318 | qed | |
| 63810 | 319 |     show "Sup {x \<in> A. P x} \<le> Sup A"
 | 
| 320 | apply (rule ccpo_Sup_least [OF chain_P]) | |
| 321 | apply (simp add: ccpo_Sup_upper [OF chain]) | |
| 322 | done | |
| 323 | qed | |
| 324 | ultimately | |
| 325 |   consider "\<exists>x. x \<in> A \<and> P x" "Sup A = Sup {x \<in> A. P x}"
 | |
| 326 |     | "\<exists>x. x \<in> A \<and> Q x" "Sup A = Sup {x \<in> A. Q x}"
 | |
| 327 | by blast | |
| 63612 | 328 | then show "P (Sup A) \<or> Q (Sup A)" | 
| 73252 | 329 | proof cases | 
| 330 | case 1 | |
| 331 | then show ?thesis | |
| 332 | using ccpo.admissibleD [OF P chain_compr [OF chain]] by force | |
| 333 | next | |
| 334 | case 2 | |
| 335 | then show ?thesis | |
| 336 | using ccpo.admissibleD [OF Q chain_compr [OF chain]] by force | |
| 337 | qed | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 338 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 339 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 340 | end | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 341 | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 342 | instance complete_lattice \<subseteq> ccpo | 
| 61169 | 343 | by standard (fast intro: Sup_upper Sup_least)+ | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 344 | |
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 345 | lemma lfp_eq_fixp: | 
| 63979 | 346 | assumes mono: "mono f" | 
| 63612 | 347 | shows "lfp f = fixp f" | 
| 73411 | 348 | proof (rule order.antisym) | 
| 67399 | 349 | from mono have f': "monotone (\<le>) (\<le>) f" | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 350 | unfolding mono_def monotone_def . | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 351 | show "lfp f \<le> fixp f" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 352 | by (rule lfp_lowerbound, subst fixp_unfold [OF f'], rule order_refl) | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 353 | show "fixp f \<le> lfp f" | 
| 63979 | 354 | by (rule fixp_lowerbound [OF f']) (simp add: lfp_fixpoint [OF mono]) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 355 | qed | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 356 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 357 | hide_const (open) iterates fixp | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 358 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 359 | end |