| author | blanchet | 
| Tue, 01 Apr 2014 10:51:29 +0200 | |
| changeset 56346 | 42533f8f4729 | 
| parent 54630 | 9061af4d5ebc | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 1 | (* Title: HOL/Complete_Partial_Order.thy | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 2 | Author: Brian Huffman, Portland State University | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 3 | Author: Alexander Krauss, TU Muenchen | 
| 
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 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 6 | header {* Chain-complete partial orders and their fixpoints *}
 | 
| 
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 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 9 | imports Product_Type | 
| 
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 | subsection {* Monotone functions *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 13 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 14 | text {* Dictionary-passing version of @{const Orderings.mono}. *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 15 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 16 | definition monotone :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 17 | where "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 18 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 19 | lemma monotoneI[intro?]: "(\<And>x y. orda x y \<Longrightarrow> ordb (f x) (f y)) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 20 | \<Longrightarrow> monotone orda ordb f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 21 | unfolding monotone_def by iprover | 
| 
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 monotoneD[dest?]: "monotone orda ordb f \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 24 | unfolding monotone_def by iprover | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 25 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 26 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 27 | subsection {* Chains *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 28 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 29 | text {* A chain is a totally-ordered set. Chains are parameterized over
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 30 | the order for maximal flexibility, since type classes are not enough. | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 31 | *} | 
| 
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 | definition | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 34 |   chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 35 | where | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 36 | "chain ord S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. 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 | 37 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 38 | lemma chainI: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 39 | 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 | 40 | shows "chain ord S" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 41 | using assms unfolding chain_def by fast | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 42 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 43 | lemma chainD: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 44 | 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 | 45 | shows "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 | 46 | using assms unfolding chain_def by fast | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 47 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 48 | lemma chainE: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 49 | 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 | 50 | obtains "ord x y" | "ord y x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 51 | using assms unfolding chain_def by fast | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 52 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 53 | lemma chain_empty: "chain ord {}"
 | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 54 | by(simp add: chain_def) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 55 | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 56 | subsection {* Chain-complete partial orders *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 57 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 58 | text {*
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 59 | A ccpo has a least upper bound for any chain. In particular, the | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 60 | empty set is a chain, so every ccpo must have a bottom element. | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 61 | *} | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 62 | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 63 | class ccpo = order + Sup + | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 64 | assumes ccpo_Sup_upper: "\<lbrakk>chain (op \<le>) A; x \<in> A\<rbrakk> \<Longrightarrow> x \<le> Sup A" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 65 | assumes ccpo_Sup_least: "\<lbrakk>chain (op \<le>) A; \<And>x. x \<in> A \<Longrightarrow> x \<le> z\<rbrakk> \<Longrightarrow> Sup A \<le> z" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 66 | begin | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 67 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 68 | subsection {* Transfinite iteration of a function *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 69 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 70 | inductive_set iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set"
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 71 | for f :: "'a \<Rightarrow> 'a" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 72 | where | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 73 | step: "x \<in> iterates f \<Longrightarrow> f x \<in> iterates f" | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 74 | | Sup: "chain (op \<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 | 75 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 76 | lemma iterates_le_f: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 77 | "x \<in> iterates f \<Longrightarrow> monotone (op \<le>) (op \<le>) f \<Longrightarrow> x \<le> f x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 78 | by (induct x rule: iterates.induct) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 79 | (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 | 80 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 81 | lemma chain_iterates: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 82 | assumes f: "monotone (op \<le>) (op \<le>) f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 83 | shows "chain (op \<le>) (iterates f)" (is "chain _ ?C") | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 84 | proof (rule chainI) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 85 | fix x y assume "x \<in> ?C" "y \<in> ?C" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 86 | 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 | 87 | proof (induct x arbitrary: y rule: iterates.induct) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 88 | fix x y assume y: "y \<in> ?C" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 89 | and IH: "\<And>z. z \<in> ?C \<Longrightarrow> x \<le> z \<or> z \<le> x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 90 | 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 | 91 | proof (induct y rule: iterates.induct) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 92 | case (step y) with IH f show ?case by (auto dest: monotoneD) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 93 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 94 | case (Sup M) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 95 | then have chM: "chain (op \<le>) M" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 96 | 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 | 97 | 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 | 98 | proof (cases "\<exists>z\<in>M. f x \<le> z") | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 99 | case True then have "f x \<le> Sup M" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 100 | apply rule | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 101 | apply (erule order_trans) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 102 | by (rule ccpo_Sup_upper[OF chM]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 103 | thus ?thesis .. | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 104 | next | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 105 | case False with IH' | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 106 | show ?thesis 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 | 107 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 108 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 109 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 110 | case (Sup M y) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 111 | show ?case | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 112 | proof (cases "\<exists>x\<in>M. y \<le> x") | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 113 | case True then have "y \<le> Sup M" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 114 | apply rule | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 115 | apply (erule order_trans) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 116 | by (rule ccpo_Sup_upper[OF Sup(1)]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 117 | thus ?thesis .. | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 118 | next | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 119 | case False with Sup | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 120 | 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 | 121 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 122 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 123 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 124 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 125 | lemma bot_in_iterates: "Sup {} \<in> iterates f"
 | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 126 | by(auto intro: iterates.Sup simp add: chain_empty) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 127 | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 128 | subsection {* Fixpoint combinator *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 129 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 130 | definition | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 131 |   fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 132 | where | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 133 | "fixp f = Sup (iterates f)" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 134 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 135 | lemma iterates_fixp: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 136 | assumes f: "monotone (op \<le>) (op \<le>) f" shows "fixp f \<in> iterates f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 137 | unfolding fixp_def | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 138 | 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 | 139 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 140 | lemma fixp_unfold: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 141 | assumes f: "monotone (op \<le>) (op \<le>) f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 142 | shows "fixp f = f (fixp f)" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 143 | proof (rule antisym) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 144 | 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 | 145 | by (intro iterates_le_f iterates_fixp f) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 146 | have "f (fixp f) \<le> Sup (iterates f)" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 147 | by (intro ccpo_Sup_upper chain_iterates f iterates.step iterates_fixp) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 148 | thus "f (fixp f) \<le> fixp f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 149 | unfolding fixp_def . | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 150 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 151 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 152 | lemma fixp_lowerbound: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 153 | assumes f: "monotone (op \<le>) (op \<le>) f" and z: "f z \<le> z" shows "fixp f \<le> z" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 154 | unfolding fixp_def | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 155 | proof (rule ccpo_Sup_least[OF chain_iterates[OF f]]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 156 | fix x assume "x \<in> iterates f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 157 | thus "x \<le> z" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 158 | proof (induct x rule: iterates.induct) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 159 | fix x assume "x \<le> z" with f have "f x \<le> f z" by (rule monotoneD) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 160 | also note z finally show "f x \<le> z" . | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 161 | qed (auto intro: ccpo_Sup_least) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 162 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 163 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 164 | end | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 165 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 166 | subsection {* Fixpoint induction *}
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 167 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 168 | setup {* Sign.map_naming (Name_Space.mandatory_path "ccpo") *}
 | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 169 | |
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 170 | definition admissible :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 171 | where "admissible lub ord P = (\<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 | 172 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 173 | lemma admissibleI: | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 174 |   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 | 175 | shows "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 | 176 | 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 | 177 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 178 | lemma admissibleD: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 179 | 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 | 180 | assumes "chain ord A" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 181 |   assumes "A \<noteq> {}"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 182 | 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 | 183 | shows "P (lub A)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 184 | 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 | 185 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 186 | setup {* Sign.map_naming Name_Space.parent_path *}
 | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 187 | |
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 188 | lemma (in ccpo) fixp_induct: | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 189 | assumes adm: "ccpo.admissible Sup (op \<le>) P" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 190 | assumes mono: "monotone (op \<le>) (op \<le>) f" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 191 |   assumes bot: "P (Sup {})"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 192 | 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 | 193 | shows "P (fixp f)" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 194 | unfolding fixp_def 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 | 195 | proof (rule ccpo.admissibleD) | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 196 |   show "iterates f \<noteq> {}" using bot_in_iterates by auto
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 197 | fix x assume "x \<in> iterates f" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 198 | thus "P x" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 199 | by (induct rule: iterates.induct) | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 200 |       (case_tac "M = {}", auto intro: step bot ccpo.admissibleD adm)
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 201 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 202 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 203 | lemma admissible_True: "ccpo.admissible lub ord (\<lambda>x. True)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 204 | 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 | 205 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 206 | (*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 | 207 | 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 | 208 | *) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 209 | lemma admissible_const: "ccpo.admissible lub ord (\<lambda>x. t)" | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 210 | by(auto intro: ccpo.admissibleI) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 211 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 212 | lemma admissible_conj: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 213 | 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 | 214 | 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 | 215 | shows "ccpo.admissible lub ord (\<lambda>x. P x \<and> Q x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 216 | 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 | 217 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 218 | lemma admissible_all: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 219 | 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 | 220 | shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y. P x y)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 221 | 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 | 222 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 223 | lemma admissible_ball: | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 224 | 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 | 225 | shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y\<in>A. P x y)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 226 | 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 | 227 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 228 | lemma chain_compr: "chain ord A \<Longrightarrow> chain ord {x \<in> A. P x}"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 229 | unfolding chain_def by fast | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 230 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 231 | context ccpo begin | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 232 | |
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 233 | lemma admissible_disj_lemma: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 234 | assumes A: "chain (op \<le>)A" | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 235 | assumes P: "\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y" | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 236 |   shows "Sup A = Sup {x \<in> A. P x}"
 | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 237 | proof (rule antisym) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 238 |   have *: "chain (op \<le>) {x \<in> A. P x}"
 | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 239 | by (rule chain_compr [OF A]) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 240 |   show "Sup A \<le> Sup {x \<in> A. P x}"
 | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 241 | apply (rule ccpo_Sup_least [OF A]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 242 | apply (drule P [rule_format], clarify) | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 243 | apply (erule order_trans) | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 244 | apply (simp add: ccpo_Sup_upper [OF *]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 245 | done | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 246 |   show "Sup {x \<in> A. P x} \<le> Sup A"
 | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 247 | apply (rule ccpo_Sup_least [OF *]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 248 | apply clarify | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 249 | apply (simp add: ccpo_Sup_upper [OF A]) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 250 | done | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 251 | qed | 
| 
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_disj: | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 254 | fixes P Q :: "'a \<Rightarrow> bool" | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 255 | assumes P: "ccpo.admissible Sup (op \<le>) (\<lambda>x. P x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 256 | assumes Q: "ccpo.admissible Sup (op \<le>) (\<lambda>x. Q x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 257 | shows "ccpo.admissible Sup (op \<le>) (\<lambda>x. P x \<or> Q x)" | 
| 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 258 | proof (rule ccpo.admissibleI) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 259 | fix A :: "'a set" assume A: "chain (op \<le>) A" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 260 |   assume "A \<noteq> {}"
 | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 261 | and "\<forall>x\<in>A. P x \<or> Q x" | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 262 | hence "(\<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)" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 263 | using chainD[OF A] by blast | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 264 |   hence "(\<exists>x. x \<in> A \<and> P x) \<and> Sup A = Sup {x \<in> A. P x} \<or> (\<exists>x. x \<in> A \<and> Q x) \<and> Sup A = Sup {x \<in> A. Q x}"
 | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 265 | using admissible_disj_lemma [OF A] by blast | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 266 | thus "P (Sup A) \<or> Q (Sup A)" | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 267 | apply (rule disjE, simp_all) | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 268 | apply (rule disjI1, rule ccpo.admissibleD [OF P chain_compr [OF A]], simp, simp) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53361diff
changeset | 269 | apply (rule disjI2, rule ccpo.admissibleD [OF Q chain_compr [OF A]], simp, simp) | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 270 | done | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 271 | qed | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 272 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 273 | end | 
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 274 | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 275 | instance complete_lattice \<subseteq> ccpo | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 276 | by default (fast intro: Sup_upper Sup_least)+ | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 277 | |
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 278 | lemma lfp_eq_fixp: | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 279 | assumes f: "mono f" shows "lfp f = fixp f" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 280 | proof (rule antisym) | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 281 | from f have f': "monotone (op \<le>) (op \<le>) f" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 282 | unfolding mono_def monotone_def . | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 283 | show "lfp f \<le> fixp f" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 284 | 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 | 285 | show "fixp f \<le> lfp f" | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 286 | by (rule fixp_lowerbound [OF f'], subst lfp_unfold [OF f], rule order_refl) | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 287 | qed | 
| 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
40252diff
changeset | 288 | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
46041diff
changeset | 289 | hide_const (open) iterates fixp | 
| 40106 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 290 | |
| 
c58951943cba
Complete_Partial_Order.thy: complete partial orders over arbitrary chains, with fixpoint theorem
 krauss parents: diff
changeset | 291 | end |