| author | nipkow | 
| Wed, 03 Oct 2018 09:46:42 +0200 | |
| changeset 69107 | c2de7a5c8de9 | 
| parent 69039 | 51005671bee5 | 
| child 69164 | 74f1b0f10b2b | 
| permissions | -rw-r--r-- | 
| 62858 | 1 | (* Title: HOL/Library/Complete_Partial_Order2.thy | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 2 | Author: Andreas Lochbihler, ETH Zurich | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 3 | *) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 4 | |
| 62837 | 5 | section \<open>Formalisation of chain-complete partial orders, continuity and admissibility\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 6 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 7 | theory Complete_Partial_Order2 imports | 
| 65366 | 8 | Main Lattice_Syntax | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 9 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 10 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 11 | lemma chain_transfer [transfer_rule]: | 
| 63343 | 12 | includes lifting_syntax | 
| 67399 | 13 | shows "((A ===> A ===> (=)) ===> rel_set A ===> (=)) Complete_Partial_Order.chain Complete_Partial_Order.chain" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 14 | unfolding chain_def[abs_def] by transfer_prover | 
| 68980 
5717fbc55521
added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
 nipkow parents: 
67399diff
changeset | 15 | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 16 | lemma linorder_chain [simp, intro!]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 17 | fixes Y :: "_ :: linorder set" | 
| 67399 | 18 | shows "Complete_Partial_Order.chain (\<le>) Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 19 | by(auto intro: chainI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 20 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 21 | lemma fun_lub_apply: "\<And>Sup. fun_lub Sup Y x = Sup ((\<lambda>f. f x) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 22 | by(simp add: fun_lub_def image_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 23 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 24 | lemma fun_lub_empty [simp]: "fun_lub lub {} = (\<lambda>_. lub {})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 25 | by(rule ext)(simp add: fun_lub_apply) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 26 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 27 | lemma chain_fun_ordD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 28 | assumes "Complete_Partial_Order.chain (fun_ord le) Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 29 | shows "Complete_Partial_Order.chain le ((\<lambda>f. f x) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 30 | by(rule chainI)(auto dest: chainD[OF assms] simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 31 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 32 | lemma chain_Diff: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 33 | "Complete_Partial_Order.chain ord A | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 34 | \<Longrightarrow> Complete_Partial_Order.chain ord (A - B)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 35 | by(erule chain_subset) blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 36 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 37 | lemma chain_rel_prodD1: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 38 | "Complete_Partial_Order.chain (rel_prod orda ordb) Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 39 | \<Longrightarrow> Complete_Partial_Order.chain orda (fst ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 40 | by(auto 4 3 simp add: chain_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 41 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 42 | lemma chain_rel_prodD2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 43 | "Complete_Partial_Order.chain (rel_prod orda ordb) Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 44 | \<Longrightarrow> Complete_Partial_Order.chain ordb (snd ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 45 | by(auto 4 3 simp add: chain_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 46 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 47 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 48 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 49 | |
| 67399 | 50 | lemma ccpo_fun: "class.ccpo (fun_lub Sup) (fun_ord (\<le>)) (mk_less (fun_ord (\<le>)))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 51 | by standard (auto 4 3 simp add: mk_less_def fun_ord_def fun_lub_apply | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 52 | intro: order.trans antisym chain_imageI ccpo_Sup_upper ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 53 | |
| 67399 | 54 | lemma ccpo_Sup_below_iff: "Complete_Partial_Order.chain (\<le>) Y \<Longrightarrow> Sup Y \<le> x \<longleftrightarrow> (\<forall>y\<in>Y. y \<le> x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 55 | by(fast intro: order_trans[OF ccpo_Sup_upper] ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 56 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 57 | lemma Sup_minus_bot: | 
| 67399 | 58 | assumes chain: "Complete_Partial_Order.chain (\<le>) A" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 59 |   shows "\<Squnion>(A - {\<Squnion>{}}) = \<Squnion>A"
 | 
| 63649 | 60 | (is "?lhs = ?rhs") | 
| 61 | proof (rule antisym) | |
| 62 | show "?lhs \<le> ?rhs" | |
| 63 | by (blast intro: ccpo_Sup_least chain_Diff[OF chain] ccpo_Sup_upper[OF chain]) | |
| 64 | show "?rhs \<le> ?lhs" | |
| 65 | proof (rule ccpo_Sup_least [OF chain]) | |
| 66 | show "x \<in> A \<Longrightarrow> x \<le> ?lhs" for x | |
| 67 |       by (cases "x = \<Squnion>{}")
 | |
| 68 | (blast intro: ccpo_Sup_least chain_empty ccpo_Sup_upper[OF chain_Diff[OF chain]])+ | |
| 69 | qed | |
| 70 | qed | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 71 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 72 | lemma mono_lub: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 73 | fixes le_b (infix "\<sqsubseteq>" 60) | 
| 67399 | 74 | assumes chain: "Complete_Partial_Order.chain (fun_ord (\<le>)) Y" | 
| 75 | and mono: "\<And>f. f \<in> Y \<Longrightarrow> monotone le_b (\<le>) f" | |
| 76 | shows "monotone (\<sqsubseteq>) (\<le>) (fun_lub Sup Y)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 77 | proof(rule monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 78 | fix x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 79 | assume "x \<sqsubseteq> y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 80 | |
| 67399 | 81 | have chain'': "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 82 | using chain by(rule chain_imageI)(simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 83 | then show "fun_lub Sup Y x \<le> fun_lub Sup Y y" unfolding fun_lub_apply | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 84 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 85 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 86 | assume "x' \<in> (\<lambda>f. f x) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 87 | then obtain f where "f \<in> Y" "x' = f x" by blast | 
| 62837 | 88 | note \<open>x' = f x\<close> also | 
| 89 | from \<open>f \<in> Y\<close> \<open>x \<sqsubseteq> y\<close> have "f x \<le> f y" by(blast dest: mono monotoneD) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 90 | also have "\<dots> \<le> \<Squnion>((\<lambda>f. f y) ` Y)" using chain'' | 
| 62837 | 91 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> Y\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 92 | finally show "x' \<le> \<Squnion>((\<lambda>f. f y) ` Y)" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 93 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 94 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 95 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 96 | context | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 97 | fixes le_b (infix "\<sqsubseteq>" 60) and Y f | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 98 | assumes chain: "Complete_Partial_Order.chain le_b Y" | 
| 67399 | 99 | and mono1: "\<And>y. y \<in> Y \<Longrightarrow> monotone le_b (\<le>) (\<lambda>x. f x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 100 | and mono2: "\<And>x a b. \<lbrakk> x \<in> Y; a \<sqsubseteq> b; a \<in> Y; b \<in> Y \<rbrakk> \<Longrightarrow> f x a \<le> f x b" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 101 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 102 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 103 | lemma Sup_mono: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 104 | assumes le: "x \<sqsubseteq> y" and x: "x \<in> Y" and y: "y \<in> Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 105 | shows "\<Squnion>(f x ` Y) \<le> \<Squnion>(f y ` Y)" (is "_ \<le> ?rhs") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 106 | proof(rule ccpo_Sup_least) | 
| 67399 | 107 | from chain show chain': "Complete_Partial_Order.chain (\<le>) (f x ` Y)" when "x \<in> Y" for x | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 108 | by(rule chain_imageI) (insert that, auto dest: mono2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 109 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 110 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 111 | assume "x' \<in> f x ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 112 | then obtain y' where "y' \<in> Y" "x' = f x y'" by blast note this(2) | 
| 62837 | 113 | also from mono1[OF \<open>y' \<in> Y\<close>] le have "\<dots> \<le> f y y'" by(rule monotoneD) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 114 | also have "\<dots> \<le> ?rhs" using chain'[OF y] | 
| 62837 | 115 | by (auto intro!: ccpo_Sup_upper simp add: \<open>y' \<in> Y\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 116 | finally show "x' \<le> ?rhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 117 | qed(rule x) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 118 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 119 | lemma diag_Sup: "\<Squnion>((\<lambda>x. \<Squnion>(f x ` Y)) ` Y) = \<Squnion>((\<lambda>x. f x x) ` Y)" (is "?lhs = ?rhs") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 120 | proof(rule antisym) | 
| 67399 | 121 | have chain1: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>(f x ` Y)) ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 122 | using chain by(rule chain_imageI)(rule Sup_mono) | 
| 67399 | 123 | have chain2: "\<And>y'. y' \<in> Y \<Longrightarrow> Complete_Partial_Order.chain (\<le>) (f y' ` Y)" using chain | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 124 | by(rule chain_imageI)(auto dest: mono2) | 
| 67399 | 125 | have chain3: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. f x x) ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 126 | using chain by(rule chain_imageI)(auto intro: monotoneD[OF mono1] mono2 order.trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 127 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 128 | show "?lhs \<le> ?rhs" using chain1 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 129 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 130 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 131 | assume "x' \<in> (\<lambda>x. \<Squnion>(f x ` Y)) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 132 | then obtain y' where "y' \<in> Y" "x' = \<Squnion>(f y' ` Y)" by blast note this(2) | 
| 62837 | 133 | also have "\<dots> \<le> ?rhs" using chain2[OF \<open>y' \<in> Y\<close>] | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 134 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 135 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 136 | assume "x \<in> f y' ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 137 | then obtain y where "y \<in> Y" and x: "x = f y' y" by blast | 
| 63040 | 138 | define y'' where "y'' = (if y \<sqsubseteq> y' then y' else y)" | 
| 62837 | 139 | from chain \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "y \<sqsubseteq> y' \<or> y' \<sqsubseteq> y" by(rule chainD) | 
| 140 | hence "f y' y \<le> f y'' y''" using \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 141 | by(auto simp add: y''_def intro: mono2 monotoneD[OF mono1]) | 
| 62837 | 142 | also from \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "y'' \<in> Y" by(simp add: y''_def) | 
| 143 | from chain3 have "f y'' y'' \<le> ?rhs" by(rule ccpo_Sup_upper)(simp add: \<open>y'' \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 144 | finally show "x \<le> ?rhs" by(simp add: x) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 145 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 146 | finally show "x' \<le> ?rhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 147 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 148 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 149 | show "?rhs \<le> ?lhs" using chain3 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 150 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 151 | fix y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 152 | assume "y \<in> (\<lambda>x. f x x) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 153 | then obtain x where "x \<in> Y" and "y = f x x" by blast note this(2) | 
| 62837 | 154 | also from chain2[OF \<open>x \<in> Y\<close>] have "\<dots> \<le> \<Squnion>(f x ` Y)" | 
| 155 | by(rule ccpo_Sup_upper)(simp add: \<open>x \<in> Y\<close>) | |
| 156 | also have "\<dots> \<le> ?lhs" by(rule ccpo_Sup_upper[OF chain1])(simp add: \<open>x \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 157 | finally show "y \<le> ?lhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 158 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 159 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 160 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 161 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 162 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 163 | lemma Sup_image_mono_le: | 
| 69038 | 164 |   fixes le_b (infix "\<sqsubseteq>" 60) and Sup_b ("\<Or>")
 | 
| 67399 | 165 | assumes ccpo: "class.ccpo Sup_b (\<sqsubseteq>) lt_b" | 
| 166 | assumes chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 167 | and mono: "\<And>x y. \<lbrakk> x \<sqsubseteq> y; x \<in> Y \<rbrakk> \<Longrightarrow> f x \<le> f y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 168 | shows "Sup (f ` Y) \<le> f (\<Or>Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 169 | proof(rule ccpo_Sup_least) | 
| 67399 | 170 | show "Complete_Partial_Order.chain (\<le>) (f ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 171 | using chain by(rule chain_imageI)(rule mono) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 172 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 173 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 174 | assume "x \<in> f ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 175 | then obtain y where "y \<in> Y" and "x = f y" by blast note this(2) | 
| 62837 | 176 | also have "y \<sqsubseteq> \<Or>Y" using ccpo chain \<open>y \<in> Y\<close> by(rule ccpo.ccpo_Sup_upper) | 
| 177 | hence "f y \<le> f (\<Or>Y)" using \<open>y \<in> Y\<close> by(rule mono) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 178 | finally show "x \<le> \<dots>" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 179 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 180 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 181 | lemma swap_Sup: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 182 | fixes le_b (infix "\<sqsubseteq>" 60) | 
| 67399 | 183 | assumes Y: "Complete_Partial_Order.chain (\<sqsubseteq>) Y" | 
| 184 | and Z: "Complete_Partial_Order.chain (fun_ord (\<le>)) Z" | |
| 185 | and mono: "\<And>f. f \<in> Z \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) f" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 186 | shows "\<Squnion>((\<lambda>x. \<Squnion>(x ` Y)) ` Z) = \<Squnion>((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 187 | (is "?lhs = ?rhs") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 188 | proof(cases "Y = {}")
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 189 | case True | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 190 | then show ?thesis | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 191 | by (simp add: image_constant_conv cong del: strong_SUP_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 192 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 193 | case False | 
| 67399 | 194 | have chain1: "\<And>f. f \<in> Z \<Longrightarrow> Complete_Partial_Order.chain (\<le>) (f ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 195 | by(rule chain_imageI[OF Y])(rule monotoneD[OF mono]) | 
| 67399 | 196 | have chain2: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>(x ` Y)) ` Z)" using Z | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 197 | proof(rule chain_imageI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 198 | fix f g | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 199 | assume "f \<in> Z" "g \<in> Z" | 
| 67399 | 200 | and "fun_ord (\<le>) f g" | 
| 62837 | 201 | from chain1[OF \<open>f \<in> Z\<close>] show "\<Squnion>(f ` Y) \<le> \<Squnion>(g ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 202 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 203 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 204 | assume "x \<in> f ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 205 | then obtain y where "y \<in> Y" "x = f y" by blast note this(2) | 
| 67399 | 206 | also have "\<dots> \<le> g y" using \<open>fun_ord (\<le>) f g\<close> by(simp add: fun_ord_def) | 
| 62837 | 207 | also have "\<dots> \<le> \<Squnion>(g ` Y)" using chain1[OF \<open>g \<in> Z\<close>] | 
| 208 | by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 209 | finally show "x \<le> \<Squnion>(g ` Y)" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 210 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 211 | qed | 
| 67399 | 212 | have chain3: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` Z)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 213 | using Z by(rule chain_imageI)(simp add: fun_ord_def) | 
| 67399 | 214 | have chain4: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 215 | using Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 216 | proof(rule chain_imageI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 217 | fix f x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 218 | assume "x \<sqsubseteq> y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 219 | show "\<Squnion>((\<lambda>f. f x) ` Z) \<le> \<Squnion>((\<lambda>f. f y) ` Z)" (is "_ \<le> ?rhs") using chain3 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 220 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 221 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 222 | assume "x' \<in> (\<lambda>f. f x) ` Z" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 223 | then obtain f where "f \<in> Z" "x' = f x" by blast note this(2) | 
| 62837 | 224 | also have "f x \<le> f y" using \<open>f \<in> Z\<close> \<open>x \<sqsubseteq> y\<close> by(rule monotoneD[OF mono]) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 225 | also have "f y \<le> ?rhs" using chain3 | 
| 62837 | 226 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> Z\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 227 | finally show "x' \<le> ?rhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 228 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 229 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 230 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 231 | from chain2 have "?lhs \<le> ?rhs" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 232 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 233 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 234 | assume "x \<in> (\<lambda>x. \<Squnion>(x ` Y)) ` Z" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 235 | then obtain f where "f \<in> Z" "x = \<Squnion>(f ` Y)" by blast note this(2) | 
| 62837 | 236 | also have "\<dots> \<le> ?rhs" using chain1[OF \<open>f \<in> Z\<close>] | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 237 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 238 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 239 | assume "x' \<in> f ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 240 | then obtain y where "y \<in> Y" "x' = f y" by blast note this(2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 241 | also have "f y \<le> \<Squnion>((\<lambda>f. f y) ` Z)" using chain3 | 
| 62837 | 242 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> Z\<close>) | 
| 243 | also have "\<dots> \<le> ?rhs" using chain4 by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 244 | finally show "x' \<le> ?rhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 245 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 246 | finally show "x \<le> ?rhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 247 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 248 | moreover | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 249 | have "?rhs \<le> ?lhs" using chain4 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 250 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 251 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 252 | assume "x \<in> (\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 253 | then obtain y where "y \<in> Y" "x = \<Squnion>((\<lambda>f. f y) ` Z)" by blast note this(2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 254 | also have "\<dots> \<le> ?lhs" using chain3 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 255 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 256 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 257 | assume "x' \<in> (\<lambda>f. f y) ` Z" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 258 | then obtain f where "f \<in> Z" "x' = f y" by blast note this(2) | 
| 62837 | 259 | also have "f y \<le> \<Squnion>(f ` Y)" using chain1[OF \<open>f \<in> Z\<close>] | 
| 260 | by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 261 | also have "\<dots> \<le> ?lhs" using chain2 | 
| 62837 | 262 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> Z\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 263 | finally show "x' \<le> ?lhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 264 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 265 | finally show "x \<le> ?lhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 266 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 267 | ultimately show "?lhs = ?rhs" by(rule antisym) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 268 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 269 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 270 | lemma fixp_mono: | 
| 67399 | 271 | assumes fg: "fun_ord (\<le>) f g" | 
| 272 | and f: "monotone (\<le>) (\<le>) f" | |
| 273 | and g: "monotone (\<le>) (\<le>) g" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 274 | shows "ccpo_class.fixp f \<le> ccpo_class.fixp g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 275 | unfolding fixp_def | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 276 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 277 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 278 | assume "x \<in> ccpo_class.iterates f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 279 | thus "x \<le> \<Squnion>ccpo_class.iterates g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 280 | proof induction | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 281 | case (step x) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 282 | from f step.IH have "f x \<le> f (\<Squnion>ccpo_class.iterates g)" by(rule monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 283 | also have "\<dots> \<le> g (\<Squnion>ccpo_class.iterates g)" using fg by(simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 284 | also have "\<dots> = \<Squnion>ccpo_class.iterates g" by(fold fixp_def fixp_unfold[OF g]) simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 285 | finally show ?case . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 286 | qed(blast intro: ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 287 | qed(rule chain_iterates[OF f]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 288 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 289 | context fixes ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubseteq>" 60) begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 290 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 291 | lemma iterates_mono: | 
| 67399 | 292 | assumes f: "f \<in> ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F" | 
| 293 | and mono: "\<And>f. monotone (\<sqsubseteq>) (\<le>) f \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) (F f)" | |
| 294 | shows "monotone (\<sqsubseteq>) (\<le>) f" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 295 | using f | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 296 | by(induction rule: ccpo.iterates.induct[OF ccpo_fun, consumes 1, case_names step Sup])(blast intro: mono mono_lub)+ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 297 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 298 | lemma fixp_preserves_mono: | 
| 67399 | 299 | assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. F f x)" | 
| 300 | and mono2: "\<And>f. monotone (\<sqsubseteq>) (\<le>) f \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) (F f)" | |
| 301 | shows "monotone (\<sqsubseteq>) (\<le>) (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) F)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 302 | (is "monotone _ _ ?fixp") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 303 | proof(rule monotoneI) | 
| 67399 | 304 | have mono: "monotone (fun_ord (\<le>)) (fun_ord (\<le>)) F" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 305 | by(rule monotoneI)(auto simp add: fun_ord_def intro: monotoneD[OF mono]) | 
| 67399 | 306 | let ?iter = "ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F" | 
| 307 | have chain: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` ?iter)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 308 | by(rule chain_imageI[OF ccpo.chain_iterates[OF ccpo_fun mono]])(simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 309 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 310 | fix x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 311 | assume "x \<sqsubseteq> y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 312 | show "?fixp x \<le> ?fixp y" | 
| 63170 | 313 | apply (simp only: ccpo.fixp_def[OF ccpo_fun] fun_lub_apply) | 
| 314 | using chain | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 315 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 316 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 317 | assume "x' \<in> (\<lambda>f. f x) ` ?iter" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 318 | then obtain f where "f \<in> ?iter" "x' = f x" by blast note this(2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 319 | also have "f x \<le> f y" | 
| 62837 | 320 | by(rule monotoneD[OF iterates_mono[OF \<open>f \<in> ?iter\<close> mono2]])(blast intro: \<open>x \<sqsubseteq> y\<close>)+ | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 321 | also have "f y \<le> \<Squnion>((\<lambda>f. f y) ` ?iter)" using chain | 
| 62837 | 322 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> ?iter\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 323 | finally show "x' \<le> \<dots>" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 324 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 325 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 326 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 327 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 328 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 329 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 330 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 331 | lemma monotone2monotone: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 332 | assumes 2: "\<And>x. monotone ordb ordc (\<lambda>y. f x y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 333 | and t: "monotone orda ordb (\<lambda>x. t x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 334 | and 1: "\<And>y. monotone orda ordc (\<lambda>x. f x y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 335 | and trans: "transp ordc" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 336 | shows "monotone orda ordc (\<lambda>x. f x (t x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 337 | by(blast intro: monotoneI transpD[OF trans] monotoneD[OF t] monotoneD[OF 2] monotoneD[OF 1]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 338 | |
| 62837 | 339 | subsection \<open>Continuity\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 340 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 341 | definition cont :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 342 | where | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 343 | "cont luba orda lubb ordb f \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 344 |   (\<forall>Y. Complete_Partial_Order.chain orda Y \<longrightarrow> Y \<noteq> {} \<longrightarrow> f (luba Y) = lubb (f ` Y))"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 345 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 346 | definition mcont :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 347 | where | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 348 | "mcont luba orda lubb ordb f \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 349 | monotone orda ordb f \<and> cont luba orda lubb ordb f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 350 | |
| 62837 | 351 | subsubsection \<open>Theorem collection \<open>cont_intro\<close>\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 352 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 353 | named_theorems cont_intro "continuity and admissibility intro rules" | 
| 62837 | 354 | ML \<open> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 355 | (* apply cont_intro rules as intro and try to solve | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 356 | the remaining of the emerging subgoals with simp *) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 357 | fun cont_intro_tac ctxt = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 358 |   REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems cont_intro})))
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 359 | THEN_ALL_NEW (SOLVED' (simp_tac ctxt)) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 360 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 361 | fun cont_intro_simproc ctxt ct = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 362 | let | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 363 | fun mk_stmt t = t | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 364 | |> HOLogic.mk_Trueprop | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 365 | |> Thm.cterm_of ctxt | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 366 | |> Goal.init | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 367 | fun mk_thm t = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 368 | case SINGLE (cont_intro_tac ctxt 1) (mk_stmt t) of | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 369 |         SOME thm => SOME (Goal.finish ctxt thm RS @{thm Eq_TrueI})
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 370 | | NONE => NONE | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 371 | in | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 372 | case Thm.term_of ct of | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 373 |       t as Const (@{const_name ccpo.admissible}, _) $ _ $ _ $ _ => mk_thm t
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 374 |     | t as Const (@{const_name mcont}, _) $ _ $ _ $ _ $ _ $ _ => mk_thm t
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 375 |     | t as Const (@{const_name monotone}, _) $ _ $ _ $ _ => mk_thm t
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 376 | | _ => NONE | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 377 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 378 | handle THM _ => NONE | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 379 | | TYPE _ => NONE | 
| 62837 | 380 | \<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 381 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 382 | simproc_setup "cont_intro" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 383 | ( "ccpo.admissible lub ord P" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 384 | | "mcont lub ord lub' ord' f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 385 | | "monotone ord ord' f" | 
| 62837 | 386 | ) = \<open>K cont_intro_simproc\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 387 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 388 | lemmas [cont_intro] = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 389 | call_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 390 | let_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 391 | if_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 392 | option.const_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 393 | tailrec.const_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 394 | bind_mono | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 395 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 396 | declare if_mono[simp] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 397 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 398 | lemma monotone_id' [cont_intro]: "monotone ord ord (\<lambda>x. x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 399 | by(simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 400 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 401 | lemma monotone_applyI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 402 | "monotone orda ordb F \<Longrightarrow> monotone (fun_ord orda) ordb (\<lambda>f. F (f x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 403 | by(rule monotoneI)(auto simp add: fun_ord_def dest: monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 404 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 405 | lemma monotone_if_fun [partial_function_mono]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 406 | "\<lbrakk> monotone (fun_ord orda) (fun_ord ordb) F; monotone (fun_ord orda) (fun_ord ordb) G \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 407 | \<Longrightarrow> monotone (fun_ord orda) (fun_ord ordb) (\<lambda>f n. if c n then F f n else G f n)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 408 | by(simp add: monotone_def fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 409 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 410 | lemma monotone_fun_apply_fun [partial_function_mono]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 411 | "monotone (fun_ord (fun_ord ord)) (fun_ord ord) (\<lambda>f n. f t (g n))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 412 | by(rule monotoneI)(simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 413 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 414 | lemma monotone_fun_ord_apply: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 415 | "monotone orda (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. monotone orda ordb (\<lambda>y. f y x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 416 | by(auto simp add: monotone_def fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 417 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 418 | context preorder begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 419 | |
| 67399 | 420 | lemma transp_le [simp, cont_intro]: "transp (\<le>)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 421 | by(rule transpI)(rule order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 422 | |
| 67399 | 423 | lemma monotone_const [simp, cont_intro]: "monotone ord (\<le>) (\<lambda>_. c)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 424 | by(rule monotoneI) simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 425 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 426 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 427 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 428 | lemma transp_le [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 429 | "class.preorder ord (mk_less ord) \<Longrightarrow> transp ord" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 430 | by(rule preorder.transp_le) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 431 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 432 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 433 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 434 | declare const_mono [cont_intro, simp] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 435 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 436 | lemma transp_le [cont_intro, simp]: "transp leq" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 437 | by(rule transpI)(rule leq_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 438 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 439 | lemma preorder [cont_intro, simp]: "class.preorder leq (mk_less leq)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 440 | by(unfold_locales)(auto simp add: mk_less_def intro: leq_refl leq_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 441 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 442 | declare ccpo[cont_intro, simp] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 443 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 444 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 445 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 446 | lemma contI [intro?]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 447 |   "(\<And>Y. \<lbrakk> Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk> \<Longrightarrow> f (luba Y) = lubb (f ` Y)) 
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 448 | \<Longrightarrow> cont luba orda lubb ordb f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 449 | unfolding cont_def by blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 450 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 451 | lemma contD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 452 |   "\<lbrakk> cont luba orda lubb ordb f; Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk> 
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 453 | \<Longrightarrow> f (luba Y) = lubb (f ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 454 | unfolding cont_def by blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 455 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 456 | lemma cont_id [simp, cont_intro]: "\<And>Sup. cont Sup ord Sup ord id" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 457 | by(rule contI) simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 458 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 459 | lemma cont_id' [simp, cont_intro]: "\<And>Sup. cont Sup ord Sup ord (\<lambda>x. x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 460 | using cont_id[unfolded id_def] . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 461 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 462 | lemma cont_applyI [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 463 | assumes cont: "cont luba orda lubb ordb g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 464 | shows "cont (fun_lub luba) (fun_ord orda) lubb ordb (\<lambda>f. g (f x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 465 | by(rule contI)(drule chain_fun_ordD[where x=x], simp add: fun_lub_apply image_image contD[OF cont]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 466 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 467 | lemma call_cont: "cont (fun_lub lub) (fun_ord ord) lub ord (\<lambda>f. f t)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 468 | by(simp add: cont_def fun_lub_apply) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 469 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 470 | lemma cont_if [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 471 | "\<lbrakk> cont luba orda lubb ordb f; cont luba orda lubb ordb g \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 472 | \<Longrightarrow> cont luba orda lubb ordb (\<lambda>x. if c then f x else g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 473 | by(cases c) simp_all | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 474 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 475 | lemma mcontI [intro?]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 476 | "\<lbrakk> monotone orda ordb f; cont luba orda lubb ordb f \<rbrakk> \<Longrightarrow> mcont luba orda lubb ordb f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 477 | by(simp add: mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 478 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 479 | lemma mcont_mono: "mcont luba orda lubb ordb f \<Longrightarrow> monotone orda ordb f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 480 | by(simp add: mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 481 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 482 | lemma mcont_cont [simp]: "mcont luba orda lubb ordb f \<Longrightarrow> cont luba orda lubb ordb f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 483 | by(simp add: mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 484 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 485 | lemma mcont_monoD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 486 | "\<lbrakk> mcont luba orda lubb ordb f; orda x y \<rbrakk> \<Longrightarrow> ordb (f x) (f y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 487 | by(auto simp add: mcont_def dest: monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 488 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 489 | lemma mcont_contD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 490 |   "\<lbrakk> mcont luba orda lubb ordb f; Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk>
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 491 | \<Longrightarrow> f (luba Y) = lubb (f ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 492 | by(auto simp add: mcont_def dest: contD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 493 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 494 | lemma mcont_call [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 495 | "mcont (fun_lub lub) (fun_ord ord) lub ord (\<lambda>f. f t)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 496 | by(simp add: mcont_def call_mono call_cont) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 497 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 498 | lemma mcont_id' [cont_intro, simp]: "mcont lub ord lub ord (\<lambda>x. x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 499 | by(simp add: mcont_def monotone_id') | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 500 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 501 | lemma mcont_applyI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 502 | "mcont luba orda lubb ordb (\<lambda>x. F x) \<Longrightarrow> mcont (fun_lub luba) (fun_ord orda) lubb ordb (\<lambda>f. F (f x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 503 | by(simp add: mcont_def monotone_applyI cont_applyI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 504 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 505 | lemma mcont_if [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 506 | "\<lbrakk> mcont luba orda lubb ordb (\<lambda>x. f x); mcont luba orda lubb ordb (\<lambda>x. g x) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 507 | \<Longrightarrow> mcont luba orda lubb ordb (\<lambda>x. if c then f x else g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 508 | by(simp add: mcont_def cont_if) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 509 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 510 | lemma cont_fun_lub_apply: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 511 | "cont luba orda (fun_lub lubb) (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. cont luba orda lubb ordb (\<lambda>y. f y x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 512 | by(simp add: cont_def fun_lub_def fun_eq_iff)(auto simp add: image_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 513 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 514 | lemma mcont_fun_lub_apply: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 515 | "mcont luba orda (fun_lub lubb) (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. mcont luba orda lubb ordb (\<lambda>y. f y x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 516 | by(auto simp add: monotone_fun_ord_apply cont_fun_lub_apply mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 517 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 518 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 519 | |
| 67399 | 520 | lemma cont_const [simp, cont_intro]: "cont luba orda Sup (\<le>) (\<lambda>x. c)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 521 | by (rule contI) (simp add: image_constant_conv cong del: strong_SUP_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 522 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 523 | lemma mcont_const [cont_intro, simp]: | 
| 67399 | 524 | "mcont luba orda Sup (\<le>) (\<lambda>x. c)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 525 | by(simp add: mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 526 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 527 | lemma cont_apply: | 
| 67399 | 528 | assumes 2: "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 529 | and t: "cont luba orda lubb ordb (\<lambda>x. t x)" | 
| 67399 | 530 | and 1: "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 531 | and mono: "monotone orda ordb (\<lambda>x. t x)" | 
| 67399 | 532 | and mono2: "\<And>x. monotone ordb (\<le>) (\<lambda>y. f x y)" | 
| 533 | and mono1: "\<And>y. monotone orda (\<le>) (\<lambda>x. f x y)" | |
| 534 | shows "cont luba orda Sup (\<le>) (\<lambda>x. f x (t x))" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 535 | proof | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 536 | fix Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 537 |   assume chain: "Complete_Partial_Order.chain orda Y" and "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 538 | moreover from chain have chain': "Complete_Partial_Order.chain ordb (t ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 539 | by(rule chain_imageI)(rule monotoneD[OF mono]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 540 | ultimately show "f (luba Y) (t (luba Y)) = \<Squnion>((\<lambda>x. f x (t x)) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 541 | by(simp add: contD[OF 1] contD[OF t] contD[OF 2] image_image) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 542 | (rule diag_Sup[OF chain], auto intro: monotone2monotone[OF mono2 mono monotone_const transpI] monotoneD[OF mono1]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 543 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 544 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 545 | lemma mcont2mcont': | 
| 67399 | 546 | "\<lbrakk> \<And>x. mcont lub' ord' Sup (\<le>) (\<lambda>y. f x y); | 
| 547 | \<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. f x y); | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 548 | mcont lub ord lub' ord' (\<lambda>y. t y) \<rbrakk> | 
| 67399 | 549 | \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x (t x))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 550 | unfolding mcont_def by(blast intro: transp_le monotone2monotone cont_apply) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 551 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 552 | lemma mcont2mcont: | 
| 67399 | 553 | "\<lbrakk>mcont lub' ord' Sup (\<le>) (\<lambda>x. f x); mcont lub ord lub' ord' (\<lambda>x. t x)\<rbrakk> | 
| 554 | \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f (t x))" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 555 | by(rule mcont2mcont'[OF _ mcont_const]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 556 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 557 | context | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 558 | fixes ord :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubseteq>" 60) | 
| 69039 | 559 |   and lub :: "'b set \<Rightarrow> 'b" ("\<Or>")
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 560 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 561 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 562 | lemma cont_fun_lub_Sup: | 
| 67399 | 563 | assumes chainM: "Complete_Partial_Order.chain (fun_ord (\<le>)) M" | 
| 564 | and mcont [rule_format]: "\<forall>f\<in>M. mcont lub (\<sqsubseteq>) Sup (\<le>) f" | |
| 565 | shows "cont lub (\<sqsubseteq>) Sup (\<le>) (fun_lub Sup M)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 566 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 567 | fix Y | 
| 67399 | 568 | assume chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 569 |     and Y: "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 570 | from swap_Sup[OF chain chainM mcont[THEN mcont_mono]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 571 | show "fun_lub Sup M (\<Or>Y) = \<Squnion>(fun_lub Sup M ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 572 | by(simp add: mcont_contD[OF mcont chain Y] fun_lub_apply cong: image_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 573 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 574 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 575 | lemma mcont_fun_lub_Sup: | 
| 67399 | 576 | "\<lbrakk> Complete_Partial_Order.chain (fun_ord (\<le>)) M; | 
| 577 | \<forall>f\<in>M. mcont lub ord Sup (\<le>) f \<rbrakk> | |
| 578 | \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (fun_lub Sup M)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 579 | by(simp add: mcont_def cont_fun_lub_Sup mono_lub) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 580 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 581 | lemma iterates_mcont: | 
| 67399 | 582 | assumes f: "f \<in> ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F" | 
| 583 | and mono: "\<And>f. mcont lub (\<sqsubseteq>) Sup (\<le>) f \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (F f)" | |
| 584 | shows "mcont lub (\<sqsubseteq>) Sup (\<le>) f" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 585 | using f | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 586 | by(induction rule: ccpo.iterates.induct[OF ccpo_fun, consumes 1, case_names step Sup])(blast intro: mono mcont_fun_lub_Sup)+ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 587 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 588 | lemma fixp_preserves_mcont: | 
| 67399 | 589 | assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. F f x)" | 
| 590 | and mcont: "\<And>f. mcont lub (\<sqsubseteq>) Sup (\<le>) f \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (F f)" | |
| 591 | shows "mcont lub (\<sqsubseteq>) Sup (\<le>) (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) F)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 592 | (is "mcont _ _ _ _ ?fixp") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 593 | unfolding mcont_def | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 594 | proof(intro conjI monotoneI contI) | 
| 67399 | 595 | have mono: "monotone (fun_ord (\<le>)) (fun_ord (\<le>)) F" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 596 | by(rule monotoneI)(auto simp add: fun_ord_def intro: monotoneD[OF mono]) | 
| 67399 | 597 | let ?iter = "ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F" | 
| 598 | have chain: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` ?iter)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 599 | by(rule chain_imageI[OF ccpo.chain_iterates[OF ccpo_fun mono]])(simp add: fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 600 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 601 |   {
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 602 | fix x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 603 | assume "x \<sqsubseteq> y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 604 | show "?fixp x \<le> ?fixp y" | 
| 63170 | 605 | apply (simp only: ccpo.fixp_def[OF ccpo_fun] fun_lub_apply) | 
| 606 | using chain | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 607 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 608 | fix x' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 609 | assume "x' \<in> (\<lambda>f. f x) ` ?iter" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 610 | then obtain f where "f \<in> ?iter" "x' = f x" by blast note this(2) | 
| 62837 | 611 | also from _ \<open>x \<sqsubseteq> y\<close> have "f x \<le> f y" | 
| 612 | by(rule mcont_monoD[OF iterates_mcont[OF \<open>f \<in> ?iter\<close> mcont]]) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 613 | also have "f y \<le> \<Squnion>((\<lambda>f. f y) ` ?iter)" using chain | 
| 62837 | 614 | by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> ?iter\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 615 | finally show "x' \<le> \<dots>" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 616 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 617 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 618 | fix Y | 
| 67399 | 619 | assume chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 620 |       and Y: "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 621 |     { fix f
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 622 | assume "f \<in> ?iter" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 623 | hence "f (\<Or>Y) = \<Squnion>(f ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 624 | using mcont chain Y by(rule mcont_contD[OF iterates_mcont]) } | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 625 | moreover have "\<Squnion>((\<lambda>f. \<Squnion>(f ` Y)) ` ?iter) = \<Squnion>((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` ?iter)) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 626 | using chain ccpo.chain_iterates[OF ccpo_fun mono] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 627 | by(rule swap_Sup)(rule mcont_mono[OF iterates_mcont[OF _ mcont]]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 628 | ultimately show "?fixp (\<Or>Y) = \<Squnion>(?fixp ` Y)" unfolding ccpo.fixp_def[OF ccpo_fun] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 629 | by(simp add: fun_lub_apply cong: image_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 630 | } | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 631 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 632 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 633 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 634 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 635 | context | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 636 |   fixes F :: "'c \<Rightarrow> 'c" and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c" and f
 | 
| 67399 | 637 | assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. U (F (C f)) x)" | 
| 638 | and eq: "f \<equiv> C (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) (\<lambda>f. U (F (C f))))" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 639 | and inverse: "\<And>f. U (C f) = f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 640 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 641 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 642 | lemma fixp_preserves_mono_uc: | 
| 67399 | 643 | assumes mono2: "\<And>f. monotone ord (\<le>) (U f) \<Longrightarrow> monotone ord (\<le>) (U (F f))" | 
| 644 | shows "monotone ord (\<le>) (U f)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 645 | using fixp_preserves_mono[OF mono mono2] by(subst eq)(simp add: inverse) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 646 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 647 | lemma fixp_preserves_mcont_uc: | 
| 67399 | 648 | assumes mcont: "\<And>f. mcont lubb ordb Sup (\<le>) (U f) \<Longrightarrow> mcont lubb ordb Sup (\<le>) (U (F f))" | 
| 649 | shows "mcont lubb ordb Sup (\<le>) (U f)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 650 | using fixp_preserves_mcont[OF mono mcont] by(subst eq)(simp add: inverse) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 651 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 652 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 653 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 654 | lemmas fixp_preserves_mono1 = fixp_preserves_mono_uc[of "\<lambda>x. x" _ "\<lambda>x. x", OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 655 | lemmas fixp_preserves_mono2 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 656 | fixp_preserves_mono_uc[of "case_prod" _ "curry", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 657 | lemmas fixp_preserves_mono3 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 658 | fixp_preserves_mono_uc[of "\<lambda>f. case_prod (case_prod f)" _ "\<lambda>f. curry (curry f)", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 659 | lemmas fixp_preserves_mono4 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 660 | fixp_preserves_mono_uc[of "\<lambda>f. case_prod (case_prod (case_prod f))" _ "\<lambda>f. curry (curry (curry f))", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 661 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 662 | lemmas fixp_preserves_mcont1 = fixp_preserves_mcont_uc[of "\<lambda>x. x" _ "\<lambda>x. x", OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 663 | lemmas fixp_preserves_mcont2 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 664 | fixp_preserves_mcont_uc[of "case_prod" _ "curry", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 665 | lemmas fixp_preserves_mcont3 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 666 | fixp_preserves_mcont_uc[of "\<lambda>f. case_prod (case_prod f)" _ "\<lambda>f. curry (curry f)", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 667 | lemmas fixp_preserves_mcont4 = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 668 | fixp_preserves_mcont_uc[of "\<lambda>f. case_prod (case_prod (case_prod f))" _ "\<lambda>f. curry (curry (curry f))", unfolded case_prod_curry curry_case_prod, OF _ _ refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 669 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 670 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 671 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 672 | lemma (in preorder) monotone_if_bot: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 673 | fixes bot | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 674 | assumes mono: "\<And>x y. \<lbrakk> x \<le> y; \<not> (x \<le> bound) \<rbrakk> \<Longrightarrow> ord (f x) (f y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 675 | and bot: "\<And>x. \<not> x \<le> bound \<Longrightarrow> ord bot (f x)" "ord bot bot" | 
| 67399 | 676 | shows "monotone (\<le>) ord (\<lambda>x. if x \<le> bound then bot else f x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 677 | by(rule monotoneI)(auto intro: bot intro: mono order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 678 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 679 | lemma (in ccpo) mcont_if_bot: | 
| 69039 | 680 |   fixes bot and lub ("\<Or>") and ord (infix "\<sqsubseteq>" 60)
 | 
| 67399 | 681 | assumes ccpo: "class.ccpo lub (\<sqsubseteq>) lt" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 682 | and mono: "\<And>x y. \<lbrakk> x \<le> y; \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f x \<sqsubseteq> f y" | 
| 67399 | 683 |   and cont: "\<And>Y. \<lbrakk> Complete_Partial_Order.chain (\<le>) Y; Y \<noteq> {}; \<And>x. x \<in> Y \<Longrightarrow> \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f (\<Squnion>Y) = \<Or>(f ` Y)"
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 684 | and bot: "\<And>x. \<not> x \<le> bound \<Longrightarrow> bot \<sqsubseteq> f x" | 
| 67399 | 685 | shows "mcont Sup (\<le>) lub (\<sqsubseteq>) (\<lambda>x. if x \<le> bound then bot else f x)" (is "mcont _ _ _ _ ?g") | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 686 | proof(intro mcontI contI) | 
| 67399 | 687 | interpret c: ccpo lub "(\<sqsubseteq>)" lt by(fact ccpo) | 
| 688 | show "monotone (\<le>) (\<sqsubseteq>) ?g" by(rule monotone_if_bot)(simp_all add: mono bot) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 689 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 690 | fix Y | 
| 67399 | 691 |   assume chain: "Complete_Partial_Order.chain (\<le>) Y" and Y: "Y \<noteq> {}"
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 692 | show "?g (\<Squnion>Y) = \<Or>(?g ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 693 |   proof(cases "Y \<subseteq> {x. x \<le> bound}")
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 694 | case True | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 695 | hence "\<Squnion>Y \<le> bound" using chain by(auto intro: ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 696 |     moreover have "Y \<inter> {x. \<not> x \<le> bound} = {}" using True by auto
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 697 | ultimately show ?thesis using True Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 698 | by (auto simp add: image_constant_conv cong del: c.strong_SUP_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 699 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 700 | case False | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 701 |     let ?Y = "Y \<inter> {x. \<not> x \<le> bound}"
 | 
| 67399 | 702 | have chain': "Complete_Partial_Order.chain (\<le>) ?Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 703 | using chain by(rule chain_subset) simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 704 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 705 | from False obtain y where ybound: "\<not> y \<le> bound" and y: "y \<in> Y" by blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 706 | hence "\<not> \<Squnion>Y \<le> bound" by (metis ccpo_Sup_upper chain order.trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 707 | hence "?g (\<Squnion>Y) = f (\<Squnion>Y)" by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 708 | also have "\<Squnion>Y \<le> \<Squnion>?Y" using chain | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 709 | proof(rule ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 710 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 711 | assume x: "x \<in> Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 712 | show "x \<le> \<Squnion>?Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 713 | proof(cases "x \<le> bound") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 714 | case True | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 715 | with chainD[OF chain x y] have "x \<le> y" using ybound by(auto intro: order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 716 | thus ?thesis by(rule order_trans)(auto intro: ccpo_Sup_upper[OF chain'] simp add: y ybound) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 717 | qed(auto intro: ccpo_Sup_upper[OF chain'] simp add: x) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 718 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 719 | hence "\<Squnion>Y = \<Squnion>?Y" by(rule antisym)(blast intro: ccpo_Sup_least[OF chain'] ccpo_Sup_upper[OF chain]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 720 | hence "f (\<Squnion>Y) = f (\<Squnion>?Y)" by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 721 | also have "f (\<Squnion>?Y) = \<Or>(f ` ?Y)" using chain' by(rule cont)(insert y ybound, auto) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 722 | also have "\<Or>(f ` ?Y) = \<Or>(?g ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 723 |     proof(cases "Y \<inter> {x. x \<le> bound} = {}")
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 724 | case True | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 725 | hence "f ` ?Y = ?g ` Y" by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 726 | thus ?thesis by(rule arg_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 727 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 728 | case False | 
| 67399 | 729 | have chain'': "Complete_Partial_Order.chain (\<sqsubseteq>) (insert bot (f ` ?Y))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 730 | using chain by(auto intro!: chainI bot dest: chainD intro: mono) | 
| 67399 | 731 | hence chain''': "Complete_Partial_Order.chain (\<sqsubseteq>) (f ` ?Y)" by(rule chain_subset) blast | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 732 | have "bot \<sqsubseteq> \<Or>(f ` ?Y)" using y ybound by(blast intro: c.order_trans[OF bot] c.ccpo_Sup_upper[OF chain''']) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 733 | hence "\<Or>(insert bot (f ` ?Y)) \<sqsubseteq> \<Or>(f ` ?Y)" using chain'' | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 734 | by(auto intro: c.ccpo_Sup_least c.ccpo_Sup_upper[OF chain''']) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 735 | with _ have "\<dots> = \<Or>(insert bot (f ` ?Y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 736 | by(rule c.antisym)(blast intro: c.ccpo_Sup_least[OF chain'''] c.ccpo_Sup_upper[OF chain'']) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 737 | also have "insert bot (f ` ?Y) = ?g ` Y" using False by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 738 | finally show ?thesis . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 739 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 740 | finally show ?thesis . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 741 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 742 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 743 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 744 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 745 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 746 | lemma mcont_const [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 747 | "mcont luba orda lub leq (\<lambda>x. c)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 748 | by(rule ccpo.mcont_const)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 749 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 750 | lemmas [cont_intro, simp] = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 751 | ccpo.cont_const[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 752 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 753 | lemma mono2mono: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 754 | assumes "monotone ordb leq (\<lambda>y. f y)" "monotone orda ordb (\<lambda>x. t x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 755 | shows "monotone orda leq (\<lambda>x. f (t x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 756 | using assms by(rule monotone2monotone) simp_all | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 757 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 758 | lemmas mcont2mcont' = ccpo.mcont2mcont'[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 759 | lemmas mcont2mcont = ccpo.mcont2mcont[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 760 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 761 | lemmas fixp_preserves_mono1 = ccpo.fixp_preserves_mono1[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 762 | lemmas fixp_preserves_mono2 = ccpo.fixp_preserves_mono2[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 763 | lemmas fixp_preserves_mono3 = ccpo.fixp_preserves_mono3[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 764 | lemmas fixp_preserves_mono4 = ccpo.fixp_preserves_mono4[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 765 | lemmas fixp_preserves_mcont1 = ccpo.fixp_preserves_mcont1[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 766 | lemmas fixp_preserves_mcont2 = ccpo.fixp_preserves_mcont2[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 767 | lemmas fixp_preserves_mcont3 = ccpo.fixp_preserves_mcont3[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 768 | lemmas fixp_preserves_mcont4 = ccpo.fixp_preserves_mcont4[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 769 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 770 | lemma monotone_if_bot: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 771 | fixes bot | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 772 | assumes g: "\<And>x. g x = (if leq x bound then bot else f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 773 | and mono: "\<And>x y. \<lbrakk> leq x y; \<not> leq x bound \<rbrakk> \<Longrightarrow> ord (f x) (f y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 774 | and bot: "\<And>x. \<not> leq x bound \<Longrightarrow> ord bot (f x)" "ord bot bot" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 775 | shows "monotone leq ord g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 776 | unfolding g[abs_def] using preorder mono bot by(rule preorder.monotone_if_bot) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 777 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 778 | lemma mcont_if_bot: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 779 | fixes bot | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 780 | assumes ccpo: "class.ccpo lub' ord (mk_less ord)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 781 | and bot: "\<And>x. \<not> leq x bound \<Longrightarrow> ord bot (f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 782 | and g: "\<And>x. g x = (if leq x bound then bot else f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 783 | and mono: "\<And>x y. \<lbrakk> leq x y; \<not> leq x bound \<rbrakk> \<Longrightarrow> ord (f x) (f y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 784 |   and cont: "\<And>Y. \<lbrakk> Complete_Partial_Order.chain leq Y; Y \<noteq> {}; \<And>x. x \<in> Y \<Longrightarrow> \<not> leq x bound \<rbrakk> \<Longrightarrow> f (lub Y) = lub' (f ` Y)"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 785 | shows "mcont lub leq lub' ord g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 786 | unfolding g[abs_def] using ccpo mono cont bot by(rule ccpo.mcont_if_bot[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 787 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 788 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 789 | |
| 62837 | 790 | subsection \<open>Admissibility\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 791 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 792 | lemma admissible_subst: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 793 | assumes adm: "ccpo.admissible luba orda (\<lambda>x. P x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 794 | and mcont: "mcont lubb ordb luba orda f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 795 | shows "ccpo.admissible lubb ordb (\<lambda>x. P (f x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 796 | apply(rule ccpo.admissibleI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 797 | apply(frule (1) mcont_contD[OF mcont]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 798 | apply(auto intro: ccpo.admissibleD[OF adm] chain_imageI dest: mcont_monoD[OF mcont]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 799 | done | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 800 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 801 | lemmas [simp, cont_intro] = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 802 | admissible_all | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 803 | admissible_ball | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 804 | admissible_const | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 805 | admissible_conj | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 806 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 807 | lemma admissible_disj' [simp, cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 808 | "\<lbrakk> class.ccpo lub ord (mk_less ord); ccpo.admissible lub ord P; ccpo.admissible lub ord Q \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 809 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<or> Q x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 810 | by(rule ccpo.admissible_disj) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 811 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 812 | lemma admissible_imp' [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 813 | "\<lbrakk> class.ccpo lub ord (mk_less ord); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 814 | ccpo.admissible lub ord (\<lambda>x. \<not> P x); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 815 | ccpo.admissible lub ord (\<lambda>x. Q x) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 816 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<longrightarrow> Q x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 817 | unfolding imp_conv_disj by(rule ccpo.admissible_disj) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 818 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 819 | lemma admissible_imp [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 820 | "(Q \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x)) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 821 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. Q \<longrightarrow> P x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 822 | by(rule ccpo.admissibleI)(auto dest: ccpo.admissibleD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 823 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 824 | lemma admissible_not_mem' [THEN admissible_subst, cont_intro, simp]: | 
| 67399 | 825 | shows admissible_not_mem: "ccpo.admissible Union (\<subseteq>) (\<lambda>A. x \<notin> A)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 826 | by(rule ccpo.admissibleI) auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 827 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 828 | lemma admissible_eqI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 829 | assumes f: "cont luba orda lub ord (\<lambda>x. f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 830 | and g: "cont luba orda lub ord (\<lambda>x. g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 831 | shows "ccpo.admissible luba orda (\<lambda>x. f x = g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 832 | apply(rule ccpo.admissibleI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 833 | apply(simp_all add: contD[OF f] contD[OF g] cong: image_cong) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 834 | done | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 835 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 836 | corollary admissible_eq_mcontI [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 837 | "\<lbrakk> mcont luba orda lub ord (\<lambda>x. f x); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 838 | mcont luba orda lub ord (\<lambda>x. g x) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 839 | \<Longrightarrow> ccpo.admissible luba orda (\<lambda>x. f x = g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 840 | by(rule admissible_eqI)(auto simp add: mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 841 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 842 | lemma admissible_iff [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 843 | "\<lbrakk> ccpo.admissible lub ord (\<lambda>x. P x \<longrightarrow> Q x); ccpo.admissible lub ord (\<lambda>x. Q x \<longrightarrow> P x) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 844 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<longleftrightarrow> Q x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 845 | by(subst iff_conv_conj_imp)(rule admissible_conj) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 846 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 847 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 848 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 849 | lemma admissible_leI: | 
| 67399 | 850 | assumes f: "mcont luba orda Sup (\<le>) (\<lambda>x. f x)" | 
| 851 | and g: "mcont luba orda Sup (\<le>) (\<lambda>x. g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 852 | shows "ccpo.admissible luba orda (\<lambda>x. f x \<le> g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 853 | proof(rule ccpo.admissibleI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 854 | fix A | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 855 | assume chain: "Complete_Partial_Order.chain orda A" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 856 | and le: "\<forall>x\<in>A. f x \<le> g x" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 857 |     and False: "A \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 858 | have "f (luba A) = \<Squnion>(f ` A)" by(simp add: mcont_contD[OF f] chain False) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 859 | also have "\<dots> \<le> \<Squnion>(g ` A)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 860 | proof(rule ccpo_Sup_least) | 
| 67399 | 861 | from chain show "Complete_Partial_Order.chain (\<le>) (f ` A)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 862 | by(rule chain_imageI)(rule mcont_monoD[OF f]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 863 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 864 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 865 | assume "x \<in> f ` A" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 866 | then obtain y where "y \<in> A" "x = f y" by blast note this(2) | 
| 62837 | 867 | also have "f y \<le> g y" using le \<open>y \<in> A\<close> by simp | 
| 67399 | 868 | also have "Complete_Partial_Order.chain (\<le>) (g ` A)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 869 | using chain by(rule chain_imageI)(rule mcont_monoD[OF g]) | 
| 62837 | 870 | hence "g y \<le> \<Squnion>(g ` A)" by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> A\<close>) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 871 | finally show "x \<le> \<dots>" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 872 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 873 | also have "\<dots> = g (luba A)" by(simp add: mcont_contD[OF g] chain False) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 874 | finally show "f (luba A) \<le> g (luba A)" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 875 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 876 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 877 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 878 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 879 | lemma admissible_leI: | 
| 69039 | 880 |   fixes ord (infix "\<sqsubseteq>" 60) and lub ("\<Or>")
 | 
| 67399 | 881 | assumes "class.ccpo lub (\<sqsubseteq>) (mk_less (\<sqsubseteq>))" | 
| 882 | and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. f x)" | |
| 883 | and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 884 | shows "ccpo.admissible luba orda (\<lambda>x. f x \<sqsubseteq> g x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 885 | using assms by(rule ccpo.admissible_leI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 886 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 887 | declare ccpo_class.admissible_leI[cont_intro] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 888 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 889 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 890 | |
| 67399 | 891 | lemma admissible_not_below: "ccpo.admissible Sup (\<le>) (\<lambda>x. \<not> (\<le>) x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 892 | by(rule ccpo.admissibleI)(simp add: ccpo_Sup_below_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 893 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 894 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 895 | |
| 67399 | 896 | lemma (in preorder) preorder [cont_intro, simp]: "class.preorder (\<le>) (mk_less (\<le>))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 897 | by(unfold_locales)(auto simp add: mk_less_def intro: order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 898 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 899 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 900 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 901 | lemmas [cont_intro, simp] = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 902 | admissible_leI[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 903 | ccpo.admissible_not_below[THEN admissible_subst, OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 904 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 905 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 906 | |
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 907 | setup \<open>Sign.map_naming (Name_Space.mandatory_path "ccpo")\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 908 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 909 | inductive compact :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 910 | for lub ord x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 911 | where compact: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 912 | "\<lbrakk> ccpo.admissible lub ord (\<lambda>y. \<not> ord x y); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 913 | ccpo.admissible lub ord (\<lambda>y. x \<noteq> y) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 914 | \<Longrightarrow> compact lub ord x" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 915 | |
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 916 | setup \<open>Sign.map_naming Name_Space.parent_path\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 917 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 918 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 919 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 920 | lemma compactI: | 
| 67399 | 921 | assumes "ccpo.admissible Sup (\<le>) (\<lambda>y. \<not> x \<le> y)" | 
| 922 | shows "ccpo.compact Sup (\<le>) x" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 923 | using assms | 
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 924 | proof(rule ccpo.compact.intros) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 925 | have neq: "(\<lambda>y. x \<noteq> y) = (\<lambda>y. \<not> x \<le> y \<or> \<not> y \<le> x)" by(auto) | 
| 67399 | 926 | show "ccpo.admissible Sup (\<le>) (\<lambda>y. x \<noteq> y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 927 | by(subst neq)(rule admissible_disj admissible_not_below assms)+ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 928 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 929 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 930 | lemma compact_bot: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 931 |   assumes "x = Sup {}"
 | 
| 67399 | 932 | shows "ccpo.compact Sup (\<le>) x" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 933 | proof(rule compactI) | 
| 67399 | 934 | show "ccpo.admissible Sup (\<le>) (\<lambda>y. \<not> x \<le> y)" using assms | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 935 | by(auto intro!: ccpo.admissibleI intro: ccpo_Sup_least chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 936 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 937 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 938 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 939 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 940 | lemma admissible_compact_neq' [THEN admissible_subst, cont_intro, simp]: | 
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 941 | shows admissible_compact_neq: "ccpo.compact lub ord k \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. k \<noteq> x)" | 
| 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 942 | by(simp add: ccpo.compact.simps) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 943 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 944 | lemma admissible_neq_compact' [THEN admissible_subst, cont_intro, simp]: | 
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 945 | shows admissible_neq_compact: "ccpo.compact lub ord k \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. x \<noteq> k)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 946 | by(subst eq_commute)(rule admissible_compact_neq) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 947 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 948 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 949 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 950 | lemmas [cont_intro, simp] = ccpo.compact_bot[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 951 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 952 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 953 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 954 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 955 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 956 | lemma fixp_strong_induct: | 
| 67399 | 957 | assumes [cont_intro]: "ccpo.admissible Sup (\<le>) P" | 
| 958 | and mono: "monotone (\<le>) (\<le>) f" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 959 |   and bot: "P (\<Squnion>{})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 960 | and step: "\<And>x. \<lbrakk> x \<le> ccpo_class.fixp f; P x \<rbrakk> \<Longrightarrow> P (f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 961 | shows "P (ccpo_class.fixp f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 962 | proof(rule fixp_induct[where P="\<lambda>x. x \<le> ccpo_class.fixp f \<and> P x", THEN conjunct2]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 963 | note [cont_intro] = admissible_leI | 
| 67399 | 964 | show "ccpo.admissible Sup (\<le>) (\<lambda>x. x \<le> ccpo_class.fixp f \<and> P x)" by simp | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 965 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 966 |   show "\<Squnion>{} \<le> ccpo_class.fixp f \<and> P (\<Squnion>{})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 967 | by(auto simp add: bot intro: ccpo_Sup_least chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 968 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 969 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 970 | assume "x \<le> ccpo_class.fixp f \<and> P x" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 971 | thus "f x \<le> ccpo_class.fixp f \<and> P (f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 972 | by(subst fixp_unfold[OF mono])(auto dest: monotoneD[OF mono] intro: step) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 973 | qed(rule mono) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 974 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 975 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 976 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 977 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 978 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 979 | lemma fixp_strong_induct_uc: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 980 | fixes F :: "'c \<Rightarrow> 'c" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 981 | and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 982 |     and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 983 |     and P :: "('b \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 984 | assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 985 | and eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 986 | and inverse: "\<And>f. U (C f) = f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 987 | and adm: "ccpo.admissible lub_fun le_fun P" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 988 |     and bot: "P (\<lambda>_. lub {})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 989 | and step: "\<And>f'. \<lbrakk> P (U f'); le_fun (U f') (U f) \<rbrakk> \<Longrightarrow> P (U (F f'))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 990 | shows "P (U f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 991 | unfolding eq inverse | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 992 | apply (rule ccpo.fixp_strong_induct[OF ccpo adm]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 993 | apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 994 | apply (rule_tac f'5="C x" in step) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 995 | apply (simp_all add: inverse eq) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 996 | done | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 997 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 998 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 999 | |
| 67399 | 1000 | subsection \<open>@{term "(=)"} as order\<close>
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1001 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1002 | definition lub_singleton :: "('a set \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1003 | where "lub_singleton lub \<longleftrightarrow> (\<forall>a. lub {a} = a)"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1004 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1005 | definition the_Sup :: "'a set \<Rightarrow> 'a" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1006 | where "the_Sup A = (THE a. a \<in> A)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1007 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1008 | lemma lub_singleton_the_Sup [cont_intro, simp]: "lub_singleton the_Sup" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1009 | by(simp add: lub_singleton_def the_Sup_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1010 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1011 | lemma (in ccpo) lub_singleton: "lub_singleton Sup" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1012 | by(simp add: lub_singleton_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1013 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1014 | lemma (in partial_function_definitions) lub_singleton [cont_intro, simp]: "lub_singleton lub" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1015 | by(rule ccpo.lub_singleton)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1016 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1017 | lemma preorder_eq [cont_intro, simp]: | 
| 67399 | 1018 | "class.preorder (=) (mk_less (=))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1019 | by(unfold_locales)(simp_all add: mk_less_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1020 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1021 | lemma monotone_eqI [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1022 | assumes "class.preorder ord (mk_less ord)" | 
| 67399 | 1023 | shows "monotone (=) ord f" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1024 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1025 | interpret preorder ord "mk_less ord" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1026 | show ?thesis by(simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1027 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1028 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1029 | lemma cont_eqI [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1030 | fixes f :: "'a \<Rightarrow> 'b" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1031 | assumes "lub_singleton lub" | 
| 67399 | 1032 | shows "cont the_Sup (=) lub ord f" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1033 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1034 | fix Y :: "'a set" | 
| 67399 | 1035 |   assume "Complete_Partial_Order.chain (=) Y" "Y \<noteq> {}"
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1036 |   then obtain a where "Y = {a}" by(auto simp add: chain_def)
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1037 | thus "f (the_Sup Y) = lub (f ` Y)" using assms | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1038 | by(simp add: the_Sup_def lub_singleton_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1039 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1040 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1041 | lemma mcont_eqI [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1042 | "\<lbrakk> class.preorder ord (mk_less ord); lub_singleton lub \<rbrakk> | 
| 67399 | 1043 | \<Longrightarrow> mcont the_Sup (=) lub ord f" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1044 | by(simp add: mcont_def cont_eqI monotone_eqI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1045 | |
| 62837 | 1046 | subsection \<open>ccpo for products\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1047 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1048 | definition prod_lub :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) set \<Rightarrow> 'a \<times> 'b"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1049 | where "prod_lub Sup_a Sup_b Y = (Sup_a (fst ` Y), Sup_b (snd ` Y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1050 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1051 | lemma lub_singleton_prod_lub [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1052 | "\<lbrakk> lub_singleton luba; lub_singleton lubb \<rbrakk> \<Longrightarrow> lub_singleton (prod_lub luba lubb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1053 | by(simp add: lub_singleton_def prod_lub_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1054 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1055 | lemma prod_lub_empty [simp]: "prod_lub luba lubb {} = (luba {}, lubb {})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1056 | by(simp add: prod_lub_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1057 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1058 | lemma preorder_rel_prodI [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1059 | assumes "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1060 | and "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1061 | shows "class.preorder (rel_prod orda ordb) (mk_less (rel_prod orda ordb))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1062 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1063 | interpret a: preorder orda "mk_less orda" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1064 | interpret b: preorder ordb "mk_less ordb" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1065 | show ?thesis by(unfold_locales)(auto simp add: mk_less_def intro: a.order_trans b.order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1066 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1067 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1068 | lemma order_rel_prodI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1069 | assumes a: "class.order orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1070 | and b: "class.order ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1071 | shows "class.order (rel_prod orda ordb) (mk_less (rel_prod orda ordb))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1072 | (is "class.order ?ord ?ord'") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1073 | proof(intro class.order.intro class.order_axioms.intro) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1074 | interpret a: order orda "mk_less orda" by(fact a) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1075 | interpret b: order ordb "mk_less ordb" by(fact b) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1076 | show "class.preorder ?ord ?ord'" by(rule preorder_rel_prodI) unfold_locales | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1077 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1078 | fix x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1079 | assume "?ord x y" "?ord y x" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1080 | thus "x = y" by(cases x y rule: prod.exhaust[case_product prod.exhaust]) auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1081 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1082 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1083 | lemma monotone_rel_prodI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1084 | assumes mono2: "\<And>a. monotone ordb ordc (\<lambda>b. f (a, b))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1085 | and mono1: "\<And>b. monotone orda ordc (\<lambda>a. f (a, b))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1086 | and a: "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1087 | and b: "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1088 | and c: "class.preorder ordc (mk_less ordc)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1089 | shows "monotone (rel_prod orda ordb) ordc f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1090 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1091 | interpret a: preorder orda "mk_less orda" by(rule a) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1092 | interpret b: preorder ordb "mk_less ordb" by(rule b) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1093 | interpret c: preorder ordc "mk_less ordc" by(rule c) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1094 | show ?thesis using mono2 mono1 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1095 | by(auto 7 2 simp add: monotone_def intro: c.order_trans) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1096 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1097 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1098 | lemma monotone_rel_prodD1: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1099 | assumes mono: "monotone (rel_prod orda ordb) ordc f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1100 | and preorder: "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1101 | shows "monotone orda ordc (\<lambda>a. f (a, b))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1102 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1103 | interpret preorder ordb "mk_less ordb" by(rule preorder) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1104 | show ?thesis using mono by(simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1105 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1106 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1107 | lemma monotone_rel_prodD2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1108 | assumes mono: "monotone (rel_prod orda ordb) ordc f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1109 | and preorder: "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1110 | shows "monotone ordb ordc (\<lambda>b. f (a, b))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1111 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1112 | interpret preorder orda "mk_less orda" by(rule preorder) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1113 | show ?thesis using mono by(simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1114 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1115 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1116 | lemma monotone_case_prodI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1117 | "\<lbrakk> \<And>a. monotone ordb ordc (f a); \<And>b. monotone orda ordc (\<lambda>a. f a b); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1118 | class.preorder orda (mk_less orda); class.preorder ordb (mk_less ordb); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1119 | class.preorder ordc (mk_less ordc) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1120 | \<Longrightarrow> monotone (rel_prod orda ordb) ordc (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1121 | by(rule monotone_rel_prodI) simp_all | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1122 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1123 | lemma monotone_case_prodD1: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1124 | assumes mono: "monotone (rel_prod orda ordb) ordc (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1125 | and preorder: "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1126 | shows "monotone orda ordc (\<lambda>a. f a b)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1127 | using monotone_rel_prodD1[OF assms] by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1128 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1129 | lemma monotone_case_prodD2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1130 | assumes mono: "monotone (rel_prod orda ordb) ordc (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1131 | and preorder: "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1132 | shows "monotone ordb ordc (f a)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1133 | using monotone_rel_prodD2[OF assms] by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1134 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1135 | context | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1136 | fixes orda ordb ordc | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1137 | assumes a: "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1138 | and b: "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1139 | and c: "class.preorder ordc (mk_less ordc)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1140 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1141 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1142 | lemma monotone_rel_prod_iff: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1143 | "monotone (rel_prod orda ordb) ordc f \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1144 | (\<forall>a. monotone ordb ordc (\<lambda>b. f (a, b))) \<and> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1145 | (\<forall>b. monotone orda ordc (\<lambda>a. f (a, b)))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1146 | using a b c by(blast intro: monotone_rel_prodI dest: monotone_rel_prodD1 monotone_rel_prodD2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1147 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1148 | lemma monotone_case_prod_iff [simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1149 | "monotone (rel_prod orda ordb) ordc (case_prod f) \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1150 | (\<forall>a. monotone ordb ordc (f a)) \<and> (\<forall>b. monotone orda ordc (\<lambda>a. f a b))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1151 | by(simp add: monotone_rel_prod_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1152 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1153 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1154 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1155 | lemma monotone_case_prod_apply_iff: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1156 | "monotone orda ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> monotone orda ordb (case_prod (\<lambda>a b. f a b y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1157 | by(simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1158 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1159 | lemma monotone_case_prod_applyD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1160 | "monotone orda ordb (\<lambda>x. (case_prod f x) y) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1161 | \<Longrightarrow> monotone orda ordb (case_prod (\<lambda>a b. f a b y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1162 | by(simp add: monotone_case_prod_apply_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1163 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1164 | lemma monotone_case_prod_applyI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1165 | "monotone orda ordb (case_prod (\<lambda>a b. f a b y)) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1166 | \<Longrightarrow> monotone orda ordb (\<lambda>x. (case_prod f x) y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1167 | by(simp add: monotone_case_prod_apply_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1168 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1169 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1170 | lemma cont_case_prod_apply_iff: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1171 | "cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1172 | by(simp add: cont_def split_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1173 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1174 | lemma cont_case_prod_applyI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1175 | "cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y)) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1176 | \<Longrightarrow> cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1177 | by(simp add: cont_case_prod_apply_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1178 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1179 | lemma cont_case_prod_applyD: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1180 | "cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1181 | \<Longrightarrow> cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1182 | by(simp add: cont_case_prod_apply_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1183 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1184 | lemma mcont_case_prod_apply_iff [simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1185 | "mcont luba orda lubb ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1186 | mcont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1187 | by(simp add: mcont_def monotone_case_prod_apply_iff cont_case_prod_apply_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1188 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1189 | lemma cont_prodD1: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1190 | assumes cont: "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1191 | and "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1192 | and luba: "lub_singleton luba" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1193 | shows "cont lubb ordb lubc ordc (\<lambda>y. f (x, y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1194 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1195 | interpret preorder orda "mk_less orda" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1196 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1197 | fix Y :: "'b set" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1198 |   let ?Y = "{x} \<times> Y"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1199 |   assume "Complete_Partial_Order.chain ordb Y" "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1200 |   hence "Complete_Partial_Order.chain (rel_prod orda ordb) ?Y" "?Y \<noteq> {}" 
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1201 | by(simp_all add: chain_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1202 | with cont have "f (prod_lub luba lubb ?Y) = lubc (f ` ?Y)" by(rule contD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1203 | moreover have "f ` ?Y = (\<lambda>y. f (x, y)) ` Y" by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1204 | ultimately show "f (x, lubb Y) = lubc ((\<lambda>y. f (x, y)) ` Y)" using luba | 
| 62837 | 1205 |     by(simp add: prod_lub_def \<open>Y \<noteq> {}\<close> lub_singleton_def)
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1206 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1207 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1208 | lemma cont_prodD2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1209 | assumes cont: "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1210 | and "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1211 | and lubb: "lub_singleton lubb" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1212 | shows "cont luba orda lubc ordc (\<lambda>x. f (x, y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1213 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1214 | interpret preorder ordb "mk_less ordb" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1215 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1216 | fix Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1217 |   assume Y: "Complete_Partial_Order.chain orda Y" "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1218 |   let ?Y = "Y \<times> {y}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1219 | have "f (luba Y, y) = f (prod_lub luba lubb ?Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1220 | using lubb by(simp add: prod_lub_def Y lub_singleton_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1221 |   also from Y have "Complete_Partial_Order.chain (rel_prod orda ordb) ?Y" "?Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1222 | by(simp_all add: chain_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1223 | with cont have "f (prod_lub luba lubb ?Y) = lubc (f ` ?Y)" by(rule contD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1224 | also have "f ` ?Y = (\<lambda>x. f (x, y)) ` Y" by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1225 | finally show "f (luba Y, y) = lubc \<dots>" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1226 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1227 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1228 | lemma cont_case_prodD1: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1229 | assumes "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1230 | and "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1231 | and "lub_singleton luba" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1232 | shows "cont lubb ordb lubc ordc (f x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1233 | using cont_prodD1[OF assms] by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1234 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1235 | lemma cont_case_prodD2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1236 | assumes "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1237 | and "class.preorder ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1238 | and "lub_singleton lubb" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1239 | shows "cont luba orda lubc ordc (\<lambda>x. f x y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1240 | using cont_prodD2[OF assms] by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1241 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1242 | context ccpo begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1243 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1244 | lemma cont_prodI: | 
| 67399 | 1245 | assumes mono: "monotone (rel_prod orda ordb) (\<le>) f" | 
| 1246 | and cont1: "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f (x, y))" | |
| 1247 | and cont2: "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f (x, y))" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1248 | and "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1249 | and "class.preorder ordb (mk_less ordb)" | 
| 67399 | 1250 | shows "cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) f" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1251 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1252 | interpret a: preorder orda "mk_less orda" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1253 | interpret b: preorder ordb "mk_less ordb" by fact | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1254 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1255 | fix Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1256 | assume chain: "Complete_Partial_Order.chain (rel_prod orda ordb) Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1257 |     and "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1258 | have "f (prod_lub luba lubb Y) = f (luba (fst ` Y), lubb (snd ` Y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1259 | by(simp add: prod_lub_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1260 | also from cont2 have "f (luba (fst ` Y), lubb (snd ` Y)) = \<Squnion>((\<lambda>x. f (x, lubb (snd ` Y))) ` fst ` Y)" | 
| 62837 | 1261 |     by(rule contD)(simp_all add: chain_rel_prodD1[OF chain] \<open>Y \<noteq> {}\<close>)
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1262 | also from cont1 have "\<And>x. f (x, lubb (snd ` Y)) = \<Squnion>((\<lambda>y. f (x, y)) ` snd ` Y)" | 
| 62837 | 1263 |     by(rule contD)(simp_all add: chain_rel_prodD2[OF chain] \<open>Y \<noteq> {}\<close>)
 | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1264 | hence "\<Squnion>((\<lambda>x. f (x, lubb (snd ` Y))) ` fst ` Y) = \<Squnion>((\<lambda>x. \<dots> x) ` fst ` Y)" by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1265 | also have "\<dots> = \<Squnion>((\<lambda>x. f (fst x, snd x)) ` Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1266 | unfolding image_image split_def using chain | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1267 | apply(rule diag_Sup) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1268 | using monotoneD[OF mono] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1269 | by(auto intro: monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1270 | finally show "f (prod_lub luba lubb Y) = \<Squnion>(f ` Y)" by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1271 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1272 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1273 | lemma cont_case_prodI: | 
| 67399 | 1274 | assumes "monotone (rel_prod orda ordb) (\<le>) (case_prod f)" | 
| 1275 | and "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)" | |
| 1276 | and "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1277 | and "class.preorder orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1278 | and "class.preorder ordb (mk_less ordb)" | 
| 67399 | 1279 | shows "cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) (case_prod f)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1280 | by(rule cont_prodI)(simp_all add: assms) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1281 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1282 | lemma cont_case_prod_iff: | 
| 67399 | 1283 | "\<lbrakk> monotone (rel_prod orda ordb) (\<le>) (case_prod f); | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1284 | class.preorder orda (mk_less orda); lub_singleton luba; | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1285 | class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk> | 
| 67399 | 1286 | \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) (case_prod f) \<longleftrightarrow> | 
| 1287 | (\<forall>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)) \<and> (\<forall>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y))" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1288 | by(blast dest: cont_case_prodD1 cont_case_prodD2 intro: cont_case_prodI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1289 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1290 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1291 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1292 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1293 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1294 | lemma mono2mono2: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1295 | assumes f: "monotone (rel_prod ordb ordc) leq (\<lambda>(x, y). f x y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1296 | and t: "monotone orda ordb (\<lambda>x. t x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1297 | and t': "monotone orda ordc (\<lambda>x. t' x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1298 | shows "monotone orda leq (\<lambda>x. f (t x) (t' x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1299 | proof(rule monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1300 | fix x y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1301 | assume "orda x y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1302 | hence "rel_prod ordb ordc (t x, t' x) (t y, t' y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1303 | using t t' by(auto dest: monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1304 | from monotoneD[OF f this] show "leq (f (t x) (t' x)) (f (t y) (t' y))" by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1305 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1306 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1307 | lemma cont_case_prodI [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1308 | "\<lbrakk> monotone (rel_prod orda ordb) leq (case_prod f); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1309 | \<And>x. cont lubb ordb lub leq (\<lambda>y. f x y); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1310 | \<And>y. cont luba orda lub leq (\<lambda>x. f x y); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1311 | class.preorder orda (mk_less orda); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1312 | class.preorder ordb (mk_less ordb) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1313 | \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1314 | by(rule ccpo.cont_case_prodI)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1315 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1316 | lemma cont_case_prod_iff: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1317 | "\<lbrakk> monotone (rel_prod orda ordb) leq (case_prod f); | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1318 | class.preorder orda (mk_less orda); lub_singleton luba; | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1319 | class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1320 | \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f) \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1321 | (\<forall>x. cont lubb ordb lub leq (\<lambda>y. f x y)) \<and> (\<forall>y. cont luba orda lub leq (\<lambda>x. f x y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1322 | by(blast dest: cont_case_prodD1 cont_case_prodD2 intro: cont_case_prodI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1323 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1324 | lemma mcont_case_prod_iff [simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1325 | "\<lbrakk> class.preorder orda (mk_less orda); lub_singleton luba; | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1326 | class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1327 | \<Longrightarrow> mcont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f) \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1328 | (\<forall>x. mcont lubb ordb lub leq (\<lambda>y. f x y)) \<and> (\<forall>y. mcont luba orda lub leq (\<lambda>x. f x y))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1329 | unfolding mcont_def by(auto simp add: cont_case_prod_iff) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1330 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1331 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1332 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1333 | lemma mono2mono_case_prod [cont_intro]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1334 | assumes "\<And>x y. monotone orda ordb (\<lambda>f. pair f x y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1335 | shows "monotone orda ordb (\<lambda>f. case_prod (pair f) x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1336 | by(rule monotoneI)(auto split: prod.split dest: monotoneD[OF assms]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1337 | |
| 62837 | 1338 | subsection \<open>Complete lattices as ccpo\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1339 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1340 | context complete_lattice begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1341 | |
| 67399 | 1342 | lemma complete_lattice_ccpo: "class.ccpo Sup (\<le>) (<)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1343 | by(unfold_locales)(fast intro: Sup_upper Sup_least)+ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1344 | |
| 67399 | 1345 | lemma complete_lattice_ccpo': "class.ccpo Sup (\<le>) (mk_less (\<le>))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1346 | by(unfold_locales)(auto simp add: mk_less_def intro: Sup_upper Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1347 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1348 | lemma complete_lattice_partial_function_definitions: | 
| 67399 | 1349 | "partial_function_definitions (\<le>) Sup" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1350 | by(unfold_locales)(auto intro: Sup_least Sup_upper) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1351 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1352 | lemma complete_lattice_partial_function_definitions_dual: | 
| 67399 | 1353 | "partial_function_definitions (\<ge>) Inf" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1354 | by(unfold_locales)(auto intro: Inf_lower Inf_greatest) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1355 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1356 | lemmas [cont_intro, simp] = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1357 | Partial_Function.ccpo[OF complete_lattice_partial_function_definitions] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1358 | Partial_Function.ccpo[OF complete_lattice_partial_function_definitions_dual] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1359 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1360 | lemma mono2mono_inf: | 
| 67399 | 1361 | assumes f: "monotone ord (\<le>) (\<lambda>x. f x)" | 
| 1362 | and g: "monotone ord (\<le>) (\<lambda>x. g x)" | |
| 1363 | shows "monotone ord (\<le>) (\<lambda>x. f x \<sqinter> g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1364 | by(auto 4 3 dest: monotoneD[OF f] monotoneD[OF g] intro: le_infI1 le_infI2 intro!: monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1365 | |
| 67399 | 1366 | lemma mcont_const [simp]: "mcont lub ord Sup (\<le>) (\<lambda>_. c)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1367 | by(rule ccpo.mcont_const[OF complete_lattice_ccpo]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1368 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1369 | lemma mono2mono_sup: | 
| 67399 | 1370 | assumes f: "monotone ord (\<le>) (\<lambda>x. f x)" | 
| 1371 | and g: "monotone ord (\<le>) (\<lambda>x. g x)" | |
| 1372 | shows "monotone ord (\<le>) (\<lambda>x. f x \<squnion> g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1373 | by(auto 4 3 intro!: monotoneI intro: sup.coboundedI1 sup.coboundedI2 dest: monotoneD[OF f] monotoneD[OF g]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1374 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1375 | lemma Sup_image_sup: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1376 |   assumes "Y \<noteq> {}"
 | 
| 67399 | 1377 | shows "\<Squnion>((\<squnion>) x ` Y) = x \<squnion> \<Squnion>Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1378 | proof(rule Sup_eqI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1379 | fix y | 
| 67399 | 1380 | assume "y \<in> (\<squnion>) x ` Y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1381 | then obtain z where "y = x \<squnion> z" and "z \<in> Y" by blast | 
| 62837 | 1382 | from \<open>z \<in> Y\<close> have "z \<le> \<Squnion>Y" by(rule Sup_upper) | 
| 1383 | with _ show "y \<le> x \<squnion> \<Squnion>Y" unfolding \<open>y = x \<squnion> z\<close> by(rule sup_mono) simp | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1384 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1385 | fix y | 
| 67399 | 1386 | assume upper: "\<And>z. z \<in> (\<squnion>) x ` Y \<Longrightarrow> z \<le> y" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1387 | show "x \<squnion> \<Squnion>Y \<le> y" unfolding Sup_insert[symmetric] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1388 | proof(rule Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1389 | fix z | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1390 | assume "z \<in> insert x Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1391 | from assms obtain z' where "z' \<in> Y" by blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1392 | let ?z = "if z \<in> Y then x \<squnion> z else x \<squnion> z'" | 
| 62837 | 1393 | have "z \<le> x \<squnion> ?z" using \<open>z' \<in> Y\<close> \<open>z \<in> insert x Y\<close> by auto | 
| 1394 | also have "\<dots> \<le> y" by(rule upper)(auto split: if_split_asm intro: \<open>z' \<in> Y\<close>) | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1395 | finally show "z \<le> y" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1396 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1397 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1398 | |
| 67399 | 1399 | lemma mcont_sup1: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>y. x \<squnion> y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1400 | by(auto 4 3 simp add: mcont_def sup.coboundedI1 sup.coboundedI2 intro!: monotoneI contI intro: Sup_image_sup[symmetric]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1401 | |
| 67399 | 1402 | lemma mcont_sup2: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>x. x \<squnion> y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1403 | by(subst sup_commute)(rule mcont_sup1) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1404 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1405 | lemma mcont2mcont_sup [cont_intro, simp]: | 
| 67399 | 1406 | "\<lbrakk> mcont lub ord Sup (\<le>) (\<lambda>x. f x); | 
| 1407 | mcont lub ord Sup (\<le>) (\<lambda>x. g x) \<rbrakk> | |
| 1408 | \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x \<squnion> g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1409 | by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_sup1 mcont_sup2 ccpo.mcont_const[OF complete_lattice_ccpo]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1410 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1411 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1412 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1413 | lemmas [cont_intro] = admissible_leI[OF complete_lattice_ccpo'] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1414 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1415 | context complete_distrib_lattice begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1416 | |
| 67399 | 1417 | lemma mcont_inf1: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>y. x \<sqinter> y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1418 | by(auto intro: monotoneI contI simp add: le_infI2 inf_Sup mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1419 | |
| 67399 | 1420 | lemma mcont_inf2: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>x. x \<sqinter> y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1421 | by(auto intro: monotoneI contI simp add: le_infI1 Sup_inf mcont_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1422 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1423 | lemma mcont2mcont_inf [cont_intro, simp]: | 
| 67399 | 1424 | "\<lbrakk> mcont lub ord Sup (\<le>) (\<lambda>x. f x); | 
| 1425 | mcont lub ord Sup (\<le>) (\<lambda>x. g x) \<rbrakk> | |
| 1426 | \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x \<sqinter> g x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1427 | by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_inf1 mcont_inf2 ccpo.mcont_const[OF complete_lattice_ccpo]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1428 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1429 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1430 | |
| 67399 | 1431 | interpretation lfp: partial_function_definitions "(\<le>) :: _ :: complete_lattice \<Rightarrow> _" Sup | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1432 | by(rule complete_lattice_partial_function_definitions) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1433 | |
| 62837 | 1434 | declaration \<open>Partial_Function.init "lfp" @{term lfp.fixp_fun} @{term lfp.mono_body}
 | 
| 1435 |   @{thm lfp.fixp_rule_uc} @{thm lfp.fixp_induct_uc} NONE\<close>
 | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1436 | |
| 67399 | 1437 | interpretation gfp: partial_function_definitions "(\<ge>) :: _ :: complete_lattice \<Rightarrow> _" Inf | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1438 | by(rule complete_lattice_partial_function_definitions_dual) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1439 | |
| 62837 | 1440 | declaration \<open>Partial_Function.init "gfp" @{term gfp.fixp_fun} @{term gfp.mono_body}
 | 
| 1441 |   @{thm gfp.fixp_rule_uc} @{thm gfp.fixp_induct_uc} NONE\<close>
 | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1442 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1443 | lemma insert_mono [partial_function_mono]: | 
| 67399 | 1444 | "monotone (fun_ord (\<subseteq>)) (\<subseteq>) A \<Longrightarrow> monotone (fun_ord (\<subseteq>)) (\<subseteq>) (\<lambda>y. insert x (A y))" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1445 | by(rule monotoneI)(auto simp add: fun_ord_def dest: monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1446 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1447 | lemma mono2mono_insert [THEN lfp.mono2mono, cont_intro, simp]: | 
| 67399 | 1448 | shows monotone_insert: "monotone (\<subseteq>) (\<subseteq>) (insert x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1449 | by(rule monotoneI) blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1450 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1451 | lemma mcont2mcont_insert[THEN lfp.mcont2mcont, cont_intro, simp]: | 
| 67399 | 1452 | shows mcont_insert: "mcont Union (\<subseteq>) Union (\<subseteq>) (insert x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1453 | by(blast intro: mcontI contI monotone_insert) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1454 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1455 | lemma mono2mono_image [THEN lfp.mono2mono, cont_intro, simp]: | 
| 67399 | 1456 | shows monotone_image: "monotone (\<subseteq>) (\<subseteq>) ((`) f)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1457 | by(rule monotoneI) blast | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1458 | |
| 67399 | 1459 | lemma cont_image: "cont Union (\<subseteq>) Union (\<subseteq>) ((`) f)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1460 | by(rule contI)(auto) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1461 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1462 | lemma mcont2mcont_image [THEN lfp.mcont2mcont, cont_intro, simp]: | 
| 67399 | 1463 | shows mcont_image: "mcont Union (\<subseteq>) Union (\<subseteq>) ((`) f)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1464 | by(blast intro: mcontI monotone_image cont_image) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1465 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1466 | context complete_lattice begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1467 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1468 | lemma monotone_Sup [cont_intro, simp]: | 
| 67399 | 1469 | "monotone ord (\<subseteq>) f \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>f x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1470 | by(blast intro: monotoneI Sup_least Sup_upper dest: monotoneD) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1471 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1472 | lemma cont_Sup: | 
| 67399 | 1473 | assumes "cont lub ord Union (\<subseteq>) f" | 
| 1474 | shows "cont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>f x)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1475 | apply(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1476 | apply(simp add: contD[OF assms]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1477 | apply(blast intro: Sup_least Sup_upper order_trans antisym) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1478 | done | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1479 | |
| 67399 | 1480 | lemma mcont_Sup: "mcont lub ord Union (\<subseteq>) f \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>f x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1481 | unfolding mcont_def by(blast intro: monotone_Sup cont_Sup) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1482 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1483 | lemma monotone_SUP: | 
| 67399 | 1484 | "\<lbrakk> monotone ord (\<subseteq>) f; \<And>y. monotone ord (\<le>) (\<lambda>x. g x y) \<rbrakk> \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1485 | by(rule monotoneI)(blast dest: monotoneD intro: Sup_upper order_trans intro!: Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1486 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1487 | lemma monotone_SUP2: | 
| 67399 | 1488 | "(\<And>y. y \<in> A \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. g x y)) \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>y\<in>A. g x y)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1489 | by(rule monotoneI)(blast intro: Sup_upper order_trans dest: monotoneD intro!: Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1490 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1491 | lemma cont_SUP: | 
| 67399 | 1492 | assumes f: "mcont lub ord Union (\<subseteq>) f" | 
| 1493 | and g: "\<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. g x y)" | |
| 1494 | shows "cont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)" | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1495 | proof(rule contI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1496 | fix Y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1497 | assume chain: "Complete_Partial_Order.chain ord Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1498 |     and Y: "Y \<noteq> {}"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1499 | show "\<Squnion>(g (lub Y) ` f (lub Y)) = \<Squnion>((\<lambda>x. \<Squnion>(g x ` f x)) ` Y)" (is "?lhs = ?rhs") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1500 | proof(rule antisym) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1501 | show "?lhs \<le> ?rhs" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1502 | proof(rule Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1503 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1504 | assume "x \<in> g (lub Y) ` f (lub Y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1505 | with mcont_contD[OF f chain Y] mcont_contD[OF g chain Y] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1506 | obtain y z where "y \<in> Y" "z \<in> f y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1507 | and x: "x = \<Squnion>((\<lambda>x. g x z) ` Y)" by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1508 | show "x \<le> ?rhs" unfolding x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1509 | proof(rule Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1510 | fix u | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1511 | assume "u \<in> (\<lambda>x. g x z) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1512 | then obtain y' where "u = g y' z" "y' \<in> Y" by auto | 
| 62837 | 1513 | from chain \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "ord y y' \<or> ord y' y" by(rule chainD) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1514 | thus "u \<le> ?rhs" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1515 | proof | 
| 62837 | 1516 | note \<open>u = g y' z\<close> also | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1517 | assume "ord y y'" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1518 | with f have "f y \<subseteq> f y'" by(rule mcont_monoD) | 
| 62837 | 1519 | with \<open>z \<in> f y\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1520 | have "g y' z \<le> \<Squnion>(g y' ` f y')" by(auto intro: Sup_upper) | 
| 62837 | 1521 | also have "\<dots> \<le> ?rhs" using \<open>y' \<in> Y\<close> by(auto intro: Sup_upper) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1522 | finally show ?thesis . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1523 | next | 
| 62837 | 1524 | note \<open>u = g y' z\<close> also | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1525 | assume "ord y' y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1526 | with g have "g y' z \<le> g y z" by(rule mcont_monoD) | 
| 62837 | 1527 | also have "\<dots> \<le> \<Squnion>(g y ` f y)" using \<open>z \<in> f y\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1528 | by(auto intro: Sup_upper) | 
| 62837 | 1529 | also have "\<dots> \<le> ?rhs" using \<open>y \<in> Y\<close> by(auto intro: Sup_upper) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1530 | finally show ?thesis . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1531 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1532 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1533 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1534 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1535 | show "?rhs \<le> ?lhs" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1536 | proof(rule Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1537 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1538 | assume "x \<in> (\<lambda>x. \<Squnion>(g x ` f x)) ` Y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1539 | then obtain y where x: "x = \<Squnion>(g y ` f y)" and "y \<in> Y" by auto | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1540 | show "x \<le> ?lhs" unfolding x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1541 | proof(rule Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1542 | fix u | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1543 | assume "u \<in> g y ` f y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1544 | then obtain z where "u = g y z" "z \<in> f y" by auto | 
| 62837 | 1545 | note \<open>u = g y z\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1546 | also have "g y z \<le> \<Squnion>((\<lambda>x. g x z) ` Y)" | 
| 62837 | 1547 | using \<open>y \<in> Y\<close> by(auto intro: Sup_upper) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1548 | also have "\<dots> = g (lub Y) z" by(simp add: mcont_contD[OF g chain Y]) | 
| 62837 | 1549 | also have "\<dots> \<le> ?lhs" using \<open>z \<in> f y\<close> \<open>y \<in> Y\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1550 | by(auto intro: Sup_upper simp add: mcont_contD[OF f chain Y]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1551 | finally show "u \<le> ?lhs" . | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1552 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1553 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1554 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1555 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1556 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1557 | lemma mcont_SUP [cont_intro, simp]: | 
| 67399 | 1558 | "\<lbrakk> mcont lub ord Union (\<subseteq>) f; \<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. g x y) \<rbrakk> | 
| 1559 | \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)" | |
| 63092 | 1560 | by(blast intro: mcontI cont_SUP monotone_SUP mcont_mono) | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1561 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1562 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1563 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1564 | lemma admissible_Ball [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1565 | "\<lbrakk> \<And>x. ccpo.admissible lub ord (\<lambda>A. P A x); | 
| 67399 | 1566 | mcont lub ord Union (\<subseteq>) f; | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1567 | class.ccpo lub ord (mk_less ord) \<rbrakk> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1568 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>A. \<forall>x\<in>f A. P A x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1569 | unfolding Ball_def by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1570 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1571 | lemma admissible_Bex'[THEN admissible_subst, cont_intro, simp]: | 
| 67399 | 1572 | shows admissible_Bex: "ccpo.admissible Union (\<subseteq>) (\<lambda>A. \<exists>x\<in>A. P x)" | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1573 | by(rule ccpo.admissibleI)(auto) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1574 | |
| 62837 | 1575 | subsection \<open>Parallel fixpoint induction\<close> | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1576 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1577 | context | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1578 | fixes luba :: "'a set \<Rightarrow> 'a" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1579 | and orda :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1580 | and lubb :: "'b set \<Rightarrow> 'b" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1581 | and ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1582 | assumes a: "class.ccpo luba orda (mk_less orda)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1583 | and b: "class.ccpo lubb ordb (mk_less ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1584 | begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1585 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1586 | interpretation a: ccpo luba orda "mk_less orda" by(rule a) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1587 | interpretation b: ccpo lubb ordb "mk_less ordb" by(rule b) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1588 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1589 | lemma ccpo_rel_prodI: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1590 | "class.ccpo (prod_lub luba lubb) (rel_prod orda ordb) (mk_less (rel_prod orda ordb))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1591 | (is "class.ccpo ?lub ?ord ?ord'") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1592 | proof(intro class.ccpo.intro class.ccpo_axioms.intro) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1593 | show "class.order ?ord ?ord'" by(rule order_rel_prodI) intro_locales | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1594 | qed(auto 4 4 simp add: prod_lub_def intro: a.ccpo_Sup_upper b.ccpo_Sup_upper a.ccpo_Sup_least b.ccpo_Sup_least rev_image_eqI dest: chain_rel_prodD1 chain_rel_prodD2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1595 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1596 | interpretation ab: ccpo "prod_lub luba lubb" "rel_prod orda ordb" "mk_less (rel_prod orda ordb)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1597 | by(rule ccpo_rel_prodI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1598 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1599 | lemma monotone_map_prod [simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1600 | "monotone (rel_prod orda ordb) (rel_prod ordc ordd) (map_prod f g) \<longleftrightarrow> | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1601 | monotone orda ordc f \<and> monotone ordb ordd g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1602 | by(auto simp add: monotone_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1603 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1604 | lemma parallel_fixp_induct: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1605 | assumes adm: "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. P (fst x) (snd x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1606 | and f: "monotone orda orda f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1607 | and g: "monotone ordb ordb g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1608 |   and bot: "P (luba {}) (lubb {})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1609 | and step: "\<And>x y. P x y \<Longrightarrow> P (f x) (g y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1610 | shows "P (ccpo.fixp luba orda f) (ccpo.fixp lubb ordb g)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1611 | proof - | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1612 | let ?lub = "prod_lub luba lubb" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1613 | and ?ord = "rel_prod orda ordb" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1614 | and ?P = "\<lambda>(x, y). P x y" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1615 | from adm have adm': "ccpo.admissible ?lub ?ord ?P" by(simp add: split_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1616 | hence "?P (ccpo.fixp (prod_lub luba lubb) (rel_prod orda ordb) (map_prod f g))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1617 | by(rule ab.fixp_induct)(auto simp add: f g step bot) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1618 | also have "ccpo.fixp (prod_lub luba lubb) (rel_prod orda ordb) (map_prod f g) = | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1619 | (ccpo.fixp luba orda f, ccpo.fixp lubb ordb g)" (is "?lhs = (?rhs1, ?rhs2)") | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1620 | proof(rule ab.antisym) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1621 | have "ccpo.admissible ?lub ?ord (\<lambda>xy. ?ord xy (?rhs1, ?rhs2))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1622 | by(rule admissible_leI[OF ccpo_rel_prodI])(auto simp add: prod_lub_def chain_empty intro: a.ccpo_Sup_least b.ccpo_Sup_least) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1623 | thus "?ord ?lhs (?rhs1, ?rhs2)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1624 | by(rule ab.fixp_induct)(auto 4 3 dest: monotoneD[OF f] monotoneD[OF g] simp add: b.fixp_unfold[OF g, symmetric] a.fixp_unfold[OF f, symmetric] f g intro: a.ccpo_Sup_least b.ccpo_Sup_least chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1625 | next | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1626 | have "ccpo.admissible luba orda (\<lambda>x. orda x (fst ?lhs))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1627 | by(rule admissible_leI[OF a])(auto intro: a.ccpo_Sup_least simp add: chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1628 | hence "orda ?rhs1 (fst ?lhs)" using f | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1629 | proof(rule a.fixp_induct) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1630 | fix x | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1631 | assume "orda x (fst ?lhs)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1632 | thus "orda (f x) (fst ?lhs)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1633 | by(subst ab.fixp_unfold)(auto simp add: f g dest: monotoneD[OF f]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1634 | qed(auto intro: a.ccpo_Sup_least chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1635 | moreover | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1636 | have "ccpo.admissible lubb ordb (\<lambda>y. ordb y (snd ?lhs))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1637 | by(rule admissible_leI[OF b])(auto intro: b.ccpo_Sup_least simp add: chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1638 | hence "ordb ?rhs2 (snd ?lhs)" using g | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1639 | proof(rule b.fixp_induct) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1640 | fix y | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1641 | assume "ordb y (snd ?lhs)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1642 | thus "ordb (g y) (snd ?lhs)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1643 | by(subst ab.fixp_unfold)(auto simp add: f g dest: monotoneD[OF g]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1644 | qed(auto intro: b.ccpo_Sup_least chain_empty) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1645 | ultimately show "?ord (?rhs1, ?rhs2) ?lhs" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1646 | by(simp add: rel_prod_conv split_beta) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1647 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1648 | finally show ?thesis by simp | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1649 | qed | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1650 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1651 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1652 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1653 | lemma parallel_fixp_induct_uc: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1654 | assumes a: "partial_function_definitions orda luba" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1655 | and b: "partial_function_definitions ordb lubb" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1656 | and F: "\<And>x. monotone (fun_ord orda) orda (\<lambda>f. U1 (F (C1 f)) x)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1657 | and G: "\<And>y. monotone (fun_ord ordb) ordb (\<lambda>g. U2 (G (C2 g)) y)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1658 | and eq1: "f \<equiv> C1 (ccpo.fixp (fun_lub luba) (fun_ord orda) (\<lambda>f. U1 (F (C1 f))))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1659 | and eq2: "g \<equiv> C2 (ccpo.fixp (fun_lub lubb) (fun_ord ordb) (\<lambda>g. U2 (G (C2 g))))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1660 | and inverse: "\<And>f. U1 (C1 f) = f" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1661 | and inverse2: "\<And>g. U2 (C2 g) = g" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1662 | and adm: "ccpo.admissible (prod_lub (fun_lub luba) (fun_lub lubb)) (rel_prod (fun_ord orda) (fun_ord ordb)) (\<lambda>x. P (fst x) (snd x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1663 |   and bot: "P (\<lambda>_. luba {}) (\<lambda>_. lubb {})"
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1664 | and step: "\<And>f g. P (U1 f) (U2 g) \<Longrightarrow> P (U1 (F f)) (U2 (G g))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1665 | shows "P (U1 f) (U2 g)" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1666 | apply(unfold eq1 eq2 inverse inverse2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1667 | apply(rule parallel_fixp_induct[OF partial_function_definitions.ccpo[OF a] partial_function_definitions.ccpo[OF b] adm]) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1668 | using F apply(simp add: monotone_def fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1669 | using G apply(simp add: monotone_def fun_ord_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1670 | apply(simp add: fun_lub_def bot) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1671 | apply(rule step, simp add: inverse inverse2) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1672 | done | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1673 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1674 | lemmas parallel_fixp_induct_1_1 = parallel_fixp_induct_uc[ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1675 | of _ _ _ _ "\<lambda>x. x" _ "\<lambda>x. x" "\<lambda>x. x" _ "\<lambda>x. x", | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1676 | OF _ _ _ _ _ _ refl refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1677 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1678 | lemmas parallel_fixp_induct_2_2 = parallel_fixp_induct_uc[ | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1679 | of _ _ _ _ "case_prod" _ "curry" "case_prod" _ "curry", | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1680 | where P="\<lambda>f g. P (curry f) (curry g)", | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1681 | unfolded case_prod_curry curry_case_prod curry_K, | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1682 | OF _ _ _ _ _ _ refl refl] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1683 | for P | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1684 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1685 | lemma monotone_fst: "monotone (rel_prod orda ordb) orda fst" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1686 | by(auto intro: monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1687 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1688 | lemma mcont_fst: "mcont (prod_lub luba lubb) (rel_prod orda ordb) luba orda fst" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1689 | by(auto intro!: mcontI monotoneI contI simp add: prod_lub_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1690 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1691 | lemma mcont2mcont_fst [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1692 | "mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) t | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1693 | \<Longrightarrow> mcont lub ord luba orda (\<lambda>x. fst (t x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1694 | by(auto intro!: mcontI monotoneI contI dest: mcont_monoD mcont_contD simp add: rel_prod_sel split_beta prod_lub_def image_image) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1695 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1696 | lemma monotone_snd: "monotone (rel_prod orda ordb) ordb snd" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1697 | by(auto intro: monotoneI) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1698 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1699 | lemma mcont_snd: "mcont (prod_lub luba lubb) (rel_prod orda ordb) lubb ordb snd" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1700 | by(auto intro!: mcontI monotoneI contI simp add: prod_lub_def) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1701 | |
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1702 | lemma mcont2mcont_snd [cont_intro, simp]: | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1703 | "mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) t | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1704 | \<Longrightarrow> mcont lub ord lubb ordb (\<lambda>x. snd (t x))" | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1705 | by(auto intro!: mcontI monotoneI contI dest: mcont_monoD mcont_contD simp add: rel_prod_sel split_beta prod_lub_def image_image) | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1706 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1707 | lemma monotone_Pair: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1708 | "\<lbrakk> monotone ord orda f; monotone ord ordb g \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1709 | \<Longrightarrow> monotone ord (rel_prod orda ordb) (\<lambda>x. (f x, g x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1710 | by(simp add: monotone_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1711 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1712 | lemma cont_Pair: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1713 | "\<lbrakk> cont lub ord luba orda f; cont lub ord lubb ordb g \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1714 | \<Longrightarrow> cont lub ord (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. (f x, g x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1715 | by(rule contI)(auto simp add: prod_lub_def image_image dest!: contD) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1716 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1717 | lemma mcont_Pair: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1718 | "\<lbrakk> mcont lub ord luba orda f; mcont lub ord lubb ordb g \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1719 | \<Longrightarrow> mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. (f x, g x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1720 | by(rule mcontI)(simp_all add: monotone_Pair mcont_mono cont_Pair) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1721 | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1722 | context partial_function_definitions begin | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1723 | text \<open>Specialised versions of @{thm [source] mcont_call} for admissibility proofs for parallel fixpoint inductions\<close>
 | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1724 | lemmas mcont_call_fst [cont_intro] = mcont_call[THEN mcont2mcont, OF mcont_fst] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1725 | lemmas mcont_call_snd [cont_intro] = mcont_call[THEN mcont2mcont, OF mcont_snd] | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1726 | end | 
| 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1727 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1728 | lemma map_option_mono [partial_function_mono]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1729 | "mono_option B \<Longrightarrow> mono_option (\<lambda>f. map_option g (B f))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1730 | unfolding map_conv_bind_option by(rule bind_mono) simp_all | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1731 | |
| 66244 
4c999b5d78e2
qualify Complete_Partial_Order2.compact
 Andreas Lochbihler parents: 
65366diff
changeset | 1732 | lemma compact_flat_lub [cont_intro]: "ccpo.compact (flat_lub x) (flat_ord x) y" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1733 | using flat_interpretation[THEN ccpo] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1734 | proof(rule ccpo.compactI[OF _ ccpo.admissibleI]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1735 | fix A | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1736 | assume chain: "Complete_Partial_Order.chain (flat_ord x) A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1737 |     and A: "A \<noteq> {}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1738 | and *: "\<forall>z\<in>A. \<not> flat_ord x y z" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1739 | from A obtain z where "z \<in> A" by blast | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1740 | with * have z: "\<not> flat_ord x y z" .. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1741 | hence y: "x \<noteq> y" "y \<noteq> z" by(auto simp add: flat_ord_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1742 |   { assume "\<not> A \<subseteq> {x}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1743 | then obtain z' where "z' \<in> A" "z' \<noteq> x" by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1744 |     then have "(THE z. z \<in> A - {x}) = z'"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1745 | by(intro the_equality)(auto dest: chainD[OF chain] simp add: flat_ord_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1746 | moreover have "z' \<noteq> y" using \<open>z' \<in> A\<close> * by(auto simp add: flat_ord_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1747 |     ultimately have "y \<noteq> (THE z. z \<in> A - {x})" by simp }
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1748 | with z show "\<not> flat_ord x y (flat_lub x A)" by(simp add: flat_ord_def flat_lub_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1749 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: 
63170diff
changeset | 1750 | |
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: diff
changeset | 1751 | end |