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