diff --git a/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy --- a/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy +++ b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy @@ -1,939 +1,939 @@ section \CCW for Arbitrary Points in the Plane\ theory Counterclockwise_2D_Arbitrary imports Counterclockwise_2D_Strict begin subsection \Interpretation of Knuth's axioms in the plane\ definition lex::"point \ point \ bool" where "lex p q \ (fst p < fst q \ fst p = fst q \ snd p < snd q \ p = q)" definition psi::"point \ point \ point \ bool" where "psi p q r \ (lex p q \ lex q r)" definition ccw::"point \ point \ point \ bool" where "ccw p q r \ ccw' p q r \ (det3 p q r = 0 \ (psi p q r \ psi q r p \ psi r p q))" interpretation ccw: linorder_list0 "ccw x" for x . lemma ccw'_imp_ccw: "ccw' a b c \ ccw a b c" by (simp add: ccw_def) lemma ccw_ncoll_imp_ccw: "ccw a b c \ \coll a b c \ ccw' a b c" by (simp add: ccw_def) lemma ccw_translate: "ccw p (p + q) (p + r) = ccw 0 q r" by (auto simp: ccw_def psi_def lex_def) lemma ccw_translate_origin: "NO_MATCH 0 p \ ccw p q r = ccw 0 (q - p) (r - p)" using ccw_translate[of p "q - p" "r - p"] by simp lemma psi_scale: "psi (r *\<^sub>R a) (r *\<^sub>R b) 0 = (if r > 0 then psi a b 0 else if r < 0 then psi 0 b a else True)" "psi (r *\<^sub>R a) 0 (r *\<^sub>R b) = (if r > 0 then psi a 0 b else if r < 0 then psi b 0 a else True)" "psi 0 (r *\<^sub>R a) (r *\<^sub>R b) = (if r > 0 then psi 0 a b else if r < 0 then psi b a 0 else True)" by (auto simp: psi_def lex_def det3_def' not_less sign_simps) lemma ccw_scale23: "ccw 0 a b \ r > 0 \ ccw 0 (r *\<^sub>R a) (r *\<^sub>R b)" by (auto simp: ccw_def psi_scale) lemma psi_notI: "distinct3 p q r \ psi p q r \ \ psi q p r" by (auto simp: algebra_simps psi_def lex_def) lemma not_lex_eq: "\ lex a b \ lex b a \ a \ b" by (auto simp: algebra_simps lex_def prod_eq_iff) lemma lex_trans: "lex a b \ lex b c \ lex a c" by (auto simp: lex_def) lemma lex_sym_eqI: "lex a b \ lex b a \ a = b" and lex_sym_eq_iff: "lex a b \ lex b a \ a = b" by (auto simp: lex_def) lemma lex_refl[simp]: "lex p p" by (metis not_lex_eq) lemma psi_disjuncts: "distinct3 p q r \ psi p q r \ psi p r q \ psi q r p \ psi q p r \ psi r p q \ psi r q p" by (auto simp: psi_def not_lex_eq) lemma nlex_ccw_left: "lex x 0 \ ccw 0 (0, 1) x" by (auto simp: ccw_def lex_def psi_def ccw'_def det3_def') interpretation ccw_system123 ccw apply unfold_locales subgoal by (force simp: ccw_def ccw'_def det3_def' algebra_simps) subgoal by (force simp: ccw_def ccw'_def det3_def' psi_def algebra_simps lex_sym_eq_iff) subgoal by (drule psi_disjuncts) (force simp: ccw_def ccw'_def det3_def' algebra_simps) done lemma lex_scaleR_nonneg: "lex a b \ r \ 0 \ lex a (a + r *\<^sub>R (b - a))" by (auto simp: lex_def) lemma lex_scale1_zero: "lex (v *\<^sub>R u) 0 = (if v > 0 then lex u 0 else if v < 0 then lex 0 u else True)" and lex_scale2_zero: "lex 0 (v *\<^sub>R u) = (if v > 0 then lex 0 u else if v < 0 then lex u 0 else True)" by (auto simp: lex_def prod_eq_iff less_eq_prod_def sign_simps) lemma nlex_add: assumes "lex a 0" "lex b 0" shows "lex (a + b) 0" using assms by (auto simp: lex_def) lemma nlex_sum: assumes "finite X" assumes "\x. x \ X \ lex (f x) 0" shows "lex (sum f X) 0" using assms by induction (auto intro!: nlex_add) lemma abs_add_nlex: assumes "coll 0 a b" assumes "lex a 0" assumes "lex b 0" shows "abs (a + b) = abs a + abs b" proof (rule antisym[OF abs_triangle_ineq]) have "fst (\a\ + \b\) \ fst \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) moreover { assume H: "fst a < 0" "fst b < 0" hence "snd b \ 0 \ snd a \ 0" using assms by (auto simp: lex_def det3_def' mult.commute) (metis mult_le_cancel_left_neg mult_zero_right)+ hence "\snd a\ + \snd b\ \ \snd a + snd b\" using H by auto } hence "snd (\a\ + \b\) \ snd \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) ultimately show "\a\ + \b\ \ \a + b\" unfolding less_eq_prod_def .. qed lemma lex_sum_list: "(\x. x \ set xs \ lex x 0) \ lex (sum_list xs) 0" by (induct xs) (auto simp: nlex_add) lemma abs_sum_list_coll: assumes coll: "list_all (coll 0 x) xs" assumes "x \ 0" assumes up: "list_all (\x. lex x 0) xs" shows "abs (sum_list xs) = sum_list (map abs xs)" using assms proof (induct xs) case (Cons y ys) hence "coll 0 x y" "coll 0 x (sum_list ys)" by (auto simp: list_all_iff intro!: coll_sum_list) hence "coll 0 y (sum_list ys)" using \x \ 0\ by (rule coll_trans) hence "\y + sum_list ys\ = abs y + abs (sum_list ys)" using Cons by (subst abs_add_nlex) (auto simp: list_all_iff lex_sum_list) thus ?case using Cons by simp qed simp lemma lex_diff1: "lex (a - b) c = lex a (c + b)" and lex_diff2: "lex c (a - b) = lex (c + b) a" by (auto simp: lex_def) lemma sum_list_eq_0_iff_nonpos: fixes xs::"'a::ordered_ab_group_add list" shows "list_all (\x. x \ 0) xs \ sum_list xs = 0 \ (\n\set xs. n = 0)" by (auto simp: list_all_iff sum_list_sum_nth sum_nonpos_eq_0_iff) (auto simp add: in_set_conv_nth) lemma sum_list_nlex_eq_zeroI: assumes nlex: "list_all (\x. lex x 0) xs" assumes "sum_list xs = 0" assumes "x \ set xs" shows "x = 0" proof - from assms(2) have z1: "sum_list (map fst xs) = 0" and z2: "sum_list (map snd xs) = 0" by (auto simp: prod_eq_iff fst_sum_list snd_sum_list) from nlex have "list_all (\x. x \ 0) (map fst xs)" by (auto simp: lex_def list_all_iff) from sum_list_eq_0_iff_nonpos[OF this] z1 nlex have z1': "list_all (\x. x = 0) (map fst xs)" and "list_all (\x. x \ 0) (map snd xs)" by (auto simp: list_all_iff lex_def) from sum_list_eq_0_iff_nonpos[OF this(2)] z2 have "list_all (\x. x = 0) (map snd xs)" by (simp add: list_all_iff) with z1' show "x = 0" by (auto simp: list_all_iff zero_prod_def assms prod_eq_iff) qed lemma sum_list_eq0I: "(\x\set xs. x = 0) \ sum_list xs = 0" by (induct xs) auto lemma sum_list_nlex_eq_zero_iff: assumes nlex: "list_all (\x. lex x 0) xs" shows "sum_list xs = 0 \ list_all ((=) 0) xs" using assms by (auto intro: sum_list_nlex_eq_zeroI sum_list_eq0I simp: list_all_iff) lemma assumes "lex p q" "lex q r" "0 \ a" "0 \ b" "0 \ c" "a + b + c = 1" assumes comb_def: "comb = a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r" shows lex_convex3: "lex p comb" "lex comb r" proof - from convex3_alt[OF assms(3-6), of p q r] obtain u v where uv: "a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r = p + u *\<^sub>R (q - p) + v *\<^sub>R (r - p)" "0 \ u" "0 \ v" "u + v \ 1" . have "lex p r" using assms by (metis lex_trans) hence "lex (v *\<^sub>R (p - r)) 0" using uv by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (u *\<^sub>R (q - p))" using \lex p q\ uv by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex p comb" unfolding comb_def uv by (simp add: lex_def prod_eq_iff algebra_simps) from comb_def have comb_def': "comb = c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p" by simp from assms have "c + b + a = 1" by simp from convex3_alt[OF assms(5,4,3) this, of r q p] obtain u v where uv: "c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p = r + u *\<^sub>R (q - r) + v *\<^sub>R (p - r)" "0 \ u" "0 \ v" "u + v \ 1" by auto have "lex (u *\<^sub>R (q - r)) 0" using uv \lex q r\ by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (v *\<^sub>R (r - p))" using uv \lex p r\ by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex comb r" unfolding comb_def' uv by (simp add: lex_def prod_eq_iff algebra_simps) qed lemma lex_convex_self2: assumes "lex p q" "0 \ a" "a \ 1" defines "r \ a *\<^sub>R p + (1 - a) *\<^sub>R q" shows "lex p r" (is ?th1) and "lex r q" (is ?th2) using lex_convex3[OF \lex p q\, of q a "1 - a" 0 r] assms by (simp_all add: r_def) lemma lex_uminus0[simp]: "lex (-a) 0 = lex 0 a" by (auto simp: lex_def) lemma lex_fst_zero_imp: "fst x = 0 \ lex x 0 \ lex y 0 \ \coll 0 x y \ ccw' 0 y x" by (auto simp: ccw'_def det3_def' lex_def sign_simps) lemma lex_ccw_left: "lex x y \ r > 0 \ ccw y (y + (0, r)) x" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps lex_def psi_def) lemma lex_translate_origin: "NO_MATCH 0 a \ lex a b = lex 0 (b - a)" by (auto simp: lex_def) subsection \Order prover setup\ definition "lexs p q \ (lex p q \ p \ q)" lemma lexs_irrefl: "\ lexs p p" and lexs_imp_lex: "lexs x y \ lex x y" and not_lexs: "(\ lexs x y) = (lex y x)" and not_lex: "(\ lex x y) = (lexs y x)" and eq_lex_refl: "x = y \ lex x y" by (auto simp: lexs_def lex_def prod_eq_iff) lemma lexs_trans: "lexs x y \ lexs y z \ lexs x z" and lexs_lex_trans: "lexs x y \ lex y z \ lexs x z" and lex_lexs_trans: "lex x y \ lexs y z \ lexs x z" and lex_neq_trans: "lex a b \ a \ b \ lexs a b" and neq_lex_trans: "a \ b \ lex a b \ lexs a b" and lexs_imp_neq: "lexs a b \ a \ b" by (auto simp: lexs_def lex_def prod_eq_iff) declare lexs_irrefl[THEN notE, order add less_reflE: linorder "(=) :: point => point => bool" lex lexs] declare lex_refl[order add le_refl: linorder "(=) :: point => point => bool" lex lexs] declare lexs_imp_lex[order add less_imp_le: linorder "(=) :: point => point => bool" lex lexs] declare not_lexs[THEN iffD2, order add not_lessI: linorder "(=) :: point => point => bool" lex lexs] declare not_lex[THEN iffD2, order add not_leI: linorder "(=) :: point => point => bool" lex lexs] declare not_lexs[THEN iffD1, order add not_lessD: linorder "(=) :: point => point => bool" lex lexs] declare not_lex[THEN iffD1, order add not_leD: linorder "(=) :: point => point => bool" lex lexs] declare lex_sym_eqI[order add eqI: linorder "(=) :: point => point => bool" lex lexs] declare eq_lex_refl[order add eqD1: linorder "(=) :: point => point => bool" lex lexs] declare sym[THEN eq_lex_refl, order add eqD2: linorder "(=) :: point => point => bool" lex lexs] declare lexs_trans[order add less_trans: linorder "(=) :: point => point => bool" lex lexs] declare lexs_lex_trans[order add less_le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_lexs_trans[order add le_less_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_trans[order add le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_neq_trans[order add le_neq_trans: linorder "(=) :: point => point => bool" lex lexs] declare neq_lex_trans[order add neq_le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lexs_imp_neq[order add less_imp_neq: linorder "(=) :: point => point => bool" lex lexs] declare eq_neq_eq_imp_neq[order add eq_neq_eq_imp_neq: linorder "(=) :: point => point => bool" lex lexs] declare not_sym[order add not_sym: linorder "(=) :: point => point => bool" lex lexs] subsection \Contradictions\ lemma assumes d: "distinct4 s p q r" shows contra1: "\(lex p q \ lex q r \ lex r s \ indelta s p q r)" (is ?th1) and contra2: "\(lex s p \ lex p q \ lex q r \ indelta s p q r)" (is ?th2) and contra3: "\(lex p r \ lex p s \ lex q r \ lex q s \ insquare p r q s)" (is ?th3) proof - { assume "det3 s p q = 0" "det3 s q r = 0" "det3 s r p = 0" "det3 p q r = 0" hence ?th1 ?th2 ?th3 using d by (auto simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume *: "\(det3 s p q = 0 \ det3 s q r = 0 \ det3 s r p = 0 \ det3 p q r = 0)" { assume d0: "det3 p q r = 0" with d have "?th1 \ ?th2" by (force simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume dp: "det3 p q r \ 0" have "?th1 \ ?th2" unfolding de_Morgan_disj[symmetric] proof (rule notI, goal_cases) case prems: 1 hence **: "indelta s p q r" by auto hence nonnegs: "det3 p q r \ 0" "0 \ det3 s q r" "0 \ det3 p s r" "0 \ det3 p q s" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps) hence det_pos: "det3 p q r > 0" using dp by simp have det_eq: "det3 s q r + det3 p s r + det3 p q s = det3 p q r" by (auto simp: ccw_def det3_def' algebra_simps) hence det_div_eq: "det3 s q r / det3 p q r + det3 p s r / det3 p q r + det3 p q s / det3 p q r = 1" using det_pos by (auto simp: field_simps) from lex_convex3[OF _ _ _ _ _ det_div_eq convex_comb_dets[OF det_pos, of s]] have "lex p s" "lex s r" using prems by (auto simp: nonnegs) with prems d show False by (simp add: lex_sym_eq_iff) qed } moreover have ?th3 proof (safe, goal_cases) case prems: 1 have nonnegs: "det3 p r q \ 0" "det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" using prems by (auto simp add: ccw_def ccw'_def less_eq_real_def) have dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" by (auto simp: det3_def') hence **: "det3 p r q = 0 \ det3 q s p = 0 \ det3 r q s = 0 \ det3 s p r = 0" using prems by (auto simp: ccw_def ccw'_def) moreover { fix p r q s assume det_pos: "det3 p r q > 0" assume dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" assume nonnegs:"det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" assume g14: "lex p r" "lex p s" "lex q r" "lex q s" assume d: "distinct4 s p q r" let ?sum = "(det3 p r q + det3 q s p) / det3 p r q" have eqs: "det3 s p r = det3 p r s" "det3 r q s = det3 s r q" "det3 q s p = - det3 p s q" by (auto simp: det3_def' algebra_simps) from convex_comb_dets[OF det_pos, of s] have "((det3 p r q / det3 p r q) *\<^sub>R s + (det3 q s p / det3 p r q) *\<^sub>R r) /\<^sub>R ?sum = ((det3 r q s / det3 p r q) *\<^sub>R p + (det3 s p r / det3 p r q) *\<^sub>R q) /\<^sub>R ?sum" unfolding eqs by (simp add: algebra_simps prod_eq_iff) hence srpq: "(det3 p r q / det3 p r q / ?sum) *\<^sub>R s + (det3 q s p / det3 p r q / ?sum) *\<^sub>R r = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" (is "?s *\<^sub>R s + ?r *\<^sub>R r = ?p *\<^sub>R p + ?q *\<^sub>R q") using det_pos by (simp add: algebra_simps inverse_eq_divide) have eqs: "?s + ?r = 1" "?p + ?q = 1" and s: "?s \ 0" "?s \ 1" and r: "?r \ 0" "?r \ 1" and p: "?p \ 0" "?p \ 1" and q: "?q \ 0" "?q \ 1" unfolding add_divide_distrib[symmetric] using det_pos nonnegs dets_eq by (auto) from eqs have eqs': "1 - ?s = ?r" "1 - ?r = ?s" "1 - ?p = ?q" "1 - ?q = ?p" by auto have comm: "?r *\<^sub>R r + ?s *\<^sub>R s = ?s *\<^sub>R s + ?r *\<^sub>R r" "?q *\<^sub>R q + ?p *\<^sub>R p = ?p *\<^sub>R p + ?q *\<^sub>R q" by simp_all define K where "K = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" note rewrs = eqs' comm srpq K_def[symmetric] from lex_convex_self2[OF _ s, of s r, unfolded rewrs] lex_convex_self2[OF _ r, of r s, unfolded rewrs] lex_convex_self2[OF _ p, of p q, unfolded rewrs] lex_convex_self2[OF _ q, of q p, unfolded rewrs] have False using g14 d det_pos by (metis lex_trans not_lex_eq) } note wlog = this from dets_eq have 1: "det3 q s p + det3 p r q = det3 s p r + det3 r q s" by simp from d have d': "distinct4 r q p s" by auto note wlog[of q s p r, OF _ 1 nonnegs(3,2,1) prems(4,3,2,1) d'] wlog[of p r q s, OF _ dets_eq nonnegs(2,3,4) prems(1-4) d] ultimately show False using nonnegs d * by (auto simp: less_eq_real_def det3_def' algebra_simps) qed ultimately have ?th1 ?th2 ?th3 by blast+ } ultimately show ?th1 ?th2 ?th3 by force+ qed lemma ccw'_subst_psi_disj: assumes "det3 t r s = 0" assumes "psi t r s \ psi t s r \ psi s r t" assumes "s \ t" assumes "ccw' t r p" shows "ccw' t s p" proof cases assume "r \ s" from assms have "r \ t" by (auto simp: det3_def' ccw'_def algebra_simps) from assms have "det3 r s t = 0" by (auto simp: algebra_simps det3_def') from coll_ex_scaling[OF assms(3) this] obtain x where s: "r = s + x *\<^sub>R (t - s)" by auto from assms(4)[simplified s] have "0 < det3 0 (s + x *\<^sub>R (t - s) - t) (p - t)" by (auto simp: algebra_simps det3_def' ccw'_def) also have "s + x *\<^sub>R (t - s) - t = (1 - x) *\<^sub>R (s - t)" by (simp add: algebra_simps) finally have ccw': "ccw' 0 ((1 - x) *\<^sub>R (s - t)) (p - t)" by (simp add: ccw'_def) hence neq: "x \ 1" by (auto simp add: det3_def' ccw'_def) have tr: "fst s < fst r \ fst t = fst s \ snd t \ snd r" by (simp add: s) from s have "fst (r - s) = fst (x *\<^sub>R (t - s))" "snd (r - s) = snd (x *\<^sub>R (t - s))" by (auto simp: ) hence "x = (if fst (t - s) = 0 then snd (r - s) / snd (t - s) else fst (r - s) / fst (t - s))" using \s \ t\ by (auto simp add: field_simps prod_eq_iff) also have "\ \ 1" using assms by (auto simp: lex_def psi_def tr) finally have "x < 1" using neq by simp thus ?thesis using ccw' by (auto simp: ccw'.translate_origin) qed (insert assms, simp) lemma lex_contr: assumes "distinct4 t s q r" assumes "lex t s" "lex s r" assumes "det3 t s r = 0" assumes "ccw' t s q" assumes "ccw' t q r" shows "False" using ccw'_subst_psi_disj[of t s r q] assms by (cases "r = t") (auto simp: det3_def' algebra_simps psi_def ccw'_def) lemma contra4: assumes "distinct4 s r q p" assumes lex: "lex q p" "lex p r" "lex r s" assumes ccw: "ccw r q s" "ccw r s p" "ccw r q p" shows False proof cases assume c: "ccw s q p" from c have *: "indelta s r q p" using assms by simp with contra1[OF assms(1)] have "\ (lex r q \ lex q p \ lex p s)" by blast hence "\ lex q p" using \ccw s q p\ contra1 cyclic assms nondegenerate by blast thus False using assms by simp next assume "\ ccw s q p" with ccw have "ccw q s p \ ccw s p r \ ccw p r q \ ccw r q s" by (metis assms(1) ccw'.cyclic ccw_def not_ccw'_eq psi_disjuncts) moreover from lex have "lex q r" "lex q s" "lex p r" "lex p s" by order+ ultimately show False using contra3[of r q p s] \distinct4 s r q p\ by blast qed lemma not_coll_ordered_lexI: assumes "l \ x0" and "lex x1 r" and "lex x1 l" and "lex r x0" and "lex l x0" and "ccw' x0 l x1" and "ccw' x0 x1 r" shows "det3 x0 l r \ 0" proof assume "coll x0 l r" from \coll x0 l r\ have 1: "coll 0 (l - x0) (r - x0)" by (simp add: det3_def' algebra_simps) from \lex r x0\ have 2: "lex (r - x0) 0" by (auto simp add: lex_def) from \lex l x0\ have 3: "lex (l - x0) 0" by (auto simp add: lex_def) from \ccw' x0 l x1\ have 4: "ccw' 0 (l - x0) (x1 - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \ccw' x0 x1 r\ have 5: "ccw' 0 (x1 - x0) (r - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \lex x1 r\ have 6: "lex 0 (r - x0 - (x1 - x0))" by (auto simp: lex_def) from \lex x1 l\ have 7: "lex 0 (l - x0 - (x1 - x0))" by (auto simp: lex_def) define r' where "r' = r - x0" define l' where "l' = l - x0" define x0' where "x0' = x1 - x0" from 1 2 3 4 5 6 7 have rs: "coll 0 l' r'" "lex r' 0" "lex l' 0" "ccw' 0 l' x0'" "ccw' 0 x0' r'" "lex 0 (r' - x0')" "lex 0 (l' - x0')" unfolding r'_def[symmetric] l'_def[symmetric] x0'_def[symmetric] by auto from assms have "l' \ 0" by (auto simp: l'_def) from coll_scale[OF \coll 0 l' _\ this] obtain y where y: "r' = y *\<^sub>R l'" by auto { assume "y > 0" with rs have False by (auto simp: det3_def' algebra_simps y ccw'_def) } moreover { assume "y < 0" with rs have False - by (auto simp: lex_def not_less sign_simps y ccw'_def) + by (auto simp: lex_def not_less algebra_simps sign_simps y ccw'_def) } moreover { assume "y = 0" from this rs have False by (simp add: ccw'_def y) } ultimately show False by arith qed interpretation ccw_system4 ccw proof unfold_locales fix p q r t assume ccw: "ccw t q r" "ccw p t r" "ccw p q t" show "ccw p q r" proof (cases "det3 t q r = 0 \ det3 p t r = 0 \ det3 p q t = 0") case True { assume "psi t q r \ psi q r t \ psi r t q" "psi p t r \ psi t r p \ psi r p t" "psi p q t \ psi q t p \ psi t p q" hence "psi p q r \ psi q r p \ psi r p q" using lex_sym_eq_iff psi_def by blast } with True ccw show ?thesis by (simp add: det3_def' algebra_simps ccw_def ccw'_def) next case False hence "0 \ det3 t q r" "0 \ det3 p t r" "0 \ det3 p q t" using ccw by (auto simp: less_eq_real_def ccw_def ccw'_def) with False show ?thesis by (auto simp: ccw_def det3_def' algebra_simps ccw'_def intro!: disjI1) qed qed lemma lex_total: "lex t q \ t \ q \ lex q t \ t \ q \ t = q" by auto lemma ccw_two_up_contra: assumes c: "ccw' t p q" "ccw' t q r" assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws have nn: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) with c det_identity[of t p q s r] have tsr: "coll t s r" and tsp: "coll t s p" by (auto simp: add_nonneg_eq_0_iff ccw'_def) moreover have trp: "coll t r p" by (metis ccw'_subst_collinear distinct not_ccw'_eq tsr tsp) ultimately have tpr: "coll t p r" by (auto simp: det3_def' algebra_simps) moreover have psi: "psi t p r \ psi t r p \ psi r p t" unfolding psi_def proof - have ntsr: "\ ccw' t s r" "\ ccw' t r s" using tsr by (auto simp: not_ccw'_eq det3_def' algebra_simps) have f8: "\ ccw' t r s" using tsr not_ccw'_eq by blast have f9: "\ ccw' t r p" using tpr by (simp add: not_ccw'_eq) have f10: "(lex t r \ lex r p \ lex r p \ lex p t \ lex p t \ lex t r)" using ccw_def ccws(6) psi_def f9 by auto have "\ ccw' t r q" using c(2) not_ccw'_eq by blast moreover have "\coll t q s" using ntsr ccw'_subst_collinear distinct c(2) by blast hence "ccw' t s q" by (meson ccw_def ccws(2) not_ccw'_eq) moreover from tsr tsp \coll t r p\ have "coll t p s" "coll t p r" "coll t r s" by (auto simp add: det3_def' algebra_simps) ultimately show "lex t p \ lex p r \ lex t r \ lex r p \ lex r p \ lex p t" by (metis ccw'_subst_psi_disj distinct ccw_def ccws(3) contra4 tsp ntsr(1) f10 lex_total psi_def trp) qed moreover from distinct have "r \ t" by auto ultimately have "ccw' t r q" using c(1) by (rule ccw'_subst_psi_disj) thus False using c(2) by (simp add: ccw'_contra) qed lemma ccw_transitive_contr: fixes t s p q r assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws distinct have *: "ccw p t r" "ccw p q t" by (metis cyclic)+ with distinct have "ccw r p q" using interior[OF _ _ ccws(5) *, of UNIV] by (auto intro: cyclic) from ccws have nonnegs: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t p q \ 0" "det3 t q r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) { assume "ccw' t p q" "ccw' t q r" "ccw' t r p" hence False using ccw_two_up_contra ccws distinct by blast } moreover { assume c: "coll t q r" "coll t r p" with distinct four_points_aligned(1)[OF c, of s] have "coll t p q" by auto hence "(psi t p q \ psi p q t \ psi q t p)" "psi t q r \ psi q r t \ psi r t q" "psi t r p \ psi r p t \ psi p t r" using ccws(4,5,6) c by (simp_all add: ccw_def ccw'_def) hence False using distinct by (auto simp: psi_def ccw'_def) } moreover { assume c: "det3 t p q = 0" "det3 t q r > 0" "det3 t r p = 0" have "\x. det3 t q r = 0 \ t = x \ r = q \ q = x \ r = p \ p = x \ r = x" by (meson c(1) c(3) distinct four_points_aligned(1)) hence False by (metis (full_types) c(2) distinct less_irrefl) } moreover { assume c: "det3 t p q = 0" "det3 t q r = 0" "det3 t r p > 0" have "\x. det3 t r p = 0 \ t = x \ r = x \ q = x \ p = x" by (meson c(1) c(2) distinct four_points_aligned(1)) hence False by (metis (no_types) c(3) distinct less_numeral_extra(3)) } moreover { assume c: "ccw' t p q" "ccw' t q r" from ccw_two_up_contra[OF this ccws distinct] have False . } moreover { assume c: "ccw' t p q" "ccw' t r p" from ccw_two_up_contra[OF this(2,1), of s] ccws distinct have False by auto } moreover { assume c: "ccw' t q r" "ccw' t r p" from ccw_two_up_contra[OF this, of s] ccws distinct have False by auto } ultimately show "False" using \0 \ det3 t p q\ \0 \ det3 t q r\\0 \ det3 t r p\ by (auto simp: less_eq_real_def ccw'_def) qed interpretation ccw: ccw_system ccw by unfold_locales (metis ccw_transitive_contr nondegenerate) lemma ccw_scaleR1: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 (e*\<^sub>Rxr) P" by (simp add: ccw_def) lemma ccw_scaleR2: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 xr (e*\<^sub>RP)" by (simp add: ccw_def) lemma ccw_translate3_aux: assumes "\coll 0 a b" assumes "x < 1" assumes "ccw 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" shows "ccw 0 a b" proof - from assms have "\ coll 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" by simp with assms have "ccw' 0 ((1 - x) *\<^sub>R a) (b - x *\<^sub>R a)" by (simp add: algebra_simps ccw_def) thus "ccw 0 a b" using \x < 1\ by (simp add: ccw_def) qed lemma ccw_translate3_minus: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a (b - x *\<^sub>R a) \ ccw 0 a b" using ccw_translate3_aux[of a b x] ccw_scaleR1[of a "b - x *\<^sub>R a" "1-x" ] by (auto simp add: algebra_simps) lemma ccw_translate3: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a b \ ccw 0 a (x *\<^sub>R a + b)" by (rule ccw_translate3_minus) (auto simp add: algebra_simps) lemma ccw_switch23: "det3 0 P Q \ 0 \ (\ ccw 0 Q P \ ccw 0 P Q)" by (auto simp: ccw_def algebra_simps not_ccw'_eq ccw'_not_coll) lemma ccw0_upward: "fst a > 0 \ snd a = 0 \ snd b > snd a \ ccw 0 a b" by (auto simp: ccw_def det3_def' algebra_simps ccw'_def) lemma ccw_uminus3[simp]: "det3 a b c \ 0 \ ccw (-a) (-b) (-c) = ccw a b c" by (auto simp: ccw_def ccw'_def algebra_simps det3_def') lemma coll_minus_eq: "coll (x - a) (x - b) (x - c) = coll a b c" by (auto simp: det3_def' algebra_simps) lemma ccw_minus3: "\ coll a b c \ ccw (x - a) (x - b) (x - c) \ ccw a b c" by (simp add: ccw_def coll_minus_eq) lemma ccw0_uminus[simp]: "\ coll 0 a b \ ccw 0 (-a) (-b) \ ccw 0 a b" using ccw_uminus3[of 0 a b] by simp lemma lex_convex2: assumes "lex p q" "lex p r" "0 \ u" "u \ 1" shows "lex p (u *\<^sub>R q + (1 - u) *\<^sub>R r)" proof cases note \lex p q\ also assume "lex q r" hence "lex q (u *\<^sub>R q + (1 - u) *\<^sub>R r)" using \0 \ u\ \u \ 1\ by (rule lex_convex_self2) finally (lex_trans) show ?thesis . next note \lex p r\ also assume "\ lex q r" hence "lex r q" by simp hence "lex r ((1 - u) *\<^sub>R r + (1 - (1 - u)) *\<^sub>R q)" using \0 \ u\ \u \ 1\ by (intro lex_convex_self2) simp_all finally (lex_trans) show ?thesis by (simp add: ac_simps) qed lemma lex_convex2': assumes "lex q p" "lex r p" "0 \ u" "u \ 1" shows "lex (u *\<^sub>R q + (1 - u) *\<^sub>R r) p" proof - have "lex (- p) (u *\<^sub>R (-q) + (1 - u) *\<^sub>R (-r))" using assms by (intro lex_convex2) (auto simp: lex_def) thus ?thesis by (auto simp: lex_def algebra_simps) qed lemma psi_convex1: assumes "psi c a b" assumes "psi d a b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi (u *\<^sub>R c + v *\<^sub>R d) a b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex2: assumes "psi a c b" assumes "psi a d b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a (u *\<^sub>R c + v *\<^sub>R d) b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex3: assumes "psi a b c" assumes "psi a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2) qed lemma coll_convex: assumes "coll a b c" "coll a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "coll a b (u *\<^sub>R c + v *\<^sub>R d)" proof cases assume "a \ b" with assms(1, 2) obtain x y where xy: "c - a = x *\<^sub>R (b - a)" "d - a = y *\<^sub>R (b - a)" by (auto simp: det3_translate_origin dest!: coll_scale) from assms have "(u + v) *\<^sub>R a = a" by simp hence "u *\<^sub>R c + v *\<^sub>R d - a = u *\<^sub>R (c - a) + v *\<^sub>R (d - a)" by (simp add: algebra_simps) also have "\ = u *\<^sub>R x *\<^sub>R (b - a) + v *\<^sub>R y *\<^sub>R (b - a)" by (simp add: xy) also have "\ = (u * x + v * y) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "coll 0 (b - a) \" by (simp add: coll_scaleR_right_eq) finally show ?thesis by (auto simp: det3_translate_origin) qed simp lemma (in ccw_vector_space) convex3: assumes "u \ 0" "v \ 0" "u + v = 1" "ccw a b d" "ccw a b c" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - have "v = 1 - u" using assms by simp hence "ccw 0 (b - a) (u *\<^sub>R (c - a) + v *\<^sub>R (d - a))" using assms by (cases "u = 0" "v = 0" rule: bool.exhaust[case_product bool.exhaust]) (auto simp add: translate_origin intro!: add3) also have "(u + v) *\<^sub>R a = a" by (simp add: assms) hence "u *\<^sub>R (c - a) + v *\<^sub>R (d - a) = u *\<^sub>R c + v *\<^sub>R d - a" by (auto simp: algebra_simps) finally show ?thesis by (simp add: translate_origin) qed lemma ccw_self[simp]: "ccw a a b" "ccw b a a" by (auto simp: ccw_def psi_def intro: cyclic) lemma ccw_sefl'[simp]: "ccw a b a" by (rule cyclic) simp lemma ccw_convex': assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" and 1: "coll a b c" assumes "ccw a b d" and 2: "\ coll a b d" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) { assume 2: "\ coll a b c" from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume 2: "coll a b c" { assume "a = b" hence ?thesis by simp } moreover { assume "v = 0" hence ?thesis by (auto simp: v assms) } moreover { assume "v \ 0" "a \ b" have "coll c a b" using 2 by (auto simp: det3_def' algebra_simps) from coll_ex_scaling[OF \a \ b\ this] obtain r where c: "c = a + r *\<^sub>R (b - a)" by auto have *: "u *\<^sub>R (a + r *\<^sub>R (b - a)) + v *\<^sub>R d - a = (u * r) *\<^sub>R (b - a) + (1 - u) *\<^sub>R (d - a)" by (auto simp: algebra_simps v) have "ccw' a b ?c" using \v \ 0\ uv abd by (simp add: ccw'.translate_origin c *) hence ?thesis by (simp add: ccw_def) } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma ccw_convex: assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" assumes "ccw a b d" assumes lex: "coll a b c \ coll a b d \ lex b a" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" { assume coll: "coll a b c \ coll a b d" hence "coll a b ?c" by (auto intro!: coll_convex assms) moreover from coll have "psi a b c \ psi b c a \ psi c a b" "psi a b d \ psi b d a \ psi d a b" using assms by (auto simp add: ccw_def ccw'_not_coll) hence "psi a b ?c \ psi b ?c a \ psi ?c a b" using coll uv lex by (auto simp: psi_def ccw_def not_lex lexs_def v intro: lex_convex2 lex_convex2') ultimately have ?thesis by (simp add: ccw_def) } moreover { assume 1: "\ coll a b d" and 2: "\ coll a b c" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume "\ coll a b d" "coll a b c" have ?thesis by (rule ccw_convex') fact+ } moreover { assume 1: "coll a b d" and 2: "\ coll a b c" have "0 \ 1 - u" using assms by (auto ) from ccw_convex'[OF this \0 \ u\ _ \ccw a b d\ 1 \ccw a b c\ 2] have ?thesis by (simp add: algebra_simps v) } ultimately show ?thesis by blast qed interpretation ccw: ccw_convex ccw S "\a b. lex b a" for S by unfold_locales (rule ccw_convex) lemma ccw_sorted_scaleR: "ccw.sortedP 0 xs \ r > 0 \ ccw.sortedP 0 (map ((*\<^sub>R) r) xs)" by (induct xs) (auto intro!: ccw.sortedP.Cons ccw_scale23 elim!: ccw.sortedP_Cons simp del: scaleR_Pair) lemma ccw_sorted_implies_ccw'_sortedP: assumes nonaligned: "\y z. y \ set Ps \ z \ set Ps \ y \ z \ \ coll 0 y z" assumes sorted: "linorder_list0.sortedP (ccw 0) Ps" assumes "distinct Ps" shows "linorder_list0.sortedP (ccw' 0 ) Ps" using assms proof (induction Ps) case (Cons P Ps) { fix p assume p: "p \ set Ps" moreover from p Cons.prems have "ccw 0 P p" by (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately have "ccw' 0 P p" using \distinct (P#Ps)\ by (intro ccw_ncoll_imp_ccw Cons) auto } moreover have "linorder_list0.sortedP (ccw' 0) Ps" using Cons.prems by (intro Cons) (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately show ?case by (auto intro!: linorder_list0.Cons ) qed (auto intro: linorder_list0.Nil) end diff --git a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy --- a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy +++ b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy @@ -1,1077 +1,1077 @@ section \Euclidean Space: Executability\ theory Executable_Euclidean_Space imports "HOL-Analysis.Analysis" "List-Index.List_Index" "HOL-Library.Float" Affine_Arithmetic_Auxiliarities begin subsection \Ordered representation of Basis and Rounding of Components\ class executable_euclidean_space = ordered_euclidean_space + fixes Basis_list eucl_down eucl_truncate_down eucl_truncate_up assumes eucl_down_def: "eucl_down p b = (\i \ Basis. round_down p (b \ i) *\<^sub>R i)" assumes eucl_truncate_down_def: "eucl_truncate_down q b = (\i \ Basis. truncate_down q (b \ i) *\<^sub>R i)" assumes eucl_truncate_up_def: "eucl_truncate_up q b = (\i \ Basis. truncate_up q (b \ i) *\<^sub>R i)" assumes Basis_list[simp]: "set Basis_list = Basis" assumes distinct_Basis_list[simp]: "distinct Basis_list" begin lemma length_Basis_list: "length Basis_list = card Basis" by (metis Basis_list distinct_Basis_list distinct_card) end lemma eucl_truncate_down_zero[simp]: "eucl_truncate_down p 0 = 0" by (auto simp: eucl_truncate_down_def truncate_down_zero) lemma eucl_truncate_up_zero[simp]: "eucl_truncate_up p 0 = 0" by (auto simp: eucl_truncate_up_def) subsection \Instantiations\ instantiation real::executable_euclidean_space begin definition Basis_list_real :: "real list" where "Basis_list_real = [1]" definition "eucl_down prec b = round_down prec b" definition "eucl_truncate_down prec b = truncate_down prec b" definition "eucl_truncate_up prec b = truncate_up prec b" instance proof qed (auto simp: Basis_list_real_def eucl_down_real_def eucl_truncate_down_real_def eucl_truncate_up_real_def) end instantiation prod::(executable_euclidean_space, executable_euclidean_space) executable_euclidean_space begin definition Basis_list_prod :: "('a \ 'b) list" where "Basis_list_prod = zip Basis_list (replicate (length (Basis_list::'a list)) 0) @ zip (replicate (length (Basis_list::'b list)) 0) Basis_list" definition "eucl_down p a = (eucl_down p (fst a), eucl_down p (snd a))" definition "eucl_truncate_down p a = (eucl_truncate_down p (fst a), eucl_truncate_down p (snd a))" definition "eucl_truncate_up p a = (eucl_truncate_up p (fst a), eucl_truncate_up p (snd a))" instance proof show "set Basis_list = (Basis::('a*'b) set)" by (auto simp: Basis_list_prod_def Basis_prod_def elim!: in_set_zipE) (auto simp: Basis_list[symmetric] in_set_zip in_set_conv_nth simp del: Basis_list) show "distinct (Basis_list::('a*'b)list)" using distinct_Basis_list[where 'a='a] distinct_Basis_list[where 'a='b] by (auto simp: Basis_list_prod_def Basis_list intro: distinct_zipI1 distinct_zipI2 elim!: in_set_zipE) qed (auto simp: eucl_down_prod_def eucl_truncate_down_prod_def eucl_truncate_up_prod_def sum_Basis_prod_eq inner_add_left inner_sum_left inner_Basis eucl_down_def eucl_truncate_down_def eucl_truncate_up_def intro!: euclidean_eqI[where 'a="'a*'b"]) end lemma eucl_truncate_down_Basis[simp]: "i \ Basis \ eucl_truncate_down e x \ i = truncate_down e (x \ i)" by (simp add: eucl_truncate_down_def) lemma eucl_truncate_down_correct: "dist (x::'a::executable_euclidean_space) (eucl_down e x) \ {0..sqrt (DIM('a)) * 2 powr of_int (- e)}" proof - have "dist x (eucl_down e x) = sqrt (\i\Basis. (dist (x \ i) (eucl_down e x \ i))\<^sup>2)" unfolding euclidean_dist_l2[where 'a='a] L2_set_def .. also have "\ \ sqrt (\i\(Basis::'a set). ((2 powr of_int (- e))\<^sup>2))" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_real_def eucl_down_def abs_round_down_le) finally show ?thesis by (simp add: real_sqrt_mult) qed lemma eucl_down: "eucl_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] round_down eucl_down_def) lemma eucl_truncate_down: "eucl_truncate_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] truncate_down) lemma eucl_truncate_down_le: "x \ y \ eucl_truncate_down w x \ (y::'a::executable_euclidean_space)" using eucl_truncate_down by (rule order.trans) lemma eucl_truncate_up_Basis[simp]: "i \ Basis \ eucl_truncate_up e x \ i = truncate_up e (x \ i)" by (simp add: eucl_truncate_up_def truncate_up_def) lemma eucl_truncate_up: "x \ eucl_truncate_up e (x::'a::executable_euclidean_space)" by (auto simp add: eucl_le[where 'a='a] round_up truncate_up_def) lemma eucl_truncate_up_le: "x \ y \ x \ eucl_truncate_up e (y::'a::executable_euclidean_space)" using _ eucl_truncate_up by (rule order.trans) lemma eucl_truncate_down_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_down p x \ eucl_truncate_down p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_down_mono) lemma eucl_truncate_up_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_up p x \ eucl_truncate_up p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_up_mono) lemma infnorm[code]: fixes x::"'a::executable_euclidean_space" shows "infnorm x = fold max (map (\i. abs (x \ i)) Basis_list) 0" by (auto simp: Max.set_eq_fold[symmetric] infnorm_Max[symmetric] infnorm_pos_le intro!: max.absorb2[symmetric]) declare Inf_real_def[code del] declare Sup_real_def[code del] declare Inf_prod_def[code del] declare Sup_prod_def[code del] declare [[code abort: "Inf::real set \ real"]] declare [[code abort: "Sup::real set \ real"]] declare [[code abort: "Inf::('a::Inf * 'b::Inf) set \ 'a * 'b"]] declare [[code abort: "Sup::('a::Sup * 'b::Sup) set \ 'a * 'b"]] lemma nth_Basis_list_in_Basis[simp]: "n < length (Basis_list::'a::executable_euclidean_space list) \ Basis_list ! n \ (Basis::'a set)" by (metis Basis_list nth_mem) subsection \Representation as list\ lemma nth_eq_iff_index: "distinct xs \ n < length xs \ xs ! n = i \ n = index xs i" using index_nth_id by fastforce lemma in_Basis_index_Basis_list: "i \ Basis \ i = Basis_list ! index Basis_list i" by simp lemmas [simp] = length_Basis_list lemma sum_Basis_sum_nth_Basis_list: "(\i\Basis. f i) = (\i(x, i)\zip xs Basis_list. x *\<^sub>R i)" lemma eucl_of_list_nth: assumes "length xs = DIM('a)" shows "eucl_of_list xs = (\iR ((Basis_list::'a list) ! i))" by (auto simp: eucl_of_list_def sum_list_sum_nth length_Basis_list assms atLeast0LessThan) lemma eucl_of_list_inner: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" by (simp add: eucl_of_list_nth[OF l] inner_sum_left assms inner_Basis nth_eq_iff_index sum.delta if_distrib cong: if_cong) lemma inner_eucl_of_list: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "i \ eucl_of_list xs = xs ! (index Basis_list i)" using eucl_of_list_inner[OF assms] by (auto simp: inner_commute) definition "list_of_eucl x = map ((\) x) Basis_list" lemma index_Basis_list_nth[simp]: "i < DIM('a::executable_euclidean_space) \ index Basis_list ((Basis_list::'a list) ! i) = i" by (simp add: index_nth_id) lemma list_of_eucl_eucl_of_list[simp]: "length xs = DIM('a::executable_euclidean_space) \ list_of_eucl (eucl_of_list xs::'a) = xs" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: nth_equalityI) lemma eucl_of_list_list_of_eucl[simp]: "eucl_of_list (list_of_eucl x) = x" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) lemma length_list_of_eucl[simp]: "length (list_of_eucl (x::'a::executable_euclidean_space)) = DIM('a)" by (auto simp: list_of_eucl_def) lemma list_of_eucl_nth[simp]: "n < DIM('a::executable_euclidean_space) \ list_of_eucl x ! n = x \ (Basis_list ! n::'a)" by (auto simp: list_of_eucl_def) lemma nth_ge_len: "n \ length xs \ xs ! n = [] ! (n - length xs)" by (induction xs arbitrary: n) auto lemma list_of_eucl_nth_if: "list_of_eucl x ! n = (if n < DIM('a::executable_euclidean_space) then x \ (Basis_list ! n::'a) else [] ! (n - DIM('a)))" apply (auto simp: list_of_eucl_def ) apply (subst nth_ge_len) apply auto done lemma list_of_eucl_eq_iff: "list_of_eucl (x::'a::executable_euclidean_space) = list_of_eucl (y::'b::executable_euclidean_space) \ (DIM('a) = DIM('b) \ (\i < DIM('b). x \ Basis_list ! i = y \ Basis_list ! i))" by (auto simp: list_eq_iff_nth_eq) lemma eucl_le_Basis_list_iff: "(x::'a::executable_euclidean_space) \ y \ (\i Basis_list ! i \ y \ Basis_list ! i)" apply (auto simp: eucl_le[where 'a='a]) subgoal for i subgoal by (auto dest!: spec[where x="index Basis_list i"]) done done lemma eucl_of_list_inj: "length xs = DIM('a::executable_euclidean_space) \ length ys = DIM('a) \ (eucl_of_list xs::'a) = eucl_of_list (ys) \ xs = ys" apply (auto intro!: nth_equalityI simp: euclidean_eq_iff[where 'a="'a"] eucl_of_list_inner) using nth_Basis_list_in_Basis[where 'a="'a"] by fastforce lemma eucl_of_list_map_plus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x + g x) xs)::'a) = eucl_of_list (map f xs) + eucl_of_list (map g xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_uminus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. - f x) xs)::'a) = - eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_left[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. r * f x) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x * r) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_divide_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x / r) xs)::'a) = eucl_of_list (map f xs) /\<^sub>R r" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner divide_simps) lemma eucl_of_list_map_const[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. c) xs)::'a) = c *\<^sub>R One" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma replicate_eq_list_of_eucl_zero: "replicate DIM('a::executable_euclidean_space) 0 = list_of_eucl (0::'a)" by (auto intro!: nth_equalityI) lemma eucl_of_list_append_zeroes[simp]: "eucl_of_list (xs @ replicate n 0) = eucl_of_list xs" unfolding eucl_of_list_def apply (auto simp: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) by (auto simp: nth_append) lemma Basis_prodD: assumes "(i, j) \ Basis" shows "i \ Basis \ j = 0 \ i = 0 \ j \ Basis" using assms by (auto simp: Basis_prod_def) lemma eucl_of_list_take_DIM[simp]: assumes "d = DIM('b::executable_euclidean_space)" shows "(eucl_of_list (take d xs)::'b) = (eucl_of_list xs)" by (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list sum_list_sum_nth assms dest!: Basis_prodD) lemma eucl_of_list_eqI: assumes "take DIM('a) (xs @ replicate (DIM('a) - length xs) 0) = take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)" shows "eucl_of_list xs = (eucl_of_list ys::'a::executable_euclidean_space)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ replicate (DIM('a) - length xs) 0))" by (simp add: ) also note assms also have "eucl_of_list (take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)) = (eucl_of_list ys::'a)" by simp finally show ?thesis . qed lemma eucl_of_list_replicate_zero[simp]: "eucl_of_list (replicate E 0) = 0" proof - have "eucl_of_list (replicate E 0) = (eucl_of_list (replicate E 0 @ replicate (DIM('a) - E) 0)::'a)" by simp also have "\ = eucl_of_list (replicate DIM('a) 0)" apply (rule eucl_of_list_eqI) by (auto simp: min_def nth_append intro!: nth_equalityI) also have "\ = 0" by (simp add: replicate_eq_list_of_eucl_zero) finally show ?thesis by simp qed lemma eucl_of_list_Nil[simp]: "eucl_of_list [] = 0" using eucl_of_list_replicate_zero[of 0] by simp lemma fst_eucl_of_list_prod: shows "fst (eucl_of_list xs::'b::executable_euclidean_space \ _) = (eucl_of_list (take DIM('b) xs)::'b)" apply (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list dest!: Basis_prodD) apply (simp add: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) subgoal by simp subgoal by auto subgoal by (auto simp: Basis_list_prod_def nth_append) subgoal by (auto simp: Basis_list_prod_def nth_append) done lemma index_zip_replicate1[simp]: "index (zip (replicate d a) bs) (a, b) = index bs b" if "d = length bs" using that by (induction bs arbitrary: d) auto lemma index_zip_replicate2[simp]: "index (zip as (replicate d b)) (a, b) = index as a" if "d = length as" using that by (induction as arbitrary: d) auto lemma index_Basis_list_prod[simp]: fixes a::"'a::executable_euclidean_space" and b::"'b::executable_euclidean_space" shows "a \ Basis \ index Basis_list (a, 0::'b) = index Basis_list a" "b \ Basis \ index Basis_list (0::'a, b) = DIM('a) + index Basis_list b" by (auto simp: Basis_list_prod_def index_append in_set_zip zip_replicate index_map_inj dest: spec[where x="index Basis_list a"]) lemma eucl_of_list_eq_takeI: assumes "(eucl_of_list (take DIM('a::executable_euclidean_space) xs)::'a) = x" shows "eucl_of_list xs = x" using eucl_of_list_take_DIM[OF refl, of xs, where 'b='a] assms by auto lemma eucl_of_list_inner_le: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs \ DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ (replicate (DIM('a) - length xs) 0)))" by (rule eucl_of_list_eq_takeI) simp also have "\ \ i = xs ! (index Basis_list i)" using assms by (subst eucl_of_list_inner) auto finally show ?thesis . qed lemma eucl_of_list_prod_if: assumes "length xs = DIM('a::executable_euclidean_space) + DIM('b::executable_euclidean_space)" shows "eucl_of_list xs = (eucl_of_list (take DIM('a) xs)::'a, eucl_of_list (drop DIM('a) xs)::'b)" apply (rule euclidean_eqI) using assms apply (auto simp: eucl_of_list_inner dest!: Basis_prodD) apply (subst eucl_of_list_inner_le) apply (auto simp: Basis_list_prod_def index_append in_set_zip) done lemma snd_eucl_of_list_prod: shows "snd (eucl_of_list xs::'b::executable_euclidean_space \ 'c::executable_euclidean_space) = (eucl_of_list (drop DIM('b) xs)::'c)" proof (cases "length xs \ DIM('b)") case True then show ?thesis by (auto simp: eucl_of_list_inner eucl_of_list_def snd_sum_list dest!: Basis_prodD) (simp add: sum_list_sum_nth Basis_list_prod_def nth_append) next case False have "xs = take DIM('b) xs @ drop DIM('b) xs" by simp also have "eucl_of_list \ = (eucl_of_list (\ @ replicate (length xs - DIM('c)) 0)::'b \ 'c)" by simp finally have "eucl_of_list xs = (eucl_of_list (xs @ replicate (DIM('b) + DIM('c) - length xs) 0)::'b \ 'c)" by simp also have "\ = eucl_of_list (take (DIM ('b \ 'c)) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by (simp add: ) finally have *: "(eucl_of_list xs::'b\'c) = eucl_of_list (take DIM('b \ 'c) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by simp show ?thesis apply (subst *) apply (subst eucl_of_list_prod_if) subgoal by simp subgoal apply simp apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (rule arg_cong[where f=eucl_of_list]) by (auto intro!: nth_equalityI simp: nth_append min_def split: if_splits) done qed lemma eucl_of_list_prod: shows "eucl_of_list xs = (eucl_of_list (take DIM('b) xs)::'b::executable_euclidean_space, eucl_of_list (drop DIM('b) xs)::'c::executable_euclidean_space)" using snd_eucl_of_list_prod[of xs, where 'b='b and 'c='c] using fst_eucl_of_list_prod[of xs, where 'b='b and 'a='c] by (auto simp del: snd_eucl_of_list_prod fst_eucl_of_list_prod simp add: prod_eq_iff) lemma eucl_of_list_real[simp]: "eucl_of_list [x] = (x::real)" by (auto simp: eucl_of_list_def Basis_list_real_def) lemma eucl_of_list_append[simp]: assumes "length xs = DIM('i::executable_euclidean_space)" assumes "length ys = DIM('j::executable_euclidean_space)" shows "eucl_of_list (xs @ ys) = (eucl_of_list xs::'i, eucl_of_list ys::'j)" using assms by (auto simp: eucl_of_list_prod) lemma list_allI: "(\x. x \ set xs \ P x) \ list_all P xs" by (auto simp: list_all_iff) lemma concat_map_nthI: assumes "\x y. x \ set xs \ y \ set (f x) \ P y" assumes "j < length (concat (map f xs))" shows "P (concat (map f xs) ! j)" proof - have "list_all P (concat (map f xs))" by (rule list_allI) (auto simp: assms) then show ?thesis by (auto simp: list_all_length assms) qed lemma map_nth_append1: assumes "length xs = d" shows "map ((!) (xs @ ys)) [0.. y = z" by (metis eucl_of_list_list_of_eucl) lemma length_Basis_list_pos[simp]: "length Basis_list > 0" by (metis length_pos_if_in_set Basis_list SOME_Basis) lemma Basis_list_nth_nonzero: "i < length (Basis_list::'a::executable_euclidean_space list) \ (Basis_list::'a list) ! i \ 0" by (auto dest!: nth_mem simp: nonzero_Basis) lemma nth_Basis_list_prod: "i < DIM('a) + DIM('b) \ (Basis_list::('a::executable_euclidean_space \ 'b::executable_euclidean_space) list) ! i = (if i < DIM('a) then (Basis_list ! i, 0) else (0, Basis_list ! (i - DIM('a))))" by (auto simp: Basis_list_nth_nonzero prod_eq_iff Basis_list_prod_def nth_append not_less) lemma eucl_of_list_if: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" "distinct xs" shows "eucl_of_list (map (\xa. if xa = x then 1 else 0) (xs::nat list)) = (if x \ set xs then Basis_list ! index xs x else 0::'a)" by (rule euclidean_eqI) (auto simp: eucl_of_list_inner inner_Basis index_nth_id) lemma take_append_take_minus_idem: "take n XS @ map ((!) XS) [n..b\Basis. f b)" by (subst sum_list_distinct_conv_sum_set) (auto simp: Basis_list distinct_Basis_list) lemma hd_Basis_list[simp]: "hd Basis_list \ Basis" unfolding Basis_list[symmetric] by (rule hd_in_set) (auto simp: set_empty[symmetric]) definition "inner_lv_rel a b = sum_list (map2 (*) a b)" lemma eucl_of_list_inner_eq: "(eucl_of_list xs::'a) \ eucl_of_list ys = inner_lv_rel xs ys" if "length xs = DIM('a::executable_euclidean_space)" "length ys = DIM('a)" using that by (subst euclidean_inner[abs_def], subst sum_list_Basis_list[symmetric]) (auto simp: eucl_of_list_inner sum_list_sum_nth index_nth_id inner_lv_rel_def) lemma euclidean_vec_componentwise: "(\(xa::'a::euclidean_space^'b::finite)\Basis. f xa) = (\a\Basis. (\b::'b\UNIV. f (axis b a)))" apply (auto simp: Basis_vec_def) apply (subst sum.swap) apply (subst sum.Union_disjoint) apply auto apply (simp add: axis_eq_axis nonzero_Basis) apply (simp add: axis_eq_axis nonzero_Basis) apply (subst sum.reindex) apply (auto intro!: injI) subgoal apply (auto simp: set_eq_iff) by (metis (full_types) all_not_in_conv inner_axis_axis inner_eq_zero_iff nonempty_Basis nonzero_Basis) apply (rule sum.cong[OF refl]) apply (auto ) apply (rule sum.reindex_cong[OF _ _ refl]) apply (auto intro!: inj_onI) using axis_eq_axis by blast lemma vec_nth_inner_scaleR_craziness: "f (x $ i \ j) *\<^sub>R j = (\xa\UNIV. f (x $ xa \ j) *\<^sub>R axis xa j) $ i" by vector (auto simp: axis_def if_distrib scaleR_vec_def sum.delta' cong: if_cong) instantiation vec :: ("{executable_euclidean_space}", enum) executable_euclidean_space begin definition Basis_list_vec :: "('a, 'b) vec list" where "Basis_list_vec = concat (map (\n. map (axis n) Basis_list) enum_class.enum)" definition eucl_down_vec :: "int \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_down_vec p x = (\ i. eucl_down p (x $ i))" definition eucl_truncate_down_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_down_vec p x = (\ i. eucl_truncate_down p (x $ i))" definition eucl_truncate_up_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_up_vec p x = (\ i. eucl_truncate_up p (x $ i))" instance proof show *: "set (Basis_list::('a, 'b) vec list) = Basis" unfolding Basis_list_vec_def Basis_vec_def apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map Basis_vec_def intro!: distinct_concat inj_onI split: if_splits) apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map enum_distinct UNIV_enum[symmetric] intro!: distinct_concat inj_onI split: if_splits) done have "length (Basis_list::('a, 'b) vec list) = CARD('b) * DIM('a)" by (auto simp: Basis_list_vec_def length_concat o_def enum_distinct sum_list_distinct_conv_sum_set UNIV_enum[symmetric]) then show "distinct (Basis_list::('a, 'b) vec list)" using * by (auto intro!: card_distinct) qed (simp_all only: vector_cart[symmetric] vec_eq_iff eucl_down_vec_def eucl_down_def eucl_truncate_down_vec_def eucl_truncate_down_def eucl_truncate_up_vec_def eucl_truncate_up_def, auto simp: euclidean_vec_componentwise inner_axis Basis_list_vec_def vec_nth_inner_scaleR_craziness intro!: sum.cong[OF refl]) end lemma concat_same_lengths_nth: assumes "\xs. xs \ set XS \ length xs = N" assumes "i < length XS * N" "N > 0" shows "concat XS ! i = XS ! (i div N) ! (i mod N)" using assms apply (induction XS arbitrary: i) apply (auto simp: nth_append nth_Cons split: nat.splits) apply (simp add: div_eq_0_iff) by (metis Suc_inject div_geq mod_geq) lemma concat_map_map_index: shows "concat (map (\n. map (f n) xs) ys) = map (\i. f (ys ! (i div length xs)) (xs ! (i mod length xs))) [0..(x, y)\zip xs (map g xs). f x y) = (\x\set xs. f x (g x))" by (force simp add: sum_list_distinct_conv_sum_set assms distinct_zipI1 split_beta' in_set_zip in_set_conv_nth inj_on_convol_ident intro!: sum.reindex_cong[where l="\x. (x, g x)"]) lemma sum_list_zip_map_of: assumes "distinct bs" assumes "length xs = length bs" shows "(\(x, y)\zip xs bs. f x y) = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" proof - have "(\(x, y)\zip xs bs. f x y) = (\(y, x)\zip bs xs. f x y)" by (subst zip_commute) (auto simp: o_def split_beta') also have "\ = (\(x, y)\zip bs (map (the o map_of (zip bs xs)) bs). f y x)" proof (rule arg_cong, rule map_cong) have "xs = (map (the \ map_of (zip bs xs)) bs)" using assms by (auto intro!: nth_equalityI simp: map_nth map_of_zip_nth) then show "zip bs xs = zip bs (map (the \ map_of (zip bs xs)) bs)" by simp qed auto also have "\ = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" using assms(1) by (subst sum_list_zip_map) (auto simp: o_def) finally show ?thesis . qed lemma vec_nth_matrix: "vec_nth (vec_nth (matrix y) i) j = vec_nth (y (axis j 1)) i" unfolding matrix_def by simp lemma matrix_eqI: assumes "\x. x \ Basis \ A *v x = B *v x" shows "(A::real^'n^'n) = B" apply vector using assms apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis matrix_vector_mul_component) lemma matrix_columnI: assumes "\i. column i A = column i B" shows "(A::real^'n^'n) = B" using assms apply vector apply (auto simp: column_def) apply vector by (metis iso_tuple_UNIV_I vec_lambda_inject) lemma vec_nth_Basis: fixes x::"real^'n" shows "x \ Basis \ vec_nth x i = (if x = axis i 1 then 1 else 0)" apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis inner_axis_axis) lemma vec_nth_eucl_of_list_eq: "length M = CARD('n) \ vec_nth (eucl_of_list M::real^'n::enum) i = M ! index Basis_list (axis i (1::real))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i 1::real^'n::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) simp qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n by (metis DIM_cart DIM_real index_Basis_list_nth mult.right_neutral prems(2) prems(3)) done done done lemma index_Basis_list_axis1: "index Basis_list (axis i (1::real)) = index enum_class.enum i" apply (auto simp: Basis_list_vec_def Basis_list_real_def ) apply (subst index_map_inj) by (auto intro!: injI simp: axis_eq_axis) lemma vec_nth_eq_list_of_eucl1: "(vec_nth (M::real^'n::enum) i) = list_of_eucl M ! (index enum_class.enum i)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq) unfolding index_Basis_list_axis1 by auto lemma enum_3[simp]: "(enum_class.enum::3 list) = [0, 1, 2]" by code_simp+ lemma three_eq_zero: "(3::3) = 0" by simp lemma forall_3': "(\i::3. P i) \ P 0 \ P 1 \ P 2" using forall_3 three_eq_zero by auto lemma euclidean_eq_list_of_euclI: "x = y" if "list_of_eucl x = list_of_eucl y" using that by (metis eucl_of_list_list_of_eucl) lemma axis_one_neq_zero[simp]: "axis xa (1::'a::zero_neq_one) \ 0" by (auto simp: axis_def vec_eq_iff) lemma eucl_of_list_vec_nth3[simp]: "(eucl_of_list [g, h, i]::real^3) $ 0 = g" "(eucl_of_list [g, h, i]::real^3) $ 1 = h" "(eucl_of_list [g, h, i]::real^3) $ 2 = i" "(eucl_of_list [g, h, i]::real^3) $ 3 = g" by (auto simp: cart_eq_inner_axis eucl_of_list_inner vec_nth_eq_list_of_eucl1 index_Basis_list_axis1) type_synonym R3 = "real*real*real" lemma Basis_list_R3: "Basis_list = [(1,0,0), (0, 1, 0), (0, 0, 1)::R3]" by (auto simp: Basis_list_prod_def Basis_list_real_def zero_prod_def) lemma Basis_list_vec3: "Basis_list = [axis 0 1::real^3, axis 1 1, axis 2 1]" by (auto simp: Basis_list_vec_def Basis_list_real_def) lemma eucl_of_list3[simp]: "eucl_of_list [a, b, c] = (a, b, c)" by (auto simp: eucl_of_list_inner Basis_list_vec_def zero_prod_def Basis_prod_def Basis_list_vec3 Basis_list_R3 intro!: euclidean_eqI[where 'a=R3]) subsection \Bounded Linear Functions\ subsection \bounded linear functions\ locale blinfun_syntax begin no_notation vec_nth (infixl "$" 90) notation blinfun_apply (infixl "$" 999) end lemma bounded_linear_via_derivative: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" \ \TODO: generalize?\ assumes "\i. ((\x. blinfun_apply (f x) i) has_derivative (\x. f' y x i)) (at y)" shows "bounded_linear (f' y x)" proof - interpret linear "f' y x" proof (unfold_locales, goal_cases) case (1 v w) from has_derivative_unique[OF assms[of "v + w", unfolded blinfun.bilinear_simps] has_derivative_add[OF assms[of v] assms[of w]], THEN fun_cong, of x] show ?case . next case (2 r v) from has_derivative_unique[OF assms[of "r *\<^sub>R v", unfolded blinfun.bilinear_simps] has_derivative_scaleR_right[OF assms[of v], of r], THEN fun_cong, of x] show ?case . qed let ?bnd = "\i\Basis. norm (f' y x i)" { fix v have "f' y x v = (\i\Basis. (v \ i) *\<^sub>R f' y x i)" by (subst euclidean_representation[symmetric]) (simp add: sum scaleR) also have "norm \ \ norm v * ?bnd" by (auto intro!: order.trans[OF norm_sum] sum_mono mult_right_mono simp: sum_distrib_left Basis_le_norm) finally have "norm (f' y x v) \ norm v * ?bnd" . } then show ?thesis by unfold_locales auto qed definition blinfun_scaleR::"('a::real_normed_vector \\<^sub>L real) \ 'b::real_normed_vector \ ('a \\<^sub>L 'b)" where "blinfun_scaleR a b = blinfun_scaleR_left b o\<^sub>L a" lemma blinfun_scaleR_transfer[transfer_rule]: "rel_fun (pcr_blinfun (=) (=)) (rel_fun (=) (pcr_blinfun (=) (=))) (\a b c. a c *\<^sub>R b) blinfun_scaleR" by (auto simp: blinfun_scaleR_def rel_fun_def pcr_blinfun_def cr_blinfun_def OO_def) lemma blinfun_scaleR_rep_eq[simp]: "blinfun_scaleR a b c = a c *\<^sub>R b" by (simp add: blinfun_scaleR_def) lemma bounded_linear_blinfun_scaleR: "bounded_linear (blinfun_scaleR a)" unfolding blinfun_scaleR_def[abs_def] by (auto intro!: bounded_linear_intros) lemma blinfun_scaleR_has_derivative[derivative_intros]: assumes "(f has_derivative f') (at x within s)" shows "((\x. blinfun_scaleR a (f x)) has_derivative (\x. blinfun_scaleR a (f' x))) (at x within s)" using bounded_linear_blinfun_scaleR assms by (rule bounded_linear.has_derivative) lemma blinfun_componentwise: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" shows "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (auto intro!: blinfun_eqI simp: blinfun.sum_left euclidean_representation blinfun.scaleR_right[symmetric] blinfun.sum_right[symmetric]) lemma blinfun_has_derivative_componentwiseI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. i \ Basis \ ((\x. f x i) has_derivative blinfun_apply (f' i)) (at x)" shows "(f has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' i x))) (at x)" by (subst blinfun_componentwise) (force intro: derivative_eq_intros assms simp: blinfun.bilinear_simps) lemma has_derivative_BlinfunI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. ((\x. f x i) has_derivative (\x. f' y x i)) (at y)" shows "(f has_derivative (\x. Blinfun (f' y x))) (at y)" proof - have 1: "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (rule blinfun_componentwise) moreover have 2: "(\ has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i))) (at y)" by (force intro: assms derivative_eq_intros) moreover interpret f': bounded_linear "f' y x" for x by (rule bounded_linear_via_derivative) (rule assms) have 3: "(\i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i)) i = f' y x i" for x i by (auto simp: if_distrib if_distribR blinfun.bilinear_simps f'.scaleR[symmetric] f'.sum[symmetric] euclidean_representation intro!: blinfun_euclidean_eqI) have 4: "blinfun_apply (Blinfun (f' y x)) = f' y x" for x apply (subst bounded_linear_Blinfun_apply) subgoal by unfold_locales subgoal by simp done show ?thesis apply (subst 1) apply (rule 2[THEN has_derivative_eq_rhs]) apply (rule ext) apply (rule blinfun_eqI) apply (subst 3) apply (subst 4) apply (rule refl) done qed lemma has_derivative_Blinfun: assumes "(f has_derivative f') F" shows "(f has_derivative Blinfun f') F" using assms by (subst bounded_linear_Blinfun_apply) auto lift_definition flip_blinfun:: "('a::real_normed_vector \\<^sub>L 'b::real_normed_vector \\<^sub>L 'c::real_normed_vector) \ 'b \\<^sub>L 'a \\<^sub>L 'c" is "\f x y. f y x" using bounded_bilinear.bounded_linear_left bounded_bilinear.bounded_linear_right bounded_bilinear.flip by auto lemma flip_blinfun_apply[simp]: "flip_blinfun f a b = f b a" by transfer simp lemma le_norm_blinfun: shows "norm (blinfun_apply f x) / norm x \ norm f" by transfer (rule le_onorm) lemma norm_flip_blinfun[simp]: "norm (flip_blinfun x) = norm x" (is "?l = ?r") proof (rule antisym) from order_trans[OF norm_blinfun, OF mult_right_mono, OF norm_blinfun, OF norm_ge_zero, of x] show "?l \ ?r" by (auto intro!: norm_blinfun_bound simp: ac_simps) have "norm (x a b) \ norm (flip_blinfun x) * norm a * norm b" for a b proof - have "norm (x a b) / norm a \ norm (flip_blinfun x b)" by (rule order_trans[OF _ le_norm_blinfun]) auto also have "\ \ norm (flip_blinfun x) * norm b" by (rule norm_blinfun) finally show ?thesis - by (auto simp add: divide_simps blinfun.bilinear_simps sign_simps split: if_split_asm) + by (auto simp add: divide_simps blinfun.bilinear_simps algebra_simps split: if_split_asm) qed then show "?r \ ?l" by (auto intro!: norm_blinfun_bound) qed lemma bounded_linear_flip_blinfun[bounded_linear]: "bounded_linear flip_blinfun" by unfold_locales (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI exI[where x=1]) lemma dist_swap2_swap2[simp]: "dist (flip_blinfun f) (flip_blinfun g) = dist f g" by (metis (no_types) bounded_linear_flip_blinfun dist_blinfun_def linear_simps(2) norm_flip_blinfun) context includes blinfun.lifting begin lift_definition blinfun_of_vmatrix::"(real^'m^'n) \ ((real^('m::finite)) \\<^sub>L (real^('n::finite)))" is "matrix_vector_mult:: ((real, 'm) vec, 'n) vec \ ((real, 'm) vec \ (real, 'n) vec)" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma matrix_blinfun_of_vmatrix[simp]: "matrix (blinfun_of_vmatrix M) = M" apply vector apply (auto simp: matrix_def) apply transfer by (metis cart_eq_inner_axis matrix_vector_mul_component) end lemma blinfun_apply_componentwise: "B = (\i\Basis. blinfun_scaleR (blinfun_inner_left i) (blinfun_apply B i))" using blinfun_componentwise[of "\x. B", unfolded fun_eq_iff] by blast lemma blinfun_apply_eq_sum: assumes [simp]: "length v = CARD('n)" shows "blinfun_apply (B::(real^'n::enum)\\<^sub>L(real^'m::enum)) (eucl_of_list v) = (\ij Basis_list ! i) * v ! j) *\<^sub>R (Basis_list ! i))" apply (subst blinfun_apply_componentwise[of B]) apply (auto intro!: euclidean_eqI[where 'a="(real,'m) vec"] simp: blinfun.bilinear_simps eucl_of_list_inner inner_sum_left inner_Basis if_distrib sum_Basis_sum_nth_Basis_list nth_eq_iff_index if_distribR cong: if_cong) apply (subst sum.swap) by (auto simp: sum.delta algebra_simps) lemma in_square_lemma[intro, simp]: "x * C + y < D * C" if "x < D" "y < C" for x::nat proof - have "x * C + y < (D - 1) * C + C" apply (rule add_le_less_mono) apply (rule mult_right_mono) using that by auto also have "\ \ D * C" using that by (auto simp: algebra_simps) finally show ?thesis . qed lemma less_square_imp_div_less[intro, simp]: "i < E * D \ i div E < D" for i::nat by (metis div_eq_0_iff div_mult2_eq gr_implies_not0 mult_not_zero) lemma in_square_lemma'[intro, simp]: "i < L \ n < N \ i * N + n < N * L" for i n::nat by (metis in_square_lemma mult.commute) lemma distinct_nth_eq_iff: "distinct xs \ x < length xs \ y < length xs \ xs ! x = xs ! y \ x = y" by (drule inj_on_nth[where I="{..i. axis (enum_class.enum ! (i div CARD('i))) (axis (enum_class.enum ! (i mod CARD('i))) 1)) ((index enum_class.enum j) * CARD('i) + index enum_class.enum i)" by (auto simp: index_less_cardi enum_UNIV) note less=in_square_lemma[OF index_less_cardj index_less_cardi, of j i] show ?thesis apply (subst *) apply (subst index_map_inj_on[where S="{.. Basis \ vec_nth (vec_nth x i) j = ((if x = axis i (axis j 1) then 1 else 0))" by (auto simp: Basis_vec_def axis_def) lemma vec_nth_eucl_of_list_eq2: "length M = CARD('n) * CARD('m) \ vec_nth (vec_nth (eucl_of_list M::real^'n::enum^'m::enum) i) j = M ! index Basis_list (axis i (axis j (1::real)))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis2) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i (axis j 1)::real^'n::enum^'m::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum^'m::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) auto qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n proof - have "n < card (Basis::(real^'n::_^'m::_) set)" by (simp add: prems(4)) then show ?thesis by (metis index_Basis_list_nth prems(2)) qed done done done lemma vec_nth_eq_list_of_eucl2: "vec_nth (vec_nth (M::real^'n::enum^'m::enum) i) j = list_of_eucl M ! (index enum_class.enum i * CARD('n) + index enum_class.enum j)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq2) unfolding index_Basis_list_axis2 by auto theorem eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list: assumes "length M = CARD('n) * CARD('m)" assumes "length v = CARD('n)" shows "(eucl_of_list M::real^'n::enum^'m::enum) *v eucl_of_list v = (\ijR Basis_list ! i)" apply (vector matrix_vector_mult_def) apply (auto simp: ) apply (subst vec_nth_eucl_of_list_eq2) apply (auto simp: assms) apply (subst vec_nth_eucl_of_list_eq) apply (auto simp: assms index_Basis_list_axis2 index_Basis_list_axis1 vec_nth_Basis sum.delta nth_eq_iff_index if_distrib cong: if_cong) subgoal for i apply (rule sum.reindex_cong[where l="nth enum_class.enum"]) apply (auto simp: enum_distinct card_UNIV_length_enum distinct_nth_eq_iff intro!: inj_onI) apply (rule image_eqI[OF ]) apply (rule nth_index[symmetric]) apply (auto simp: enum_UNIV) by (auto simp: algebra_simps enum_UNIV enum_distinct index_nth_id) subgoal for i using index_less[of i "enum_class.enum" "CARD('n)"] by (auto simp: enum_UNIV card_UNIV_length_enum) done lemma index_enum_less[intro, simp]: "index enum_class.enum (i::'n::enum) < CARD('n)" by (auto intro!: index_less simp: enum_UNIV card_UNIV_length_enum) lemmas [intro, simp] = enum_distinct lemmas [simp] = card_UNIV_length_enum[symmetric] enum_UNIV lemma sum_index_enum_eq: "(\(k::'n::enum)\UNIV. f (index enum_class.enum k)) = (\i *) section \An alternative Sturm sequences\ theory Extended_Sturm imports "Sturm_Tarski.Sturm_Tarski" "Winding_Number_Eval.Cauchy_Index_Theorem" begin text \The main purpose of this theory is to provide an effective way to compute @{term "cindexE a b f"} when @{term f} is a rational function. The idea is similar to and based on the evaluation of @{term cindex_poly} through @{thm cindex_poly_changes_itv_mods}.\ text \ This alternative version of remainder sequences is inspired by the paper "The Fundamental Theorem of Algebra made effective: an elementary real-algebraic proof via Sturm chains" by Michael Eisermann. \ hide_const Permutations.sign subsection \Misc\ lemma is_unit_pCons_ex_iff: fixes p::"'a::field poly" shows "is_unit p \ (\a. a\0 \ p=[:a:])" using is_unit_poly_iff is_unit_triv by auto lemma poly_gcd_iff: "poly (gcd p q) x=0 \ poly p x=0 \ poly q x=0 " by (simp add: poly_eq_0_iff_dvd) lemma eventually_poly_nz_at_within: fixes x :: "'a::{idom,euclidean_space} " assumes "p\0" shows "eventually (\x. poly p x\0) (at x within S)" proof - have "eventually (\x. poly p x\0) (at x within S) = (\\<^sub>F x in (at x within S). \y\proots p. x \ y)" apply (rule eventually_subst,rule eventuallyI) by (auto simp add:proots_def) also have "... = (\y\proots p. \\<^sub>F x in (at x within S). x \ y)" apply (subst eventually_ball_finite_distrib) using \p\0\ by auto also have "..." unfolding eventually_at by (metis gt_ex not_less_iff_gr_or_eq zero_less_dist_iff) finally show ?thesis . qed lemma sgn_power: fixes x::"'a::linordered_idom" shows "sgn (x^n) = (if n=0 then 1 else if even n then \sgn x\ else sgn x)" apply (induct n) by (auto simp add:sgn_mult) lemma poly_divide_filterlim_at_top: fixes p q::"real poly" defines "ll\( if degree q 0 then at_top else at_bot)" assumes "p\0" "q\0" shows "filterlim (\x. poly q x / poly p x) ll at_top" proof - define pp where "pp=(\x. poly p x / x^(degree p))" define qq where "qq=(\x. poly q x / x^(degree q))" define dd where "dd=(\x::real. if degree p>degree q then 1/x^(degree p - degree q) else x^(degree q - degree p))" have divide_cong:"\\<^sub>Fx in at_top. poly q x / poly p x = qq x / pp x * dd x" proof (rule eventually_at_top_linorderI[of 1]) fix x assume "(x::real)\1" then have "x\0" by auto then show "poly q x / poly p x = qq x / pp x * dd x" unfolding qq_def pp_def dd_def using assms by (auto simp add:field_simps power_diff) qed have qqpp_tendsto:"((\x. qq x / pp x) \ lead_coeff q / lead_coeff p) at_top" proof - have "(qq \ lead_coeff q) at_top" unfolding qq_def using poly_divide_tendsto_aux[of q] by (auto elim!:filterlim_mono simp:at_top_le_at_infinity) moreover have "(pp \ lead_coeff p) at_top" unfolding pp_def using poly_divide_tendsto_aux[of p] by (auto elim!:filterlim_mono simp:at_top_le_at_infinity) ultimately show ?thesis using \p\0\ by (auto intro!:tendsto_eq_intros) qed have ?thesis when "degree qx. poly q x / poly p x) (at 0) at_top" proof (rule filterlim_atI) show "((\x. poly q x / poly p x) \ 0) at_top" using poly_divide_tendsto_0_at_infinity[OF that] by (auto elim:filterlim_mono simp:at_top_le_at_infinity) have "\\<^sub>F x in at_top. poly q x \0" "\\<^sub>F x in at_top. poly p x \0" using poly_eventually_not_zero[OF \q\0\] poly_eventually_not_zero[OF \p\0\] filter_leD[OF at_top_le_at_infinity] by auto then show "\\<^sub>F x in at_top. poly q x / poly p x \ 0" apply eventually_elim by auto qed then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q=degree p" proof - have "((\x. poly q x / poly p x) \ lead_coeff q / lead_coeff p) at_top" using divide_cong qqpp_tendsto that unfolding dd_def by (auto dest:tendsto_cong) then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q>degree p" "sgn_pos_inf q * sgn_pos_inf p > 0" proof - have "filterlim (\x. (qq x / pp x) * dd x) at_top at_top" proof (subst filterlim_tendsto_pos_mult_at_top_iff[OF qqpp_tendsto]) show "0 < lead_coeff q / lead_coeff p" using that(2) unfolding sgn_pos_inf_def by (simp add: zero_less_divide_iff zero_less_mult_iff) show "filterlim dd at_top at_top" unfolding dd_def using that(1) by (auto intro!:filterlim_pow_at_top simp:filterlim_ident) qed then have "LIM x at_top. poly q x / poly p x :> at_top" using filterlim_cong[OF _ _ divide_cong] by blast then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q>degree p" "\ sgn_pos_inf q * sgn_pos_inf p > 0" proof - have "filterlim (\x. (qq x / pp x) * dd x) at_bot at_top" proof (subst filterlim_tendsto_neg_mult_at_bot_iff[OF qqpp_tendsto]) show "lead_coeff q / lead_coeff p < 0" using that(2) \p\0\ \q\0\ unfolding sgn_pos_inf_def by (metis divide_eq_0_iff divide_sgn leading_coeff_0_iff linorder_neqE_linordered_idom sgn_divide sgn_greater) show "filterlim dd at_top at_top" unfolding dd_def using that(1) by (auto intro!:filterlim_pow_at_top simp:filterlim_ident) qed then have "LIM x at_top. poly q x / poly p x :> at_bot" using filterlim_cong[OF _ _ divide_cong] by blast then show ?thesis unfolding ll_def using that by auto qed ultimately show ?thesis by linarith qed lemma poly_divide_filterlim_at_bot: fixes p q::"real poly" defines "ll\( if degree q 0 then at_top else at_bot)" assumes "p\0" "q\0" shows "filterlim (\x. poly q x / poly p x) ll at_bot" proof - define pp where "pp=(\x. poly p x / x^(degree p))" define qq where "qq=(\x. poly q x / x^(degree q))" define dd where "dd=(\x::real. if degree p>degree q then 1/x^(degree p - degree q) else x^(degree q - degree p))" have divide_cong:"\\<^sub>Fx in at_bot. poly q x / poly p x = qq x / pp x * dd x" proof (rule eventually_at_bot_linorderI[of "-1"]) fix x assume "(x::real)\-1" then have "x\0" by auto then show "poly q x / poly p x = qq x / pp x * dd x" unfolding qq_def pp_def dd_def using assms by (auto simp add:field_simps power_diff) qed have qqpp_tendsto:"((\x. qq x / pp x) \ lead_coeff q / lead_coeff p) at_bot" proof - have "(qq \ lead_coeff q) at_bot" unfolding qq_def using poly_divide_tendsto_aux[of q] by (auto elim!:filterlim_mono simp:at_bot_le_at_infinity) moreover have "(pp \ lead_coeff p) at_bot" unfolding pp_def using poly_divide_tendsto_aux[of p] by (auto elim!:filterlim_mono simp:at_bot_le_at_infinity) ultimately show ?thesis using \p\0\ by (auto intro!:tendsto_eq_intros) qed have ?thesis when "degree qx. poly q x / poly p x) (at 0) at_bot" proof (rule filterlim_atI) show "((\x. poly q x / poly p x) \ 0) at_bot" using poly_divide_tendsto_0_at_infinity[OF that] by (auto elim:filterlim_mono simp:at_bot_le_at_infinity) have "\\<^sub>F x in at_bot. poly q x \0" "\\<^sub>F x in at_bot. poly p x \0" using poly_eventually_not_zero[OF \q\0\] poly_eventually_not_zero[OF \p\0\] filter_leD[OF at_bot_le_at_infinity] by auto then show "\\<^sub>F x in at_bot. poly q x / poly p x \ 0" by eventually_elim auto qed then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q=degree p" proof - have "((\x. poly q x / poly p x) \ lead_coeff q / lead_coeff p) at_bot" using divide_cong qqpp_tendsto that unfolding dd_def by (auto dest:tendsto_cong) then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q>degree p" "sgn_neg_inf q * sgn_neg_inf p > 0" proof - define cc where "cc=lead_coeff q / lead_coeff p" have "(cc > 0 \ even (degree q - degree p)) \ (cc<0 \ odd (degree q - degree p))" proof - have "even (degree q - degree p) \ (even (degree q) \ even (degree p)) \ (odd (degree q) \ odd (degree p))" using \degree q>degree p\ by auto then show ?thesis using that \p\0\ \q\0\ unfolding sgn_neg_inf_def cc_def zero_less_mult_iff divide_less_0_iff zero_less_divide_iff apply (simp add:if_split[of "(<) 0"] if_split[of "(>) 0"]) by argo qed moreover have "filterlim (\x. (qq x / pp x) * dd x) at_top at_bot" when "cc>0" "even (degree q - degree p)" proof (subst filterlim_tendsto_pos_mult_at_top_iff[OF qqpp_tendsto]) show "0 < lead_coeff q / lead_coeff p" using \cc>0\ unfolding cc_def by auto show "filterlim dd at_top at_bot" unfolding dd_def using \degree q>degree p\ that(2) by (auto intro!:filterlim_pow_at_bot_even simp:filterlim_ident) qed moreover have "filterlim (\x. (qq x / pp x) * dd x) at_top at_bot" when "cc<0" "odd (degree q - degree p)" proof (subst filterlim_tendsto_neg_mult_at_top_iff[OF qqpp_tendsto]) show "0 > lead_coeff q / lead_coeff p" using \cc<0\ unfolding cc_def by auto show "filterlim dd at_bot at_bot" unfolding dd_def using \degree q>degree p\ that(2) by (auto intro!:filterlim_pow_at_bot_odd simp:filterlim_ident) qed ultimately have "filterlim (\x. (qq x / pp x) * dd x) at_top at_bot" by blast then have "LIM x at_bot. poly q x / poly p x :> at_top" using filterlim_cong[OF _ _ divide_cong] by blast then show ?thesis unfolding ll_def using that by auto qed moreover have ?thesis when "degree q>degree p" "\ sgn_neg_inf q * sgn_neg_inf p > 0" proof - define cc where "cc=lead_coeff q / lead_coeff p" have "(cc < 0 \ even (degree q - degree p)) \ (cc > 0 \ odd (degree q - degree p))" proof - have "even (degree q - degree p) \ (even (degree q) \ even (degree p)) \ (odd (degree q) \ odd (degree p))" using \degree q>degree p\ by auto then show ?thesis using that \p\0\ \q\0\ unfolding sgn_neg_inf_def cc_def zero_less_mult_iff divide_less_0_iff zero_less_divide_iff apply (simp add:if_split[of "(<) 0"] if_split[of "(>) 0"]) by (metis leading_coeff_0_iff linorder_neqE_linordered_idom) qed moreover have "filterlim (\x. (qq x / pp x) * dd x) at_bot at_bot" when "cc<0" "even (degree q - degree p)" proof (subst filterlim_tendsto_neg_mult_at_bot_iff[OF qqpp_tendsto]) show "0 > lead_coeff q / lead_coeff p" using \cc<0\ unfolding cc_def by auto show "filterlim dd at_top at_bot" unfolding dd_def using \degree q>degree p\ that(2) by (auto intro!:filterlim_pow_at_bot_even simp:filterlim_ident) qed moreover have "filterlim (\x. (qq x / pp x) * dd x) at_bot at_bot" when "cc>0" "odd (degree q - degree p)" proof (subst filterlim_tendsto_pos_mult_at_bot_iff[OF qqpp_tendsto]) show "0 < lead_coeff q / lead_coeff p" using \cc>0\ unfolding cc_def by auto show "filterlim dd at_bot at_bot" unfolding dd_def using \degree q>degree p\ that(2) by (auto intro!:filterlim_pow_at_bot_odd simp:filterlim_ident) qed ultimately have "filterlim (\x. (qq x / pp x) * dd x) at_bot at_bot" by blast then have "LIM x at_bot. poly q x / poly p x :> at_bot" using filterlim_cong[OF _ _ divide_cong] by blast then show ?thesis unfolding ll_def using that by auto qed ultimately show ?thesis by linarith qed subsection \Alternative definition of cross\ definition cross_alt :: "real poly \real poly \ real \ real \ int" where "cross_alt p q a b=\sign (poly p a) - sign (poly q a)\ - \sign (poly p b) - sign(poly q b)\" lemma cross_alt_coprime_0: assumes "coprime p q" "p=0\q=0" shows "cross_alt p q a b=0" proof - have ?thesis when "p=0" proof - have "is_unit q" using that \coprime p q\ by simp then obtain a where "a\0" "q=[:a:]" using is_unit_pCons_ex_iff by blast then show ?thesis using that unfolding cross_alt_def by auto qed moreover have ?thesis when "q=0" proof - have "is_unit p" using that \coprime p q\ by simp then obtain a where "a\0" "p=[:a:]" using is_unit_pCons_ex_iff by blast then show ?thesis using that unfolding cross_alt_def by auto qed ultimately show ?thesis using \p=0\q=0\ by auto qed lemma cross_alt_0[simp]: "cross_alt 0 0 a b=0" unfolding cross_alt_def by auto lemma cross_alt_poly_commute: "cross_alt p q a b = cross_alt q p a b" unfolding cross_alt_def by auto lemma cross_alt_clear_n: assumes "coprime p q" shows "cross_alt p q a b = cross_alt 1 (p*q) a b" proof - have "\sign (poly p a) - sign (poly q a)\ = \1 - sign (poly p a) * sign (poly q a)\" proof (cases "poly p a=0 \ poly q a=0") case True then have False using assms using coprime_poly_0 by blast then show ?thesis by auto next case False then show ?thesis unfolding Sturm_Tarski.sign_def by force qed moreover have "\sign (poly p b) - sign (poly q b)\ = \1 - sign (poly p b) * sign (poly q b)\" proof (cases "poly p b=0 \ poly q b=0") case True then have False using assms using coprime_poly_0 by blast then show ?thesis by auto next case False then show ?thesis unfolding Sturm_Tarski.sign_def by force qed - ultimately show ?thesis - unfolding cross_alt_def - by (simp only:sign_times poly_mult poly_1 sign_simps) + ultimately show ?thesis + by (simp add: cross_alt_def sign_times) qed subsection \Alternative sign variation sequencse\ fun changes_alt:: "('a ::linordered_idom) list \ int" where "changes_alt [] = 0"| "changes_alt [_] = 0" | "changes_alt (x1#x2#xs) = abs(sign x1 - sign x2) + changes_alt (x2#xs)" definition changes_alt_poly_at::"('a ::linordered_idom) poly list \ 'a \ int" where "changes_alt_poly_at ps a= changes_alt (map (\p. poly p a) ps)" definition changes_alt_itv_smods:: "real \ real \real poly \ real poly \ int" where "changes_alt_itv_smods a b p q= (let ps= smods p q in changes_alt_poly_at ps a - changes_alt_poly_at ps b)" lemma changes_alt_itv_smods_rec: assumes "a q = 0 \ q dvd p") case True moreover have "p=0 \ q=0 \ ?thesis" using cross_alt_coprime_0[OF \coprime p q\] unfolding changes_alt_itv_smods_def changes_alt_poly_at_def by fastforce moreover have "\p\0;q\0;p mod q = 0\ \ ?thesis" unfolding changes_alt_itv_smods_def changes_alt_poly_at_def cross_alt_def by (simp add:sign_times) ultimately show ?thesis by auto (auto elim: dvdE) next case False hence "p\0" "q\0" "p mod q\0" by auto then obtain ps where ps:"smods p q=p#q#-(p mod q)#ps" "smods q (-(p mod q)) = q#-(p mod q)#ps" by auto define changes_diff where "changes_diff\\x. changes_alt_poly_at (p#q#-(p mod q)#ps) x - changes_alt_poly_at (q#-(p mod q)#ps) x" have "changes_diff a - changes_diff b=cross_alt p q a b" unfolding changes_diff_def changes_alt_poly_at_def cross_alt_def by simp thus ?thesis unfolding changes_alt_itv_smods_def changes_diff_def changes_alt_poly_at_def ps by force qed subsection \jumpF on polynomials\ definition jumpF_polyR:: "real poly \ real poly \ real \ real" where "jumpF_polyR q p a = jumpF (\x. poly q x / poly p x) (at_right a)" definition jumpF_polyL:: "real poly \ real poly \ real \ real" where "jumpF_polyL q p a = jumpF (\x. poly q x / poly p x) (at_left a)" definition jumpF_poly_top:: "real poly \ real poly \ real" where "jumpF_poly_top q p = jumpF (\x. poly q x / poly p x) at_top" definition jumpF_poly_bot:: "real poly \ real poly \ real" where "jumpF_poly_bot q p = jumpF (\x. poly q x / poly p x) at_bot" lemma jumpF_polyR_0[simp]: "jumpF_polyR 0 p a = 0" "jumpF_polyR q 0 a = 0" unfolding jumpF_polyR_def by auto lemma jumpF_polyL_0[simp]: "jumpF_polyL 0 p a = 0" "jumpF_polyL q 0 a = 0" unfolding jumpF_polyL_def by auto lemma jumpF_polyR_mult_cancel: assumes "p'\0" shows "jumpF_polyR (p' * q) (p' * p) a = jumpF_polyR q p a" unfolding jumpF_polyR_def proof (rule jumpF_cong) obtain ub where "a < ub" "\z. a < z \ z \ ub \ poly p' z \ 0" using next_non_root_interval[OF \p'\0\,of a] by auto then show "\\<^sub>F x in at_right a. poly (p' * q) x / poly (p' * p) x = poly q x / poly p x" apply (unfold eventually_at_right) apply (intro exI[where x=ub]) by auto qed simp lemma jumpF_polyL_mult_cancel: assumes "p'\0" shows "jumpF_polyL (p' * q) (p' * p) a = jumpF_polyL q p a" unfolding jumpF_polyL_def proof (rule jumpF_cong) obtain lb where "lb < a" "\z. lb \ z \ z < a \ poly p' z \ 0 " using last_non_root_interval[OF \p'\0\,of a] by auto then show "\\<^sub>F x in at_left a. poly (p' * q) x / poly (p' * p) x = poly q x / poly p x" apply (unfold eventually_at_left) apply (intro exI[where x=lb]) by auto qed simp lemma jumpF_poly_noroot: assumes "poly p a\0" shows "jumpF_polyL q p a = 0" "jumpF_polyR q p a = 0" subgoal unfolding jumpF_polyL_def using assms apply (intro jumpF_not_infinity) by (auto intro!:continuous_intros) subgoal unfolding jumpF_polyR_def using assms apply (intro jumpF_not_infinity) by (auto intro!:continuous_intros) done lemma jumpF_polyR_coprime: assumes "coprime p q" shows "jumpF_polyR q p x = (if p \ 0 \ q \ 0 \ poly p x=0 then if sign_r_pos p x \ poly q x>0 then 1/2 else - 1/2 else 0)" proof (cases "p=0 \ q=0 \ poly p x\0") case True then show ?thesis using jumpF_poly_noroot by fastforce next case False then have asm:"p\0" "q\0" "poly p x=0" by auto then have "poly q x\0" using assms using coprime_poly_0 by blast have ?thesis when "sign_r_pos p x \ poly q x>0" proof - have "(poly p has_sgnx sgn (poly q x)) (at_right x)" by (metis False \poly q x \ 0\ add.inverse_neutral has_sgnx_imp_sgnx less_not_sym neg_less_iff_less poly_has_sgnx_values(2) sgn_if sign_r_pos_sgnx_iff that trivial_limit_at_right_real zero_less_one) then have "LIM x at_right x. poly q x / poly p x :> at_top" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "poly q x"]) apply (auto simp add:\poly q x\0\) by (metis asm(3) poly_tendsto(3)) then have "jumpF_polyR q p x = 1/2" unfolding jumpF_polyR_def jumpF_def by auto then show ?thesis using that False by auto qed moreover have ?thesis when "\ (sign_r_pos p x \ poly q x>0)" proof - have "(poly p has_sgnx - sgn (poly q x)) (at_right x)" proof - have "(0::real) < 1 \ \ (1::real) < 0 \ sign_r_pos p x \ (poly p has_sgnx - sgn (poly q x)) (at_right x)" by simp then show ?thesis by (metis (no_types) False \poly q x \ 0\ add.inverse_inverse has_sgnx_imp_sgnx neg_less_0_iff_less poly_has_sgnx_values(2) sgn_if sgn_less sign_r_pos_sgnx_iff that trivial_limit_at_right_real) qed then have "LIM x at_right x. poly q x / poly p x :> at_bot" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "poly q x"]) apply (auto simp add:\poly q x\0\) by (metis asm(3) poly_tendsto(3)) then have "jumpF_polyR q p x = - 1/2" unfolding jumpF_polyR_def jumpF_def by auto then show ?thesis using that False by auto qed ultimately show ?thesis by auto qed lemma jumpF_polyL_coprime: assumes "coprime p q" shows "jumpF_polyL q p x = (if p \ 0 \ q \ 0 \ poly p x=0 then if even (order x p) \ sign_r_pos p x \ poly q x>0 then 1/2 else - 1/2 else 0)" proof (cases "p=0 \ q=0 \ poly p x\0") case True then show ?thesis using jumpF_poly_noroot by fastforce next case False then have asm:"p\0" "q\0" "poly p x=0" by auto then have "poly q x\0" using assms using coprime_poly_0 by blast have ?thesis when "even (order x p) \ sign_r_pos p x \ poly q x>0" proof - consider (lt) "poly q x>0" | (gt) " poly q x<0" using \poly q x\0\ by linarith then have "sgnx (poly p) (at_left x) = sgn (poly q x)" apply cases subgoal using that sign_r_pos_sgnx_iff poly_sgnx_values[OF \p\0\,of x] apply (subst poly_sgnx_left_right[OF \p\0\]) by auto subgoal using that sign_r_pos_sgnx_iff poly_sgnx_values[OF \p\0\,of x] apply (subst poly_sgnx_left_right[OF \p\0\]) by auto done then have "(poly p has_sgnx sgn (poly q x)) (at_left x)" by (metis sgnx_able_poly(2) sgnx_able_sgnx) then have "LIM x at_left x. poly q x / poly p x :> at_top" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "poly q x"]) apply (auto simp add:\poly q x\0\) by (metis asm(3) poly_tendsto(2)) then have "jumpF_polyL q p x = 1/2" unfolding jumpF_polyL_def jumpF_def by auto then show ?thesis using that False by auto qed moreover have ?thesis when "\ (even (order x p) \ sign_r_pos p x \ poly q x>0)" proof - consider (lt) "poly q x>0" | (gt) " poly q x<0" using \poly q x\0\ by linarith then have "sgnx (poly p) (at_left x) = - sgn (poly q x)" apply cases subgoal using that sign_r_pos_sgnx_iff poly_sgnx_values[OF \p\0\,of x] apply (subst poly_sgnx_left_right[OF \p\0\]) by auto subgoal using that sign_r_pos_sgnx_iff poly_sgnx_values[OF \p\0\,of x] apply (subst poly_sgnx_left_right[OF \p\0\]) by auto done then have "(poly p has_sgnx - sgn (poly q x)) (at_left x)" by (metis sgnx_able_poly(2) sgnx_able_sgnx) then have "LIM x at_left x. poly q x / poly p x :> at_bot" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "poly q x"]) apply (auto simp add:\poly q x\0\) by (metis asm(3) poly_tendsto(2)) then have "jumpF_polyL q p x = - 1/2" unfolding jumpF_polyL_def jumpF_def by auto then show ?thesis using that False by auto qed ultimately show ?thesis by auto qed lemma jumpF_times: assumes tendsto:"(f \ c) F" and "c\0" "F\bot" shows "jumpF (\x. f x * g x) F = sgn c * jumpF g F" proof - have "c>0 \ c<0" using \c\0\ by auto moreover have ?thesis when "c>0" proof - note filterlim_tendsto_pos_mult_at_top_iff[OF tendsto \c>0\,of g] moreover note filterlim_tendsto_pos_mult_at_bot_iff[OF tendsto \c>0\,of g] moreover have "sgn c = 1" using \c>0\ by auto ultimately show ?thesis unfolding jumpF_def by auto qed moreover have ?thesis when "c<0" proof - define atbot where "atbot = filterlim g at_bot F" define attop where "attop = filterlim g at_top F" have "jumpF (\x. f x * g x) F = (if atbot then 1 / 2 else if attop then - 1 / 2 else 0)" proof - note filterlim_tendsto_neg_mult_at_top_iff[OF tendsto \c<0\,of g] moreover note filterlim_tendsto_neg_mult_at_bot_iff[OF tendsto \c<0\,of g] ultimately show ?thesis unfolding jumpF_def atbot_def attop_def by auto qed also have "... = - (if attop then 1 / 2 else if atbot then - 1 / 2 else 0)" proof - have False when atbot attop using filterlim_at_top_at_bot[OF _ _ \F\bot\] that unfolding atbot_def attop_def by auto then show ?thesis by fastforce qed also have "... = sgn c * jumpF g F" using \c<0\ unfolding jumpF_def attop_def atbot_def by auto finally show ?thesis . qed ultimately show ?thesis by auto qed lemma jumpF_polyR_inverse_add: assumes "coprime p q" shows "jumpF_polyR q p x + jumpF_polyR p q x = jumpF_polyR 1 (q*p) x" proof (cases "p=0 \ q=0") case True then show ?thesis by auto next case False have jumpF_add: "jumpF_polyR q p x= jumpF_polyR 1 (q*p) x" when "poly p x=0" "coprime p q" for p q proof (cases "p=0") case True then show ?thesis by auto next case False have "poly q x\0" using that coprime_poly_0 by blast then have "q\0" by auto moreover have "sign_r_pos p x = (0 < poly q x) \ sign_r_pos (q * p) x" using sign_r_pos_mult[OF \q\0\ \p\0\] sign_r_pos_rec[OF \q\0\] \poly q x\0\ by auto ultimately show ?thesis using \poly p x=0\ unfolding jumpF_polyR_coprime[OF \coprime p q\,of x] jumpF_polyR_coprime[of "q*p" 1 x,simplified] by auto qed have False when "poly p x=0" "poly q x=0" using \coprime p q\ that coprime_poly_0 by blast moreover have ?thesis when "poly p x=0" "poly q x\0" proof - have "jumpF_polyR p q x = 0" using jumpF_poly_noroot[OF \poly q x\0\] by auto then show ?thesis using jumpF_add[OF \poly p x=0\ \coprime p q\] by auto qed moreover have ?thesis when "poly p x\0" "poly q x=0" proof - have "jumpF_polyR q p x = 0" using jumpF_poly_noroot[OF \poly p x\0\] by auto then show ?thesis using jumpF_add[OF \poly q x=0\,of p] \coprime p q\ by (simp add: ac_simps) qed moreover have ?thesis when "poly p x\0" "poly q x\0" by (simp add: jumpF_poly_noroot(2) that(1) that(2)) ultimately show ?thesis by auto qed lemma jumpF_polyL_inverse_add: assumes "coprime p q" shows "jumpF_polyL q p x + jumpF_polyL p q x = jumpF_polyL 1 (q*p) x" proof (cases "p=0 \ q=0") case True then show ?thesis by auto next case False have jumpF_add: "jumpF_polyL q p x= jumpF_polyL 1 (q*p) x" when "poly p x=0" "coprime p q" for p q proof (cases "p=0") case True then show ?thesis by auto next case False have "poly q x\0" using that coprime_poly_0 by blast then have "q\0" by auto moreover have "sign_r_pos p x = (0 < poly q x) \ sign_r_pos (q * p) x" using sign_r_pos_mult[OF \q\0\ \p\0\] sign_r_pos_rec[OF \q\0\] \poly q x\0\ by auto moreover have "order x p = order x (q * p)" by (metis \poly q x \ 0\ add_cancel_right_left divisors_zero order_mult order_root) ultimately show ?thesis using \poly p x=0\ unfolding jumpF_polyL_coprime[OF \coprime p q\,of x] jumpF_polyL_coprime[of "q*p" 1 x,simplified] by auto qed have False when "poly p x=0" "poly q x=0" using \coprime p q\ that coprime_poly_0 by blast moreover have ?thesis when "poly p x=0" "poly q x\0" proof - have "jumpF_polyL p q x = 0" using jumpF_poly_noroot[OF \poly q x\0\] by auto then show ?thesis using jumpF_add[OF \poly p x=0\ \coprime p q\] by auto qed moreover have ?thesis when "poly p x\0" "poly q x=0" proof - have "jumpF_polyL q p x = 0" using jumpF_poly_noroot[OF \poly p x\0\] by auto then show ?thesis using jumpF_add[OF \poly q x=0\,of p] \coprime p q\ by (simp add: ac_simps) qed moreover have ?thesis when "poly p x\0" "poly q x\0" by (simp add: jumpF_poly_noroot that(1) that(2)) ultimately show ?thesis by auto qed lemma jumpF_polyL_smult_1: "jumpF_polyL (smult c q) p x = sgn c * jumpF_polyL q p x" proof (cases "c=0") case True then show ?thesis by auto next case False then show ?thesis unfolding jumpF_polyL_def apply (subst jumpF_times[of "\_. c",symmetric]) by auto qed lemma jumpF_polyR_smult_1: "jumpF_polyR (smult c q) p x = sgn c * jumpF_polyR q p x" proof (cases "c=0") case True then show ?thesis by auto next case False then show ?thesis unfolding jumpF_polyR_def using False apply (subst jumpF_times[of "\_. c",symmetric]) by auto qed lemma shows jumpF_polyR_mod:"jumpF_polyR q p x = jumpF_polyR (q mod p) p x" and jumpF_polyL_mod:"jumpF_polyL q p x = jumpF_polyL (q mod p) p x" proof - define f where "f=(\x. poly (q div p) x)" define g where "g=(\x. poly (q mod p) x / poly p x)" have jumpF_eq:"jumpF (\x. poly q x / poly p x) (at y within S) = jumpF g (at y within S)" when "p\0" for y S proof - let ?F = "at y within S" have "\\<^sub>F x in at y within S. poly p x \ 0" using eventually_poly_nz_at_within[OF \p\0\,of y S] . then have "eventually (\x. (poly q x / poly p x) = (f x+ g x)) ?F" proof (rule eventually_mono) fix x assume P: "poly p x \ 0" have "poly q x = poly (q div p * p + q mod p) x" by simp also have "\ = f x * poly p x + poly (q mod p) x" by (simp only: poly_add poly_mult f_def g_def) moreover have "poly (q mod p) x = g x * poly p x" using P by (simp add: g_def) ultimately show "poly q x / poly p x = f x + g x" using P by simp qed then have "jumpF (\x. poly q x / poly p x) ?F = jumpF (\x. f x+ g x) ?F" by (intro jumpF_cong,auto) also have "... = jumpF g ?F" proof - have "(f \ f y) (at y within S)" unfolding f_def by (intro tendsto_intros) from filterlim_tendsto_add_at_bot_iff[OF this,of g] filterlim_tendsto_add_at_top_iff[OF this,of g] show ?thesis unfolding jumpF_def by auto qed finally show ?thesis . qed show "jumpF_polyR q p x = jumpF_polyR (q mod p) p x" apply (cases "p=0") subgoal by auto subgoal using jumpF_eq unfolding g_def jumpF_polyR_def by auto done show "jumpF_polyL q p x = jumpF_polyL (q mod p) p x" apply (cases "p=0") subgoal by auto subgoal using jumpF_eq unfolding g_def jumpF_polyL_def by auto done qed lemma jumpF_poly_top_0[simp]: "jumpF_poly_top 0 p = 0" "jumpF_poly_top q 0 = 0" unfolding jumpF_poly_top_def by auto lemma jumpF_poly_bot_0[simp]: "jumpF_poly_bot 0 p = 0" "jumpF_poly_bot q 0 = 0" unfolding jumpF_poly_bot_def by auto lemma jumpF_poly_top_code: "jumpF_poly_top q p = (if p\0 \ q\0 \ degree q>degree p then if sgn_pos_inf q * sgn_pos_inf p > 0 then 1/2 else -1/2 else 0)" proof (cases "p\0 \ q\0 \ degree q>degree p") case True have ?thesis when "sgn_pos_inf q * sgn_pos_inf p > 0" proof - have "LIM x at_top. poly q x / poly p x :> at_top" using poly_divide_filterlim_at_top[of p q] True that by auto then have "jumpF (\x. poly q x / poly p x) at_top = 1/2" unfolding jumpF_def by auto then show ?thesis unfolding jumpF_poly_top_def using that True by auto qed moreover have ?thesis when "\ sgn_pos_inf q * sgn_pos_inf p > 0" proof - have "LIM x at_top. poly q x / poly p x :> at_bot" using poly_divide_filterlim_at_top[of p q] True that by auto then have "jumpF (\x. poly q x / poly p x) at_top = - 1/2" unfolding jumpF_def by auto then show ?thesis unfolding jumpF_poly_top_def using that True by auto qed ultimately show ?thesis by auto next case False define P where "P= (\ (LIM x at_top. poly q x / poly p x :> at_bot) \ \ (LIM x at_top. poly q x / poly p x :> at_top))" have P when "p=0 \ q=0" unfolding P_def using that by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds) moreover have P when "p\0" "q\0" "degree p> degree q" proof - have "LIM x at_top. poly q x / poly p x :> at 0" using poly_divide_filterlim_at_top[OF that(1,2)] that(3) by auto then show ?thesis unfolding P_def by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds simp:filterlim_at) qed moreover have P when "p\0" "q\0" "degree p = degree q" proof - have "((\x. poly q x / poly p x) \ lead_coeff q / lead_coeff p) at_top" using poly_divide_filterlim_at_top[OF that(1,2)] using that by auto then show ?thesis unfolding P_def by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds) qed ultimately have P using False by fastforce then have "jumpF (\x. poly q x / poly p x) at_top = 0" unfolding jumpF_def P_def by auto then show ?thesis unfolding jumpF_poly_top_def using False by presburger qed lemma jumpF_poly_bot_code: "jumpF_poly_bot q p = (if p\0 \ q\0 \ degree q>degree p then if sgn_neg_inf q * sgn_neg_inf p > 0 then 1/2 else -1/2 else 0)" proof (cases "p\0 \ q\0 \ degree q>degree p") case True have ?thesis when "sgn_neg_inf q * sgn_neg_inf p > 0" proof - have "LIM x at_bot. poly q x / poly p x :> at_top" using poly_divide_filterlim_at_bot[of p q] True that by auto then have "jumpF (\x. poly q x / poly p x) at_bot = 1/2" unfolding jumpF_def by auto then show ?thesis unfolding jumpF_poly_bot_def using that True by auto qed moreover have ?thesis when "\ sgn_neg_inf q * sgn_neg_inf p > 0" proof - have "LIM x at_bot. poly q x / poly p x :> at_bot" using poly_divide_filterlim_at_bot[of p q] True that by auto then have "jumpF (\x. poly q x / poly p x) at_bot = - 1/2" unfolding jumpF_def by auto then show ?thesis unfolding jumpF_poly_bot_def using that True by auto qed ultimately show ?thesis by auto next case False define P where "P= (\ (LIM x at_bot. poly q x / poly p x :> at_bot) \ \ (LIM x at_bot. poly q x / poly p x :> at_top))" have P when "p=0 \ q=0" unfolding P_def using that by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds) moreover have P when "p\0" "q\0" "degree p> degree q" proof - have "LIM x at_bot. poly q x / poly p x :> at 0" using poly_divide_filterlim_at_bot[OF that(1,2)] that(3) by auto then show ?thesis unfolding P_def by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds simp:filterlim_at) qed moreover have P when "p\0" "q\0" "degree p = degree q" proof - have "((\x. poly q x / poly p x) \ lead_coeff q / lead_coeff p) at_bot" using poly_divide_filterlim_at_bot[OF that(1,2)] using that by auto then show ?thesis unfolding P_def by (auto elim!:filterlim_at_bot_nhds filterlim_at_top_nhds) qed ultimately have P using False by fastforce then have "jumpF (\x. poly q x / poly p x) at_bot = 0" unfolding jumpF_def P_def by auto then show ?thesis unfolding jumpF_poly_bot_def using False by presburger qed subsection \The extended Cauchy index on polynomials\ definition cindex_polyE:: "real \ real \ real poly \ real poly \ real" where "cindex_polyE a b q p = jumpF_polyR q p a + cindex_poly a b q p - jumpF_polyL q p b" definition cindex_poly_ubd::"real poly \ real poly \ int" where "cindex_poly_ubd q p = (THE l. (\\<^sub>F r in at_top. cindexE (-r) r (\x. poly q x/poly p x) = of_int l))" lemma cindex_polyE_0[simp]: "cindex_polyE a b 0 p = 0" "cindex_polyE a b q 0 = 0" unfolding cindex_polyE_def by auto lemma cindex_polyE_mult_cancel: fixes p q p'::"real poly" assumes "p'\ 0" shows "cindex_polyE a b (p' * q ) (p' * p) = cindex_polyE a b q p" unfolding cindex_polyE_def using cindex_poly_mult[OF \p'\0\] jumpF_polyL_mult_cancel[OF \p'\0\] jumpF_polyR_mult_cancel[OF \p'\0\] by simp lemma cindexE_eq_cindex_polyE: assumes "ax. poly q x/poly p x) = cindex_polyE a b q p" proof (cases "p=0 \ q=0") case True then show ?thesis by (auto simp add: cindexE_constI) next case False then have "p\0" "q\0" by auto define g where "g=gcd p q" define p' q' where "p'=p div g" and "q' = q div g" define f' where "f'=(\x. poly q' x / poly p' x)" have "g\0" using False g_def by auto have pq_f:"p=g*p'" "q=g*q'" and "coprime p' q'" unfolding g_def p'_def q'_def apply simp_all using False div_gcd_coprime by blast have "cindexE a b (\x. poly q x/poly p x) = cindexE a b (\x. poly q' x/poly p' x)" proof - define f where "f=(\x. poly q x / poly p x)" define f' where "f'=(\x. poly q' x / poly p' x)" have "jumpF f (at_right x) = jumpF f' (at_right x)" for x proof (rule jumpF_cong) obtain ub where "x < ub" "\z. x < z \ z \ ub \ poly g z \ 0" using next_non_root_interval[OF \g\0\,of x] by auto then show "\\<^sub>F x in at_right x. f x = f' x" unfolding eventually_at_right f_def f'_def pq_f apply (intro exI[where x=ub]) by auto qed simp moreover have "jumpF f (at_left x) = jumpF f' (at_left x)" for x proof (rule jumpF_cong) obtain lb where "lb < x" "\z. lb \ z \ z < x \ poly g z \ 0 " using last_non_root_interval[OF \g\0\,of x] by auto then show "\\<^sub>F x in at_left x. f x = f' x" unfolding eventually_at_left f_def f'_def pq_f apply (intro exI[where x=lb]) by auto qed simp ultimately show ?thesis unfolding cindexE_def apply (fold f_def f'_def) by auto qed also have "... = jumpF f' (at_right a) + real_of_int (cindex a b f') - jumpF f' (at_left b)" unfolding f'_def apply (rule cindex_eq_cindexE_divide) subgoal using \a . subgoal using False poly_roots_finite pq_f(1) pq_f(2) by fastforce subgoal using \coprime p' q'\ poly_gcd_iff by force subgoal by (auto intro:continuous_intros) subgoal by (auto intro:continuous_intros) done also have "... = cindex_polyE a b q' p'" using cindex_eq_cindex_poly unfolding cindex_polyE_def jumpF_polyR_def jumpF_polyL_def f'_def by auto also have "... = cindex_polyE a b q p" using cindex_polyE_mult_cancel[OF \g\0\] unfolding pq_f by auto finally show ?thesis . qed lemma cindex_polyE_cross: fixes p::"real poly" and a b::real assumes "a0" and noroot:"{x. a< x\ x< b \ poly p x=0 } = {}" proof - have "cindex_polyE a b 1 p = jumpF_polyR 1 p a - jumpF_polyL 1 p b" proof - have "cindex_poly a b 1 p = 0" unfolding cindex_poly_def apply (rule sum.neutral) using that by auto then show ?thesis unfolding cindex_polyE_def by auto qed also have "... = cross_alt 1 p a b / 2" proof - define f where "f = (\x. 1 / poly p x)" define ja where "ja = jumpF f (at_right a)" define jb where "jb = jumpF f (at_left b)" define right where "right = (\R. R ja (0::real) \ (continuous (at_right a) f \ R (poly p a) 0))" define left where "left = (\R. R jb (0::real) \ (continuous (at_left b) f \ R (poly p b) 0))" note ja_alt=jumpF_polyR_coprime[of p 1 a,unfolded jumpF_polyR_def,simplified,folded f_def ja_def] note jb_alt=jumpF_polyL_coprime[of p 1 b,unfolded jumpF_polyL_def,simplified,folded f_def jb_def] have [simp]:"0 < ja \ jumpF_polyR 1 p a = 1/2" "0 > ja \ jumpF_polyR 1 p a = -1/2" "0 < jb \ jumpF_polyL 1 p b = 1/2" "0 > jb \ jumpF_polyL 1 p b = -1/2" unfolding ja_def jb_def jumpF_polyR_def jumpF_polyL_def f_def jumpF_def by auto have [simp]: "poly p a \0 \ continuous (at_right a) f" "poly p b \0 \ continuous (at_left b) f" unfolding f_def by (auto intro!: continuous_intros ) have not_right_left: False when "(right greater \ left less \ right less \ left greater)" proof - have [simp]: "f a > 0 \ poly p a > 0" "f a < 0 \ poly p a < 0" "f b > 0 \ poly p b > 0" "f b < 0 \ poly p b < 0" unfolding f_def by auto have "continuous_on {a<..x>a. x < b \ f x = 0" apply (elim jumpF_IVT[OF \a,of f]) using that unfolding right_def left_def by (fold ja_def jb_def,auto) then show False using noroot using f_def by auto qed have ?thesis when "poly p a>0 \ poly p b>0 \ poly p a<0 \ poly p b<0" using that jumpF_poly_noroot unfolding cross_alt_def by auto moreover have False when "poly p a>0 \ poly p b<0 \ poly p a<0 \ poly p b>0" apply (rule not_right_left) unfolding right_def left_def using that by auto moreover have ?thesis when "poly p a=0" "poly p b>0 \ poly p b <0" proof - have "ja>0 \ ja < 0" using ja_alt \p\0\ \poly p a=0\ by argo moreover have False when "ja > 0 \ poly p b<0 \ ja < 0 \ poly p b>0" apply (rule not_right_left) unfolding right_def left_def using that by fastforce moreover have ?thesis when "ja >0 \ poly p b>0 \ ja < 0 \ poly p b<0" using that jumpF_poly_noroot \poly p a=0\ unfolding cross_alt_def by auto ultimately show ?thesis using that jumpF_poly_noroot unfolding cross_alt_def by auto qed moreover have ?thesis when "poly p b=0" "poly p a>0 \ poly p a <0" proof - have "jb>0 \ jb < 0" using jb_alt \p\0\ \poly p b=0\ by argo moreover have False when "jb > 0 \ poly p a<0 \ jb < 0 \ poly p a>0" apply (rule not_right_left) unfolding right_def left_def using that by fastforce moreover have ?thesis when "jb >0 \ poly p a>0 \ jb < 0 \ poly p a<0" using that jumpF_poly_noroot \poly p b=0\ unfolding cross_alt_def by auto ultimately show ?thesis using that jumpF_poly_noroot unfolding cross_alt_def by auto qed moreover have ?thesis when "poly p a=0" "poly p b=0" proof - have "jb>0 \ jb < 0" using jb_alt \p\0\ \poly p b=0\ by argo moreover have "ja>0 \ ja < 0" using ja_alt \p\0\ \poly p a=0\ by argo moreover have False when "ja>0 \ jb<0 \ ja<0 \ jb>0" apply (rule not_right_left) unfolding right_def left_def using that by fastforce moreover have ?thesis when "ja>0 \ jb>0 \ ja<0 \ jb<0" using that jumpF_poly_noroot \poly p b=0\ \poly p a=0\ unfolding cross_alt_def by auto ultimately show ?thesis by blast qed ultimately show ?thesis by argo qed finally show ?thesis . qed moreover have ?case when "p\0" and no_empty:"{x. a< x\ x< b \ poly p x=0 } \ {}" proof - define roots where "roots\{x. a< x\ x< b \ poly p x=0 }" have "finite roots" unfolding roots_def using poly_roots_finite[OF \p\0\] by auto define max_r where "max_r\Max roots" hence "poly p max_r=0" and "afinite roots\] no_empty unfolding roots_def by auto define max_rp where "max_rp\[:-max_r,1:]^order max_r p" then obtain p' where p'_def:"p=p'*max_rp" and "\ [:-max_r,1:] dvd p'" by (metis \p\0\ mult.commute order_decomp) hence "p'\0" and "max_rp\0" and max_r_nz:"poly p' max_r\0"(*and "poly p' a\0" and "poly p' b\0" *) (*and "poly max_rp a\0" and "poly max_rp b\0"*) using \p\0\ by (auto simp add: dvd_iff_poly_eq_0) define max_r_sign where "max_r_sign\if odd(order max_r p) then -1 else 1::int" define roots' where "roots'\{x. a< x\ x< b \ poly p' x=0}" have "cindex_polyE a b 1 p = jumpF_polyR 1 p a + (\x\roots. jump_poly 1 p x) - jumpF_polyL 1 p b" unfolding cindex_polyE_def cindex_poly_def roots_def by (simp,meson) also have "... = max_r_sign * cindex_poly a b 1 p' + jump_poly 1 p max_r + max_r_sign * jumpF_polyR 1 p' a - jumpF_polyL 1 p' b" proof - have "(\x\roots. jump_poly 1 p x) = max_r_sign * cindex_poly a b 1 p' + jump_poly 1 p max_r" proof - have "(\x\roots. jump_poly 1 p x)= (\x\roots'. jump_poly 1 p x)+ jump_poly 1 p max_r" proof - have "roots = insert max_r roots'" unfolding roots_def roots'_def p'_def using \poly p max_r=0\ \a \max_r \p\0\ order_root apply (subst max_rp_def) by auto moreover have "finite roots'" unfolding roots'_def using poly_roots_finite[OF \p'\0\] by auto moreover have "max_r \ roots'" unfolding roots'_def using max_r_nz by auto ultimately show ?thesis using sum.insert[of roots' max_r] by auto qed moreover have "(\x\roots'. jump_poly 1 p x) = max_r_sign * cindex_poly a b 1 p'" proof - have "(\x\roots'. jump_poly 1 p x) = (\x\roots'. max_r_sign * jump_poly 1 p' x)" proof (rule sum.cong,rule refl) fix x assume "x \ roots'" hence "x\max_r" using max_r_nz unfolding roots'_def by auto hence "poly max_rp x\0" using poly_power_n_eq unfolding max_rp_def by auto hence "order x max_rp=0" by (metis order_root) moreover have "jump_poly 1 max_rp x=0" using \poly max_rp x\0\ by (metis jump_poly_not_root) moreover have "x\roots" using \x \ roots'\ unfolding roots_def roots'_def p'_def by auto hence "xfinite roots\,of x] \x\max_r\ by (fold max_r_def,auto) hence "sign (poly max_rp x) = max_r_sign" using \poly max_rp x \ 0\ unfolding max_r_sign_def max_rp_def sign_def by (subst poly_power,simp add:linorder_class.not_less zero_less_power_eq) ultimately show "jump_poly 1 p x = max_r_sign * jump_poly 1 p' x" using jump_poly_1_mult[of p' x max_rp] unfolding p'_def by (simp add: \poly max_rp x \ 0\) qed also have "... = max_r_sign * (\x\roots'. jump_poly 1 p' x)" by (simp add: sum_distrib_left) also have "... = max_r_sign * cindex_poly a b 1 p'" unfolding cindex_poly_def roots'_def by meson finally show ?thesis . qed ultimately show ?thesis by simp qed moreover have "jumpF_polyR 1 p a = max_r_sign * jumpF_polyR 1 p' a" proof - define f where "f = (\x. 1 / poly max_rp x)" define g where "g = (\x. 1 / poly p' x)" have "jumpF_polyR 1 p a = jumpF (\x. f x * g x) (at_right a)" unfolding jumpF_polyR_def f_def g_def p'_def by (auto simp add:field_simps) also have "... = sgn (f a) * jumpF g (at_right a)" proof (rule jumpF_times) have [simp]: "poly max_rp a \0" unfolding max_rp_def using \max_r>a\ by auto show "(f \ f a) (at_right a)" "f a \ 0" unfolding f_def by (auto intro:tendsto_intros) qed auto also have "... = max_r_sign * jumpF_polyR 1 p' a" proof - have "sgn (f a) = max_r_sign" unfolding max_r_sign_def f_def max_rp_def using \a by (auto simp add:sgn_power) then show ?thesis unfolding jumpF_polyR_def g_def by auto qed finally show ?thesis . qed moreover have "jumpF_polyL 1 p b = jumpF_polyL 1 p' b" proof - define f where "f = (\x. 1 / poly max_rp x)" define g where "g = (\x. 1 / poly p' x)" have "jumpF_polyL 1 p b = jumpF (\x. f x * g x) (at_left b)" unfolding jumpF_polyL_def f_def g_def p'_def by (auto simp add:field_simps) also have "... = sgn (f b) * jumpF g (at_left b)" proof (rule jumpF_times) have [simp]: "poly max_rp b \0" unfolding max_rp_def using \max_r by auto show "(f \ f b) (at_left b)" "f b \ 0" unfolding f_def by (auto intro:tendsto_intros) qed auto also have "... = jumpF_polyL 1 p' b" proof - have "sgn (f b) = 1" unfolding max_r_sign_def f_def max_rp_def using \b>max_r\ by (auto simp add:sgn_power) then show ?thesis unfolding jumpF_polyL_def g_def by auto qed finally show ?thesis . qed ultimately show ?thesis by auto qed also have "... = max_r_sign * cindex_polyE a b 1 p' + jump_poly 1 p max_r + (max_r_sign - 1) * jumpF_polyL 1 p' b" unfolding cindex_polyE_def roots'_def by (auto simp add:algebra_simps) also have "... = max_r_sign * cross_alt 1 p' a b / 2 + jump_poly 1 p max_r + (max_r_sign - 1) * jumpF_polyL 1 p' b" proof - have "degree max_rp>0" unfolding max_rp_def degree_linear_power using \poly p max_r=0\ order_root \p\0\ by blast then have "degree p'p'\0\ \max_rp\0\] by auto from induct[rule_format, OF this] have "cindex_polyE a b 1 p' = real_of_int (cross_alt 1 p' a b) / 2" by auto then show ?thesis by auto qed also have "... = real_of_int (cross_alt 1 p a b) / 2" proof - have sjump_p:"jump_poly 1 p max_r = (if odd (order max_r p) then sign (poly p' max_r) else 0)" proof - note max_r_nz moreover then have "poly max_rp max_r=0" using \poly p max_r = 0\ p'_def by auto ultimately have "jump_poly 1 p max_r = sign (poly p' max_r) * jump_poly 1 max_rp max_r" unfolding p'_def using jump_poly_1_mult[of p' max_r max_rp] by auto also have "... = (if odd (order max_r max_rp) then sign (poly p' max_r) else 0)" proof - have "sign_r_pos max_rp max_r" unfolding max_rp_def using sign_r_pos_power by auto then show ?thesis using \max_rp\0\ unfolding jump_poly_def by auto qed also have "... = (if odd (order max_r p) then sign (poly p' max_r) else 0)" proof - have "order max_r p'=0" by (simp add: \poly p' max_r \ 0\ order_0I) then have "order max_r max_rp = order max_r p" unfolding p'_def using \p'\0\ \max_rp\0\ apply (subst order_mult) by auto then show ?thesis by auto qed finally show ?thesis . qed have ?thesis when "even (order max_r p)" proof - have "sign (poly p x) = sign (poly p' x)" when "x\max_r" for x proof - have "sign (poly max_rp x) = 1" unfolding max_rp_def using \even (order max_r p)\ that apply (simp add:sign_power ) by (simp add: Sturm_Tarski.sign_def) then show ?thesis unfolding p'_def by (simp add:sign_times) qed from this[of a] this[of b] \a \max_r have "cross_alt 1 p' a b = cross_alt 1 p a b" unfolding cross_alt_def by auto then show ?thesis using that unfolding max_r_sign_def sjump_p by auto qed moreover have ?thesis when "odd (order max_r p)" proof - let ?thesis2 = "sign (poly p' max_r) * 2 - cross_alt 1 p' a b - 4 * jumpF_polyL 1 p' b = cross_alt 1 p a b" have ?thesis2 when "poly p' b=0" proof - have "jumpF_polyL 1 p' b = 1/2 \ jumpF_polyL 1 p' b=-1/2" using jumpF_polyL_coprime[of p' 1 b,simplified] \p'\0\ \poly p' b=0\ by auto moreover have "poly p' max_r>0 \ poly p' max_r<0" using max_r_nz by auto moreover have False when "poly p' max_r>0 \ jumpF_polyL 1 p' b=-1/2 \ poly p' max_r<0 \ jumpF_polyL 1 p' b=1/2" proof - define f where "f= (\x. 1/ poly p' x)" have noroots:"poly p' x\0" when "x\{max_r<.. poly p' x \ 0" then have "poly p x =0" unfolding p'_def by auto then have "x\roots" unfolding roots_def using that \a by auto then have "x\max_r" using Max_ge[OF \finite roots\] unfolding max_r_def by auto moreover have "x>max_r" using that by auto ultimately show False by auto qed have "continuous_on {max_r<..0 \ jumpF f (at_left b)<0 \ f max_r<0 \ jumpF f (at_left b)>0" using that unfolding f_def jumpF_polyL_def by auto ultimately have "\x>max_r. x < b \ f x = 0" apply (intro jumpF_IVT[OF \max_r]) by auto then show False using noroots unfolding f_def by auto qed moreover have ?thesis when "poly p' max_r>0 \ jumpF_polyL 1 p' b=1/2 \ poly p' max_r<0 \ jumpF_polyL 1 p' b=-1/2" proof - have "poly max_rp a < 0" "poly max_rp b>0" unfolding max_rp_def using \odd (order max_r p)\ \a \max_r by (simp_all add:zero_less_power_eq) then have "cross_alt 1 p a b = - cross_alt 1 p' a b" unfolding cross_alt_def p'_def using \poly p' b=0\ apply (simp add:sign_times) by (simp add: Sturm_Tarski.sign_def) with that show ?thesis by auto qed ultimately show ?thesis by blast qed moreover have ?thesis2 when "poly p' b\0" proof - have [simp]:"jumpF_polyL 1 p' b = 0" using jumpF_polyL_coprime[of p' 1 b,simplified] \poly p' b\0\ by auto have [simp]:"poly max_rp a < 0" "poly max_rp b>0" unfolding max_rp_def using \odd (order max_r p)\ \a \max_r by (simp_all add:zero_less_power_eq) have "poly p' b>0 \ poly p' b<0" using \poly p' b\0\ by auto moreover have "poly p' max_r>0 \ poly p' max_r<0" using max_r_nz by auto moreover have ?thesis when "poly p' b>0 \ poly p' max_r>0 " using that unfolding cross_alt_def p'_def apply (simp add:sign_times) by (simp add: Sturm_Tarski.sign_def) moreover have ?thesis when "poly p' b<0 \ poly p' max_r<0" using that unfolding cross_alt_def p'_def apply (simp add:sign_times) by (simp add: Sturm_Tarski.sign_def) moreover have False when "poly p' b>0 \ poly p' max_r<0 \ poly p' b<0 \ poly p' max_r>0" proof - have "\x>max_r. x < b \ poly p' x = 0" apply (rule poly_IVT[OF \max_r,of p']) using that mult_less_0_iff by blast then obtain x where "max_rroots" using \a unfolding roots_def by auto then have "x\max_r" unfolding max_r_def using Max_ge[OF \finite roots\] by auto then show False using \max_r by auto qed ultimately show ?thesis by blast qed ultimately have ?thesis2 by auto then show ?thesis unfolding max_r_sign_def sjump_p using that by simp qed ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?case by fast qed lemma cindex_polyE_inverse_add: fixes p q::"real poly" assumes cp:"coprime p q" shows "cindex_polyE a b q p + cindex_polyE a b p q=cindex_polyE a b 1 (q*p)" unfolding cindex_polyE_def using cindex_poly_inverse_add[OF cp,symmetric] jumpF_polyR_inverse_add[OF cp,symmetric] jumpF_polyL_inverse_add[OF cp,symmetric] by auto lemma cindex_polyE_inverse_add_cross: fixes p q::"real poly" assumes "a < b" "coprime p q" shows "cindex_polyE a b q p + cindex_polyE a b p q = cross_alt p q a b / 2" apply (subst cindex_polyE_inverse_add[OF \coprime p q\]) apply (subst cindex_polyE_cross[OF \a]) apply (subst mult.commute) apply (subst cross_alt_clear_n[OF \coprime p q\]) by simp lemma cindex_polyE_smult_1: fixes p q::"real poly" and c::real shows "cindex_polyE a b (smult c q) p = (sgn c) * cindex_polyE a b q p" unfolding cindex_polyE_def jumpF_polyL_smult_1 jumpF_polyR_smult_1 cindex_poly_smult_1 by (auto simp add:sgn_sign_eq[symmetric] algebra_simps) lemma cindex_polyE_mod: fixes p q::"real poly" shows "cindex_polyE a b q p = cindex_polyE a b (q mod p) p" unfolding cindex_polyE_def apply (subst cindex_poly_mod) apply (subst jumpF_polyR_mod) apply (subst jumpF_polyL_mod) by simp lemma cindex_polyE_rec: fixes p q::"real poly" assumes "a < b" "coprime p q" shows "cindex_polyE a b q p = cross_alt q p a b/2 + cindex_polyE a b (- (p mod q)) q" proof - note cindex_polyE_inverse_add_cross[OF assms] moreover have "cindex_polyE a b (- (p mod q)) q = - cindex_polyE a b p q" using cindex_polyE_mod cindex_polyE_smult_1[of a b "-1"] by auto ultimately show ?thesis by (auto simp add:field_simps cross_alt_poly_commute) qed lemma cindex_polyE_changes_alt_itv_mods: assumes "acoprime p q\ proof (induct "smods p q" arbitrary:p q) case Nil then have "p=0" by (metis smods_nil_eq) then show ?case by (simp add:changes_alt_itv_smods_def changes_alt_poly_at_def) next case (Cons x xs) then have "p\0" by auto have ?case when "q=0" using that by (simp add:changes_alt_itv_smods_def changes_alt_poly_at_def) moreover have ?case when "q\0" proof - define r where "r\- (p mod q)" obtain ps where ps:"smods p q=p#q#ps" "smods q r=q#ps" and "xs=q#ps" unfolding r_def using \q\0\ \p\0\ \x # xs = smods p q\ by (metis list.inject smods.simps) from Cons.prems \q \ 0\ have "coprime q r" by (simp add: r_def ac_simps) then have "cindex_polyE a b r q = real_of_int (changes_alt_itv_smods a b q r) / 2" apply (rule_tac Cons.hyps(1)) using ps \xs=q#ps\ by simp_all moreover have "changes_alt_itv_smods a b p q = cross_alt p q a b + changes_alt_itv_smods a b q r" using changes_alt_itv_smods_rec[OF \a \coprime p q\,folded r_def] . moreover have "cindex_polyE a b q p = real_of_int (cross_alt q p a b) / 2 + cindex_polyE a b r q" using cindex_polyE_rec[OF \a \coprime p q\,folded r_def] . ultimately show ?case by (auto simp add:field_simps cross_alt_poly_commute) qed ultimately show ?case by blast qed lemma cindex_poly_ubd_eventually: shows "\\<^sub>F r in at_top. cindexE (-r) r (\x. poly q x/poly p x) = of_int (cindex_poly_ubd q p)" proof - define f where "f=(\x. poly q x/poly p x)" obtain R where R_def: "R>0" "proots p \ {-R<..0" proof (cases "p=0") case True then show ?thesis using that[of 1] by auto next case False then have "finite (proots p)" by auto from finite_ball_include[OF this,of 0] obtain r where "r>0" and r_ball:"proots p \ ball 0 r" by auto have "proots p \ {-r<.. proots p" then have "x\ball 0 r" using r_ball by auto then have "abs x {- r<..r>0\ by auto qed then show ?thesis using that[of r] False \r>0\ by auto qed define l where "l=(if p=0 then 0 else cindex_poly (-R) R q p)" define P where "P=(\l. (\\<^sub>F r in at_top. cindexE (-r) r f = of_int l))" have "P l " proof (cases "p=0 ") case True then show ?thesis unfolding P_def f_def l_def using True by (auto intro!: eventuallyI cindexE_constI) next case False have "P l" unfolding P_def proof (rule eventually_at_top_linorderI[of R]) fix r assume "R \ r" then have "cindexE (- r) r f = cindex_polyE (-r) r q p" unfolding f_def using R_def[OF \p\0\] by (auto intro: cindexE_eq_cindex_polyE) also have "... = of_int (cindex_poly (- r) r q p)" proof - have "jumpF_polyR q p (- r) = 0" apply (rule jumpF_poly_noroot) using \R\r\ R_def[OF \p\0\] by auto moreover have "jumpF_polyL q p r = 0" apply (rule jumpF_poly_noroot) using \R\r\ R_def[OF \p\0\] by auto ultimately show ?thesis unfolding cindex_polyE_def by auto qed also have "... = of_int (cindex_poly (- R) R q p)" proof - define rs where "rs={x. poly p x = 0 \ - r < x \ x < r}" define Rs where "Rs={x. poly p x = 0 \ - R < x \ x < R}" have "rs=Rs" using R_def[OF \p\0\] \R\r\ unfolding rs_def Rs_def by force then show ?thesis unfolding cindex_poly_def by (fold rs_def Rs_def,auto) qed also have "... = of_int l" unfolding l_def using False by auto finally show "cindexE (- r) r f = real_of_int l" . qed then show ?thesis unfolding P_def by auto qed moreover have "x=l" when "P x" for x proof - have "\\<^sub>F r in at_top. cindexE (- r) r f = real_of_int x" "\\<^sub>F r in at_top. cindexE (- r) r f = real_of_int l" using \P x\ \P l\ unfolding P_def by auto from eventually_conj[OF this] have "\\<^sub>F r::real in at_top. real_of_int x = real_of_int l" by (elim eventually_mono,auto) then have "real_of_int x = real_of_int l" by auto then show ?thesis by simp qed ultimately have "P (THE x. P x)" using theI[of P l] by blast then show ?thesis unfolding P_def f_def cindex_poly_ubd_def by auto qed lemma cindex_poly_ubd_0: assumes "p=0 \ q=0" shows "cindex_poly_ubd q p = 0" proof - have "\\<^sub>F r in at_top. cindexE (-r) r (\x. poly q x/poly p x) = 0" apply (rule eventuallyI) using assms by (auto intro:cindexE_constI) from eventually_conj[OF this cindex_poly_ubd_eventually[of q p]] have "\\<^sub>F r::real in at_top. (cindex_poly_ubd q p) = (0::int)" apply (elim eventually_mono) by auto then show ?thesis by auto qed lemma cindex_poly_ubd_code: shows "cindex_poly_ubd q p = changes_R_smods p q" proof (cases "p=0") case True then show ?thesis using cindex_poly_ubd_0 by auto next case False define ps where "ps\smods p q" have "p\set ps" using ps_def \p\0\ by auto obtain lb where lb:"\p\set ps. \x. poly p x=0 \ x>lb" and lb_sgn:"\x\lb. \p\set ps. sgn (poly p x) = sgn_neg_inf p" and "lb<0" using root_list_lb[OF no_0_in_smods,of p q,folded ps_def] by auto obtain ub where ub:"\p\set ps. \x. poly p x=0 \ xx\ub. \p\set ps. sgn (poly p x) = sgn_pos_inf p" and "ub>0" using root_list_ub[OF no_0_in_smods,of p q,folded ps_def] by auto define f where "f=(\t. poly q t/ poly p t)" define P where "P=(\l. (\\<^sub>F r in at_top. cindexE (-r) r f = of_int l))" have "P (changes_R_smods p q)" unfolding P_def proof (rule eventually_at_top_linorderI[of "max \lb\ \ub\ + 1"]) fix r assume r_asm:"r\max \lb\ \ub\ + 1" have "cindexE (- r) r f = cindex_polyE (-r) r q p" unfolding f_def using r_asm by (auto intro: cindexE_eq_cindex_polyE) also have "... = of_int (cindex_poly (- r) r q p)" proof - have "jumpF_polyR q p (- r) = 0" apply (rule jumpF_poly_noroot) using r_asm lb[rule_format,OF \p\set ps\,of "-r"] by linarith moreover have "jumpF_polyL q p r = 0" apply (rule jumpF_poly_noroot) using r_asm ub[rule_format,OF \p\set ps\,of "r"] by linarith ultimately show ?thesis unfolding cindex_polyE_def by auto qed also have "... = of_int (changes_itv_smods (- r) r p q)" apply (rule cindex_poly_changes_itv_mods[THEN arg_cong]) using r_asm lb[rule_format,OF \p\set ps\,of "-r"] ub[rule_format,OF \p\set ps\,of "r"] by linarith+ also have "... = of_int (changes_R_smods p q)" proof - have "map (sgn \ (\p. poly p (-r))) ps = map sgn_neg_inf ps" and "map (sgn \ (\p. poly p r)) ps = map sgn_pos_inf ps" using lb_sgn[THEN spec,of "-r",simplified] ub_sgn[THEN spec,of r,simplified] r_asm by auto hence "changes_poly_at ps (-r)=changes_poly_neg_inf ps \ changes_poly_at ps r=changes_poly_pos_inf ps" unfolding changes_poly_neg_inf_def changes_poly_at_def changes_poly_pos_inf_def by (subst (1 3) changes_map_sgn_eq,metis map_map) thus ?thesis unfolding changes_R_smods_def changes_itv_smods_def ps_def by metis qed finally show "cindexE (- r) r f = of_int (changes_R_smods p q)" . qed moreover have "x = changes_R_smods p q" when "P x" for x proof - have "\\<^sub>F r in at_top. cindexE (- r) r f = real_of_int (changes_R_smods p q)" "\\<^sub>F r in at_top. cindexE (- r) r f = real_of_int x" using \P (changes_R_smods p q)\ \P x\ unfolding P_def by auto from eventually_conj[OF this] have "\\<^sub>F (r::real) in at_top. of_int x = of_int (changes_R_smods p q)" by (elim eventually_mono,auto) then have "of_int x = of_int (changes_R_smods p q)" using eventually_const_iff by auto then show ?thesis using of_int_eq_iff by blast qed ultimately have "(THE x. P x) = changes_R_smods p q" using the_equality[of P "changes_R_smods p q"] by blast then show ?thesis unfolding cindex_poly_ubd_def P_def f_def by auto qed lemma cindexE_ubd_poly: "cindexE_ubd (\x. poly q x/poly p x) = cindex_poly_ubd q p" proof (cases "p=0") case True then show ?thesis using cindex_poly_ubd_0 unfolding cindexE_ubd_def by auto next case False define mx mn where "mx = Max {x. poly p x = 0}" and "mn = Min {x. poly p x=0}" define rr where "rr= 1+ (max \mx\ \mn\)" have rr:"-rr < x \ x< rr" when "poly p x = 0 " for x proof - have "finite {x. poly p x = 0}" using \p\0\ poly_roots_finite by blast then have "mn \ x" "x\mx" using Max_ge Min_le that unfolding mn_def mx_def by simp_all then show ?thesis unfolding rr_def by auto qed define f where "f=(\x. poly q x / poly p x)" have "\\<^sub>F r in at_top. cindexE (- r) r f = cindexE_ubd f" proof (rule eventually_at_top_linorderI[of rr]) fix r assume "r\rr" define R1 R2 where "R1={x. jumpF f (at_right x) \ 0 \ - r \ x \ x < r}" and "R2 = {x. jumpF f (at_right x) \ 0}" define L1 L2 where "L1={x. jumpF f (at_left x) \ 0 \ - r < x \ x \ r}" and "L2={x. jumpF f (at_left x) \ 0}" have "R1=R2" proof - have "jumpF f (at_right x) = 0" when "\ (- r \ x \ x < r)" for x proof - have "jumpF f (at_right x) = jumpF_polyR q p x" unfolding f_def jumpF_polyR_def by simp also have "... = 0" apply (rule jumpF_poly_noroot) using that \r\rr\ by (auto dest:rr) finally show ?thesis . qed then show ?thesis unfolding R1_def R2_def by blast qed moreover have "L1=L2" proof - have "jumpF f (at_left x) = 0" when "\ (- r < x \ x \ r)" for x proof - have "jumpF f (at_left x) = jumpF_polyL q p x" unfolding f_def jumpF_polyL_def by simp also have "... = 0" apply (rule jumpF_poly_noroot) using that \r\rr\ by (auto dest:rr) finally show ?thesis . qed then show ?thesis unfolding L1_def L2_def by blast qed ultimately show "cindexE (- r) r f = cindexE_ubd f" unfolding cindexE_def cindexE_ubd_def apply (fold R1_def R2_def L1_def L2_def) by auto qed moreover have "\\<^sub>F r in at_top. cindexE (- r) r f = cindex_poly_ubd q p" using cindex_poly_ubd_eventually unfolding f_def by auto ultimately have "\\<^sub>F r in at_top. cindexE (- r) r f = cindexE_ubd f \ cindexE (- r) r f = cindex_poly_ubd q p" using eventually_conj by auto then have "\\<^sub>F (r::real) in at_top. cindexE_ubd f = cindex_poly_ubd q p" by (elim eventually_mono) auto then show ?thesis unfolding f_def by auto qed end diff --git a/thys/DiscretePricing/CRR_Model.thy b/thys/DiscretePricing/CRR_Model.thy --- a/thys/DiscretePricing/CRR_Model.thy +++ b/thys/DiscretePricing/CRR_Model.thy @@ -1,3271 +1,3252 @@ (* Title: CRR_Model.thy Author: Mnacho Echenim, Univ. Grenoble Alpes *) section \The Cox Ross Rubinstein model\ text \This section defines the Cox-Ross-Rubinstein model of a financial market, and charcterizes a risk-neutral probability space for this market. This, together with the proof that every derivative is attainable, permits to obtain a formula to explicitely compute the fair price of any derivative.\ theory CRR_Model imports Fair_Price begin locale CRR_hyps = prob_grw + rsk_free_asset + fixes stk assumes stocks: "stocks Mkt = {stk, risk_free_asset}" and stk_price: "prices Mkt stk = geom_proc" and S0_positive: "0 < init" and down_positive: "0 < d" and down_lt_up: "d < u" and psgt: "0 < p" and pslt: "p < 1" locale CRR_market = CRR_hyps + fixes G assumes stock_filtration:"G = stoch_proc_filt M geom_proc borel" subsection \Preliminary results on the market\ lemma (in CRR_market) case_asset: assumes "asset \ stocks Mkt" shows "asset = stk \ asset = risk_free_asset" proof (rule ccontr) assume "\ (asset = stk \ asset = risk_free_asset)" hence "asset \ stk \ asset \ risk_free_asset" by simp moreover have "asset \ {stk, risk_free_asset}" using assms stocks by simp ultimately show False by auto qed lemma (in CRR_market) assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows bernoulli_gen_filtration: "filtration N G" and bernoulli_sigma_finite: "\n. sigma_finite_subalgebra N (G n)" proof - show "filtration N G" proof - have "disc_filtr M (stoch_proc_filt M geom_proc borel)" proof (rule stoch_proc_filt_disc_filtr) fix i show "random_variable borel (geom_proc i)" by (simp add: geom_rand_walk_borel_measurable) qed hence "filtration M G" using stock_filtration by (simp add: filtration_def disc_filtr_def) have "filt_equiv nat_filtration M N" using pslt psgt by (simp add: assms bernoulli_stream_equiv) hence "sets N = sets M" unfolding filt_equiv_def by simp thus ?thesis unfolding filtration_def by (metis filtration_def \Filtration.filtration M G\ sets_eq_imp_space_eq subalgebra_def) qed show "\n. sigma_finite_subalgebra N (G n)" using assms unfolding subalgebra_def using filtration_def subalgebra_sigma_finite by (metis \Filtration.filtration N G\ bernoulli_stream_def prob_space.prob_space_stream_space prob_space.subalgebra_sigma_finite prob_space_measure_pmf) qed sublocale CRR_market \ rfr_disc_equity_market _ G proof (unfold_locales) show "disc_filtr M G \ sets (G \) = {{}, space M}" proof show "sets (G \) = {{}, space M}" using infinite_cts_filtration.stoch_proc_filt_triv_init stock_filtration geometric_process geom_rand_walk_borel_adapted by (meson infinite_coin_toss_space_axioms infinite_cts_filtration_axioms.intro infinite_cts_filtration_def init_triv_filt_def) show "disc_filtr M G" by (metis Filtration.filtration_def bernoulli bernoulli_gen_filtration disc_filtr_def psgt pslt) qed show "\asset\stocks Mkt. borel_adapt_stoch_proc G (prices Mkt asset)" proof - have "borel_adapt_stoch_proc G (prices Mkt stk)" using stk_price stock_filtration stoch_proc_filt_adapt by (simp add: stoch_proc_filt_adapt geom_rand_walk_borel_measurable) moreover have "borel_adapt_stoch_proc G (prices Mkt risk_free_asset)" using \disc_filtr M G \ sets (G \) = {{}, space M}\ disc_filtr_prob_space.disc_rfr_proc_borel_adapted disc_filtr_prob_space.intro disc_filtr_prob_space_axioms.intro prob_space_axioms rf_price by fastforce moreover have "disc_filtr_prob_space M G" proof (unfold_locales) show "disc_filtr M G" by (simp add: \disc_filtr M G \ sets (G \) = {{}, space M}\) qed ultimately show ?thesis using stocks by force qed qed lemma (in CRR_market) two_stocks: shows "stk \ risk_free_asset" proof (rule ccontr) assume "\stk \ risk_free_asset" hence "disc_rfr_proc r = prices Mkt stk" using rf_price by simp also have "... = geom_proc" using stk_price by simp finally have eqf: "disc_rfr_proc r = geom_proc" . hence "\w. disc_rfr_proc r 0 w = geom_proc 0 w" by simp hence "1 = init" using geometric_process by simp have eqfs: "\w. disc_rfr_proc r (Suc 0) w = geom_proc (Suc 0) w" using eqf by simp hence "disc_rfr_proc r (Suc 0) (sconst True) = geom_proc (Suc 0) (sconst True)" by simp hence "1+r = u" using geometric_process \1 = init\ by simp have "disc_rfr_proc r (Suc 0) (sconst False) = geom_proc (Suc 0) (sconst False)" using eqfs by simp hence "1+r = d" using geometric_process \1 = init\ by simp show False using \1+r = u\ \1+r = d\ down_lt_up by simp qed lemma (in CRR_market) stock_pf_vp_expand: assumes "stock_portfolio Mkt pf" shows "val_process Mkt pf n w = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" proof - have "val_process Mkt pf n w =(sum (\x. ((prices Mkt) x n w) * (pf x (Suc n) w)) (stocks Mkt))" proof (rule subset_val_process') show "finite (stocks Mkt)" using stocks by auto show "support_set pf \ stocks Mkt" using assms unfolding stock_portfolio_def by simp qed also have "... = (\x\ {stk, risk_free_asset}. ((prices Mkt) x n w) * (pf x (Suc n) w))" using stocks by simp also have "... = prices Mkt stk n w * pf stk (Suc n) w + (\ x\ {risk_free_asset}. ((prices Mkt) x n w) * (pf x (Suc n) w))" by (simp add:two_stocks) also have "... = prices Mkt stk n w * pf stk (Suc n) w + prices Mkt risk_free_asset n w * pf risk_free_asset (Suc n) w" by simp also have "... = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using rf_price stk_price by simp finally show ?thesis . qed lemma (in CRR_market) stock_pf_uvp_expand: assumes "stock_portfolio Mkt pf" shows "cls_val_process Mkt pf (Suc n) w = geom_proc (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" proof - have "cls_val_process Mkt pf (Suc n) w =(sum (\x. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w)) (stocks Mkt))" proof (rule subset_cls_val_process') show "finite (stocks Mkt)" using stocks by auto show "support_set pf \ stocks Mkt" using assms unfolding stock_portfolio_def by simp qed also have "... = (\x\ {stk, risk_free_asset}. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w))" using stocks by simp also have "... = prices Mkt stk (Suc n) w * pf stk (Suc n) w + (\ x\ {risk_free_asset}. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w))" by (simp add:two_stocks) also have "... = prices Mkt stk (Suc n) w * pf stk (Suc n) w + prices Mkt risk_free_asset (Suc n) w * pf risk_free_asset (Suc n) w" by simp also have "... = geom_proc (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using rf_price stk_price by simp finally show ?thesis . qed lemma (in CRR_market) pos_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "d < 1+r" and "0 < pf stk (Suc n) (spick w n False)" and "val_process Mkt pf n (spick w n False) \ 0" shows "cls_val_process Mkt pf (Suc n) (spick w n False) < 0" proof - define wnf where "wnf = spick w n False" have "cls_val_process Mkt pf (Suc n) (spick w n False) = geom_proc (Suc n) wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms by (simp add:stock_pf_uvp_expand) also have "... = d * geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using geometric_process spickI[of n w False] by simp also have "... = d * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" by simp also have "... < (1+r) * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms geom_rand_walk_strictly_positive S0_positive down_positive down_lt_up by simp also have "... = (1+r) * (geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf)" by (simp add: distrib_left) also have "... = (1+r) * val_process Mkt pf n wnf" using stock_pf_vp_expand assms by simp also have "... \ 0" proof - have "0 < 1+r" using assms down_positive by simp moreover have "val_process Mkt pf n wnf \ 0" using assms unfolding wnf_def by simp ultimately show "(1+r) * (val_process Mkt pf n wnf) \ 0" unfolding wnf_def using less_eq_real_def[of 0 "1+r"] mult_nonneg_nonpos[of "1+r" "val_process Mkt pf n (spick w n False)"] by simp qed finally show ?thesis . qed lemma (in CRR_market) neg_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "1+r < u" and "pf stk (Suc n) (spick w n True) < 0" and "val_process Mkt pf n (spick w n True) \ 0" shows "cls_val_process Mkt pf (Suc n) (spick w n True) < 0" proof - define wnf where "wnf = spick w n True" have "cls_val_process Mkt pf (Suc n) (spick w n True) = geom_proc (Suc n) wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms by (simp add:stock_pf_uvp_expand) also have "... = u * geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using geometric_process spickI[of n w True] by simp also have "... = u * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" by simp also have "... < (1+r) * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms geom_rand_walk_strictly_positive S0_positive down_positive down_lt_up by simp also have "... = (1+r) * (geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf)" by (simp add: distrib_left) also have "... = (1+r) * val_process Mkt pf n wnf" using stock_pf_vp_expand assms by simp also have "... \ 0" proof - have "0 < 1+r" using acceptable_rate by simp moreover have "val_process Mkt pf n wnf \ 0" using assms unfolding wnf_def by simp ultimately show "(1+r) * (val_process Mkt pf n wnf) \ 0" unfolding wnf_def using less_eq_real_def[of 0 "1+r"] mult_nonneg_nonpos[of "1+r" "val_process Mkt pf n (spick w n True)"] by simp qed finally show ?thesis . qed lemma (in CRR_market) zero_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "pf stk (Suc n) w = 0" and "pf risk_free_asset (Suc n) w \ 0" and "val_process Mkt pf n w \ 0" shows "cls_val_process Mkt pf (Suc n) w < 0" proof - have "cls_val_process Mkt pf (Suc n) w = S (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using assms by (simp add:stock_pf_uvp_expand) also have "... = disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using assms by simp also have "... = (1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" by simp also have "... < 0" proof - have "0 < 1+r" using acceptable_rate by simp moreover have "0 < disc_rfr_proc r n w" using acceptable_rate by (simp add: disc_rfr_proc_positive) ultimately have "0 < (1+r) * disc_rfr_proc r n w" by simp have 1: "0< pf risk_free_asset (Suc n) w \ 0 <(1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" proof (intro impI) assume "0 < pf risk_free_asset (Suc n) w" thus "0 < (1 + r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using \0 < (1+r) * disc_rfr_proc r n w\ by simp qed have 2: "pf risk_free_asset (Suc n) w < 0 \ (1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w < 0" proof (intro impI) assume "pf risk_free_asset (Suc n) w < 0" thus "(1 + r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w < 0" using \0 < (1+r) * disc_rfr_proc r n w\ by (simp add:mult_pos_neg) qed have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using assms by (simp add:stock_pf_vp_expand) also have "... = disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using assms by simp finally have "0\ disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" . have "0< pf risk_free_asset (Suc n) w \ pf risk_free_asset (Suc n) w < 0" using assms by linarith thus ?thesis using "2" \0 < disc_rfr_proc r n w\ \disc_rfr_proc r n w * pf risk_free_asset (Suc n) w \ 0\ mult_pos_pos by fastforce qed finally show ?thesis . qed lemma (in CRR_market) neg_pf_exists: assumes "stock_portfolio Mkt pf" and "trading_strategy pf" and "1+r < u" and "d < 1+r" and "val_process Mkt pf n w \ 0" and "pf stk (Suc n) w \ 0 \ pf risk_free_asset (Suc n) w \ 0" shows "\y. cls_val_process Mkt pf (Suc n) y < 0" proof - have "borel_predict_stoch_proc G (pf stk)" proof (rule inc_predict_support_trading_strat') show "trading_strategy pf" using assms by simp show "stk \ support_set pf \ {stk}" by simp qed hence "pf stk (Suc n) \ borel_measurable (G n)" unfolding predict_stoch_proc_def by simp have "val_process Mkt pf n \ borel_measurable (G n)" proof - have "borel_adapt_stoch_proc G (val_process Mkt pf)" using assms using support_adapt_def ats_val_process_adapted readable unfolding stock_portfolio_def by blast thus ?thesis unfolding adapt_stoch_proc_def by simp qed define wn where "wn = pseudo_proj_True n w" show ?thesis proof (cases "pf stk (Suc n) w \ 0") case True show ?thesis proof (cases "pf stk (Suc n) w > 0") case True have "0 0 < pf stk (Suc n) w\ by simp also have "... = pf stk (Suc n) wn" unfolding wn_def using \pf stk (Suc n) \ borel_measurable (G n)\ stoch_proc_subalg_nat_filt[of geom_proc] geometric_process nat_filtration_info stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = pf stk (Suc n) (spick wn n False)" using \pf stk (Suc n) \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed moreover have "0 \ val_process Mkt pf n (spick wn n False)" proof - have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = val_process Mkt pf n wn" unfolding wn_def using \val_process Mkt pf n \ borel_measurable (G n)\ nat_filtration_info stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = val_process Mkt pf n (spick wn n False)" using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed ultimately have "cls_val_process Mkt pf (Suc n) (spick wn n False) < 0" using assms by (simp add:pos_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto next case False have "0 >pf stk (Suc n) (spick wn n True)" proof - have "0 > pf stk (Suc n) w" using \\ 0 < pf stk (Suc n) w\ \pf stk (Suc n) w \ 0\ by simp also have "pf stk (Suc n) w = pf stk (Suc n) wn" unfolding wn_def using \pf stk (Suc n) \ borel_measurable (G n)\ nat_filtration_info stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = pf stk (Suc n) (spick wn n True)" using \pf stk (Suc n) \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed moreover have "0 \ val_process Mkt pf n (spick wn n True)" proof - have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = val_process Mkt pf n wn" unfolding wn_def using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = val_process Mkt pf n (spick wn n True)" using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed ultimately have "cls_val_process Mkt pf (Suc n) (spick wn n True) < 0" using assms by (simp add:neg_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto qed next case False hence "pf risk_free_asset (Suc n) w \ 0" using assms by simp hence "cls_val_process Mkt pf (Suc n) w < 0" using False assms by (auto simp add:zero_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto qed qed lemma (in CRR_market) non_zero_components: assumes "val_process Mkt pf n y \ 0" and "stock_portfolio Mkt pf" shows "pf stk (Suc n) y \ 0 \ pf risk_free_asset (Suc n) y \ 0" proof (rule ccontr) assume "\(pf stk (Suc n) y \ 0 \ pf risk_free_asset (Suc n) y \ 0)" hence "pf stk (Suc n) y = 0" "pf risk_free_asset (Suc n) y = 0" by auto have "val_process Mkt pf n y = geom_proc n y * pf stk (Suc n) y + disc_rfr_proc r n y * pf risk_free_asset (Suc n) y" using \stock_portfolio Mkt pf\ stock_pf_vp_expand[of pf n] by simp also have "... = 0" using \pf stk (Suc n) y = 0\ \pf risk_free_asset (Suc n) y = 0\ by simp finally have "val_process Mkt pf n y = 0" . moreover have "val_process Mkt pf n y \ 0" using assms by simp ultimately show False by simp qed lemma (in CRR_market) neg_pf_Suc: assumes "stock_portfolio Mkt pf" and "trading_strategy pf" and "self_financing Mkt pf" and "1+r < u" and "d < 1+r" and "cls_val_process Mkt pf n w < 0" shows "n \ m \ \y. cls_val_process Mkt pf m y < 0" proof (induct m) case 0 assume "n \ 0" hence "n=0" by simp thus "\y. cls_val_process Mkt pf 0 y < 0" using assms by auto next case (Suc m) assume "n \ Suc m" thus "\y. cls_val_process Mkt pf (Suc m) y < 0" proof (cases "n < Suc m") case False hence "n = Suc m" using \n \ Suc m\ by simp thus "\y. cls_val_process Mkt pf (Suc m) y < 0" using assms by auto next case True hence "n \ m" by simp hence "\y. cls_val_process Mkt pf m y < 0" using Suc by simp from this obtain y where "cls_val_process Mkt pf m y < 0" by auto hence "val_process Mkt pf m y < 0" using assms by (simp add:self_financingE) hence "val_process Mkt pf m y \ 0" by simp have "val_process Mkt pf m y \ 0" using \val_process Mkt pf m y < 0\ by simp hence "pf stk (Suc m) y \ 0 \ pf risk_free_asset (Suc m) y \ 0" using assms non_zero_components by simp thus "\y. cls_val_process Mkt pf (Suc m) y < 0" using neg_pf_exists[of pf m y] assms \val_process Mkt pf m y \ 0\ by simp qed qed lemma (in CRR_market) viable_if: assumes "1+r < u" and "d < 1+r" shows "viable_market Mkt" unfolding viable_market_def proof (rule ccontr) assume "\(\p. stock_portfolio Mkt p \ \ arbitrage_process Mkt p)" hence "\p. stock_portfolio Mkt p \ arbitrage_process Mkt p" by simp from this obtain pf where "stock_portfolio Mkt pf" and "arbitrage_process Mkt pf" by auto have "(\ m. (self_financing Mkt pf) \ (trading_strategy pf) \ (\w \ space M. cls_val_process Mkt pf 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt pf m w) \ 0 < \

(w in M. cls_val_process Mkt pf m w > 0))" using \arbitrage_process Mkt pf\ using arbitrage_processE by simp from this obtain m where "self_financing Mkt pf" and "(trading_strategy pf)" and "(\w \ space M. cls_val_process Mkt pf 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt pf m w)" and "0 < \

(w in M. cls_val_process Mkt pf m w > 0)" by auto have "{w\ space M. cls_val_process Mkt pf m w > 0} \ {}" using \0 < \

(w in M. cls_val_process Mkt pf m w > 0)\ by force hence "\w\ space M. cls_val_process Mkt pf m w > 0" by auto from this obtain y where "y\ space M" and "cls_val_process Mkt pf m y > 0" by auto define A where "A = {n::nat. n \ m \ cls_val_process Mkt pf n y > 0}" have "finite A" unfolding A_def by auto have "m \ A" using \cls_val_process Mkt pf m y > 0\ unfolding A_def by simp hence "A \ {}" by auto hence "Min A \ A" using \finite A\ by simp have "Min A \ m" using \finite A\ \m\ A\ by simp have "0 < Min A" proof - have "cls_val_process Mkt pf 0 y = 0" using \y\ space M\ \\w \ space M. cls_val_process Mkt pf 0 w = 0\ by simp hence "0\ A" unfolding A_def by simp moreover have "0 \ Min A" by simp ultimately show ?thesis using \Min A \ A\ neq0_conv by fastforce qed hence "\l. Suc l = Min A" using Suc_diff_1 by blast from this obtain l where "Suc l = Min A" by auto have "cls_val_process Mkt pf l y \ 0" proof - have "l < Min A" using \Suc l = Min A\ by simp hence "l\ A" using \finite A\ \A \ {}\ by auto moreover have "l \ m" using \Suc l = Min A\ \m\ A\ \finite A\ \A \ {}\ \l < Min A\ by auto ultimately show ?thesis unfolding A_def by auto qed hence "val_process Mkt pf l y \ 0" using \self_financing Mkt pf\ by (simp add:self_financingE) moreover have "pf stk (Suc l) y \ 0 \ pf risk_free_asset (Suc l) y \ 0" proof (rule ccontr) assume "\(pf stk (Suc l) y \ 0 \ pf risk_free_asset (Suc l) y \ 0)" hence "pf stk (Suc l) y = 0" "pf risk_free_asset (Suc l) y = 0" by auto have "cls_val_process Mkt pf (Min A) y = geom_proc (Suc l) y * pf stk (Suc l) y + disc_rfr_proc r (Suc l) y * pf risk_free_asset (Suc l) y" using \stock_portfolio Mkt pf\ \Suc l = Min A\ stock_pf_uvp_expand[of pf l] by simp also have "... = 0" using \pf stk (Suc l) y = 0\ \pf risk_free_asset (Suc l) y = 0\ by simp finally have "cls_val_process Mkt pf (Min A) y = 0" . moreover have "cls_val_process Mkt pf (Min A) y > 0" using \Min A \ A\ unfolding A_def by simp ultimately show False by simp qed ultimately have "\z. cls_val_process Mkt pf (Suc l) z < 0" using assms \stock_portfolio Mkt pf\ \trading_strategy pf\ by (simp add:neg_pf_exists) from this obtain z where "cls_val_process Mkt pf (Suc l) z < 0" by auto hence "\x'. cls_val_process Mkt pf m x' < 0" using neg_pf_Suc assms \trading_strategy pf\ \self_financing Mkt pf\ \Suc l = Min A\ \Min A \ m\ \stock_portfolio Mkt pf\ by simp from this obtain x' where "cls_val_process Mkt pf m x' < 0" by auto have "x'\ space M" using bernoulli_stream_space bernoulli by auto hence "x'\ {w\ space M. \0 \ cls_val_process Mkt pf m w}" using \cls_val_process Mkt pf m x' < 0\ by auto from \AE w in M. 0 \ cls_val_process Mkt pf m w\ obtain N where "{w\ space M. \0 \ cls_val_process Mkt pf m w} \ N" and "emeasure M N = 0" and "N\ sets M" using AE_E by auto have "{w\ space M. (stake m w = stake m x')} \ N" proof fix x assume "x \ {w \ space M. stake m w = stake m x'}" hence "x\ space M" and "stake m x = stake m x'" by auto have "cls_val_process Mkt pf m \ borel_measurable (G m)" proof - have "borel_adapt_stoch_proc G (cls_val_process Mkt pf)" using \trading_strategy pf\ \stock_portfolio Mkt pf\ by (meson support_adapt_def readable stock_portfolio_def subsetCE cls_val_process_adapted) thus ?thesis unfolding adapt_stoch_proc_def by simp qed hence "cls_val_process Mkt pf m x' = cls_val_process Mkt pf m x" using \stake m x = stake m x'\ borel_measurable_stake[of "cls_val_process Mkt pf m" m x x'] pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) hence "cls_val_process Mkt pf m x < 0" using \cls_val_process Mkt pf m x' < 0\ by simp thus "x\ N" using \{w\ space M. \0 \ cls_val_process Mkt pf m w} \ N\ \x\ space M\ \cls_val_process Mkt pf (Suc l) z < 0\ by auto qed moreover have "emeasure M {w\ space M. (stake m w = stake m x')} \ 0" using bernoulli_stream_pref_prob_neq_zero psgt pslt by simp ultimately show False using \emeasure M N = 0\ \N \ events\ emeasure_eq_0 by blast qed lemma (in CRR_market) viable_only_if_d: assumes "viable_market Mkt" shows "d < 1+r" proof (rule ccontr) assume "\ d < 1+r" hence "1+r \ d" by simp define arb_pf where "arb_pf = (\ (x::'a) (n::nat) w. 0::real)(stk:= (\ n w. 1), risk_free_asset := (\ n w. - geom_proc 0 w))" have "support_set arb_pf = {stk, risk_free_asset}" proof show "support_set arb_pf \ {stk, risk_free_asset}" by (simp add: arb_pf_def subset_iff support_set_def) have "stk\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks by simp moreover have "risk_free_asset\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks geometric_process S0_positive by simp ultimately show "{stk, risk_free_asset}\ support_set arb_pf" by simp qed hence "stock_portfolio Mkt arb_pf" using stocks by (simp add: portfolio_def stock_portfolio_def) have "arbitrage_process Mkt arb_pf" proof (rule arbitrage_processI, intro exI conjI) show "self_financing Mkt arb_pf" unfolding arb_pf_def using \support_set arb_pf = {stk, risk_free_asset}\ by (simp add: static_portfolio_self_financing) show "trading_strategy arb_pf" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio arb_pf" unfolding portfolio_def using \support_set arb_pf = {stk, risk_free_asset}\ by simp fix asset assume "asset\ support_set arb_pf" show "borel_predict_stoch_proc G (arb_pf asset)" proof (cases "asset = stk") case True hence "arb_pf asset = (\ n w. 1)" unfolding arb_pf_def by (simp add: two_stocks) show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. 1)\ by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. 1)\ by simp qed qed next case False hence "arb_pf asset = (\ n w. - geom_proc 0 w)" using \support_set arb_pf = {stk, risk_free_asset}\ \asset \ support_set arb_pf\ unfolding arb_pf_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. - geom_proc 0 w)\ geometric_process by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. - geom_proc 0 w)\ geometric_process by simp qed qed qed qed show "\w\space M. cls_val_process Mkt arb_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt arb_pf 0 w = geom_proc 0 w * arb_pf stk (Suc 0) w + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_vp_expand \stock_portfolio Mkt arb_pf\ using \self_financing Mkt arb_pf\ self_financingE by fastforce also have "... = geom_proc 0 w * (1) + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = geom_proc 0 w + arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc 0 w - geom_proc 0 w" unfolding arb_pf_def by simp also have "... = 0" by simp finally show "cls_val_process Mkt arb_pf 0 w = 0" . qed have dev: "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" proof (intro ballI) fix w assume "w\ space M" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w * arb_pf stk (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_uvp_expand \stock_portfolio Mkt arb_pf\ by simp also have "... = geom_proc (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = geom_proc (Suc 0) w + (1+r) * arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" by (simp add:arb_pf_def) finally show "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" . qed have iniT: "\w\ space M. snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w > 0" proof (intro ballI impI) fix w assume "w\ space M" and "snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = u * geom_proc 0 w - (1+r) * geom_proc 0 w" using \snth w 0\ geometric_process by simp also have "... = (u - (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... > 0" using S0_positive \1 + r \ d\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w > 0" . qed have iniF: "\w\ space M. \snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (intro ballI impI) fix w assume "w\ space M" and "\snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = d * geom_proc 0 w - (1+r) * geom_proc 0 w" using \\snth w 0\ geometric_process by simp also have "... = (d - (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... \ 0" using S0_positive \1 + r \ d\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" . qed have "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof fix w assume "w\ space M" show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (cases "snth w 0") case True thus ?thesis using \w\ space M\ iniT by auto next case False thus ?thesis using \w\ space M\ iniF by simp qed qed thus "AE w in M. 0 \ cls_val_process Mkt arb_pf (Suc 0) w" by simp show "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" proof - have "cls_val_process Mkt arb_pf (Suc 0) \ borel_measurable M" using borel_adapt_stoch_proc_borel_measurable cls_val_process_adapted \trading_strategy arb_pf\ \stock_portfolio Mkt arb_pf\ using support_adapt_def readable unfolding stock_portfolio_def by blast hence set_event:"{w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w} \ sets M" using borel_measurable_iff_greater by blast have "\n. emeasure M {w \ space M. w !! n} = ennreal p" using bernoulli p_gt_0 p_lt_1 bernoulli_stream_component_probability[of M p] by auto hence "emeasure M {w \ space M. w !! 0} = ennreal p" by blast moreover have "{w \ space M. w !! 0} \ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" proof fix w assume "w\ {w \ space M. w !! 0}" hence "w \ space M" and "w !! 0" by auto note wprops = this hence "0 < cls_val_process Mkt arb_pf 1 w" using iniT by simp thus "w\ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using wprops by simp qed ultimately have "p \ emeasure M {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using emeasure_mono set_event by fastforce hence "p \ prob {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" by (simp add: emeasure_eq_measure) thus "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" using psgt by simp qed qed thus False using assms unfolding viable_market_def using \stock_portfolio Mkt arb_pf\ by simp qed lemma (in CRR_market) viable_only_if_u: assumes "viable_market Mkt" shows "1+r < u" proof (rule ccontr) assume "\ 1+r < u" hence "u \ 1+r" by simp define arb_pf where "arb_pf = (\ (x::'a) (n::nat) w. 0::real)(stk:= (\ n w. -1), risk_free_asset := (\ n w. geom_proc 0 w))" have "support_set arb_pf = {stk, risk_free_asset}" proof show "support_set arb_pf \ {stk, risk_free_asset}" by (simp add: arb_pf_def subset_iff support_set_def) have "stk\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks by simp moreover have "risk_free_asset\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks geometric_process S0_positive by simp ultimately show "{stk, risk_free_asset}\ support_set arb_pf" by simp qed hence "stock_portfolio Mkt arb_pf" using stocks by (simp add: portfolio_def stock_portfolio_def) have "arbitrage_process Mkt arb_pf" proof (rule arbitrage_processI, intro exI conjI) show "self_financing Mkt arb_pf" unfolding arb_pf_def using \support_set arb_pf = {stk, risk_free_asset}\ by (simp add: static_portfolio_self_financing) show "trading_strategy arb_pf" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio arb_pf" unfolding portfolio_def using \support_set arb_pf = {stk, risk_free_asset}\ by simp fix asset assume "asset\ support_set arb_pf" show "borel_predict_stoch_proc G (arb_pf asset)" proof (cases "asset = stk") case True hence "arb_pf asset = (\ n w. -1)" unfolding arb_pf_def by (simp add: two_stocks) show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. -1)\ by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. -1)\ by simp qed qed next case False hence "arb_pf asset = (\ n w. geom_proc 0 w)" using \support_set arb_pf = {stk, risk_free_asset}\ \asset \ support_set arb_pf\ unfolding arb_pf_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. geom_proc 0 w)\ geometric_process by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. geom_proc 0 w)\ geometric_process by simp qed qed qed qed show "\w\space M. cls_val_process Mkt arb_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt arb_pf 0 w = geom_proc 0 w * arb_pf stk (Suc 0) w + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_vp_expand \stock_portfolio Mkt arb_pf\ using \self_financing Mkt arb_pf\ self_financingE by fastforce also have "... = geom_proc 0 w * (-1) + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = -geom_proc 0 w + arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc 0 w - geom_proc 0 w" unfolding arb_pf_def by simp also have "... = 0" by simp finally show "cls_val_process Mkt arb_pf 0 w = 0" . qed have dev: "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" proof (intro ballI) fix w assume "w\ space M" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w * arb_pf stk (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_uvp_expand \stock_portfolio Mkt arb_pf\ by simp also have "... = -geom_proc (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = -geom_proc (Suc 0) w + (1+r) * arb_pf risk_free_asset (Suc 0) w" by simp also have "... = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" by (simp add:arb_pf_def) finally show "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" . qed have iniT: "\w\ space M. snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (intro ballI impI) fix w assume "w\ space M" and "snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = - u * geom_proc 0 w + (1+r) * geom_proc 0 w" using \snth w 0\ geometric_process by simp also have "... = (-u + (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... \ 0" using S0_positive \u\ 1 + r\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" . qed have iniF: "\w\ space M. \snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w > 0" proof (intro ballI impI) fix w assume "w\ space M" and "\snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = -d * geom_proc 0 w + (1+r) * geom_proc 0 w" using \\snth w 0\ geometric_process by simp also have "... = (-d + (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... > 0" using S0_positive \u <= 1 + r\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w > 0" . qed have "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof fix w assume "w\ space M" show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (cases "snth w 0") case True thus ?thesis using \w\ space M\ iniT by simp next case False thus ?thesis using \w\ space M\ iniF by auto qed qed thus "AE w in M. 0 \ cls_val_process Mkt arb_pf (Suc 0) w" by simp show "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" proof - have "cls_val_process Mkt arb_pf (Suc 0) \ borel_measurable M" using borel_adapt_stoch_proc_borel_measurable cls_val_process_adapted \trading_strategy arb_pf\ \stock_portfolio Mkt arb_pf\ using support_adapt_def readable unfolding stock_portfolio_def by blast hence set_event:"{w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w} \ sets M" using borel_measurable_iff_greater by blast have "\n. emeasure M {w \ space M. \w !! n} = ennreal (1-p)" using bernoulli p_gt_0 p_lt_1 bernoulli_stream_component_probability_compl[of M p] by auto hence "emeasure M {w \ space M. \w !! 0} = ennreal (1-p)" by blast moreover have "{w \ space M. \w !! 0} \ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" proof fix w assume "w\ {w \ space M. \w !! 0}" hence "w \ space M" and "\w !! 0" by auto note wprops = this hence "0 < cls_val_process Mkt arb_pf 1 w" using iniF by simp thus "w\ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using wprops by simp qed ultimately have "1-p \ emeasure M {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using emeasure_mono set_event by fastforce hence "1-p \ prob {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" by (simp add: emeasure_eq_measure) thus "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" using pslt by simp qed qed thus False using assms unfolding viable_market_def using \stock_portfolio Mkt arb_pf\ by simp qed lemma (in CRR_market) viable_iff: shows "viable_market Mkt \ (d < 1+r \ 1+r < u)" using viable_if viable_only_if_d viable_only_if_u by auto subsection \Risk-neutral probability space for the geometric random walk\ lemma (in CRR_market) stock_price_borel_measurable: shows "borel_adapt_stoch_proc G (prices Mkt stk)" proof - have "borel_adapt_stoch_proc (stoch_proc_filt M geom_proc borel) (prices Mkt stk)" by (simp add: geom_rand_walk_borel_measurable stk_price stoch_proc_filt_adapt) thus ?thesis by (simp add:stock_filtration) qed lemma (in CRR_market) risk_free_asset_martingale: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "martingale N G (discounted_value r (prices Mkt risk_free_asset))" proof - have "filtration N G" by (simp add: assms bernoulli_gen_filtration) moreover have "\n. sigma_finite_subalgebra N (G n)" by (simp add: assms bernoulli_sigma_finite) moreover have "finite_measure N" using assms bernoulli_stream_def prob_space.prob_space_stream_space prob_space_def prob_space_measure_pmf by auto moreover have "discounted_value r (prices Mkt risk_free_asset) = (\ n w. 1)" using discounted_rfr by auto ultimately show ?thesis using finite_measure.constant_martingale by simp qed lemma (in infinite_coin_toss_space) nat_filtration_from_eq_sets: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "sets (infinite_coin_toss_space.nat_filtration N n) = sets (nat_filtration n)" proof - have "sigma_sets (space (bernoulli_stream q)) {pseudo_proj_True n -` B \ space N |B. B \ sets (bernoulli_stream q)} = sigma_sets (space (bernoulli_stream p)) {pseudo_proj_True n -` B \ space M |B. B \ sets (bernoulli_stream p)}" proof - have "sets N = events" by (metis assms(1) bernoulli_stream_def infinite_coin_toss_space_axioms infinite_coin_toss_space_def sets_measure_pmf sets_stream_space_cong) then show ?thesis using assms(1) bernoulli_stream_space infinite_coin_toss_space_axioms infinite_coin_toss_space_def by auto qed thus ?thesis using infinite_coin_toss_space.nat_filtration_sets using assms(1) assms(2) assms(3) infinite_coin_toss_space_axioms infinite_coin_toss_space_def by auto qed lemma (in CRR_market) geom_proc_integrable: assumes "N = bernoulli_stream q" and "0 \ q" and "q \ 1" shows "integrable N (geom_proc n)" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by unfold_locales show "geom_proc n \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using geometric_process prob_grw.geom_rand_walk_borel_adapted[of q N geom_proc u d init] by (metis \infinite_coin_toss_space q N\ geom_rand_walk_pseudo_proj_True infinite_coin_toss_space.nat_filtration_borel_measurable_characterization prob_grw.geom_rand_walk_borel_measurable prob_grw_axioms prob_grw_def) qed lemma (in CRR_market) CRR_infinite_cts_filtration: shows "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) lemma (in CRR_market) proj_stoch_proc_geom_disc_fct: shows "disc_fct (proj_stoch_proc geom_proc n)" unfolding disc_fct_def using CRR_infinite_cts_filtration by (simp add: countable_finite geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_set_finite_range) lemma (in CRR_market) proj_stoch_proc_geom_rng: assumes "N = bernoulli_stream q" shows "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" proof - have "random_variable (stream_space borel) (proj_stoch_proc geom_proc n)" using CRR_infinite_cts_filtration using geom_rand_walk_borel_adapted nat_discrete_filtration proj_stoch_measurable_if_adapted by blast then show ?thesis using assms(1) bernoulli bernoulli_stream_def by auto qed lemma (in CRR_market) proj_stoch_proc_geom_open_set: shows "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume "r\ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof show "infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast show "range (proj_stoch_proc geom_proc n) \ infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r = {r}" using infinite_cts_filtration.stream_space_single_preimage \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast qed qed lemma (in CRR_market) bernoulli_AE_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof (rule finite_measure.charact_cond_exp') have "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) show "finite_measure N" using assms by (simp add: bernoulli_stream_def prob_space.finite_measure prob_space.prob_space_stream_space prob_space_measure_pmf) show "disc_fct (proj_stoch_proc geom_proc n)" using proj_stoch_proc_geom_disc_fct by simp show "integrable N X" using assms by simp show "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" using assms proj_stoch_proc_geom_rng by simp show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" using proj_stoch_proc_geom_open_set by simp qed lemma (in CRR_market) geom_proc_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_AE_cond_exp) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in CRR_market) expl_cond_eq_sets: assumes "N = bernoulli_stream q" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule expl_cond_exp_borel) show "proj_stoch_proc geom_proc n \ space N \ space (stream_space borel)" proof - have "random_variable (stream_space borel) (proj_stoch_proc geom_proc n)" using CRR_infinite_cts_filtration geom_rand_walk_borel_adapted proj_stoch_measurable_if_adapted nat_discrete_filtration by blast then show ?thesis by (simp add: assms(1) bernoulli bernoulli_stream_space measurable_def) qed show "disc_fct (proj_stoch_proc geom_proc n)" unfolding disc_fct_def using CRR_infinite_cts_filtration by (simp add: countable_finite geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_set_finite_range) show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume "r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof show "infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast show "range (proj_stoch_proc geom_proc n) \ infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r = {r}" using infinite_cts_filtration.stream_space_single_preimage \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast qed qed qed lemma (in CRR_market) bernoulli_real_cond_exp_AE: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof - have "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_AE_cond_exp by simp show "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" proof - have "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by simp moreover have "subalgebra (infinite_coin_toss_space.nat_filtration N n) (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_coin_toss_space.stoch_proc_subalg_nat_filt[of q N geom_proc n] infinite_cts_filtration.stoch_proc_filt_gen[of q N] by (metis \infinite_coin_toss_space q N\ infinite_cts_filtration_axioms.intro infinite_cts_filtration_def prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) ultimately show ?thesis using measurable_from_subalg by blast qed show "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" proof - have "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by (simp add: expl_cond_eq_sets assms) moreover have "subalgebra (infinite_coin_toss_space.nat_filtration N n) (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_coin_toss_space.stoch_proc_subalg_nat_filt[of q N geom_proc n] infinite_cts_filtration.stoch_proc_filt_gen[of q N] by (metis \infinite_coin_toss_space q N\ infinite_cts_filtration_axioms.intro infinite_cts_filtration_def prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) ultimately show ?thesis using measurable_from_subalg by blast qed show "0 < q" and "q < 1" using assms by auto qed thus ?thesis by simp qed lemma (in CRR_market) geom_proc_real_cond_exp_AE: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_real_cond_exp_AE) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in CRR_market) geom_proc_stoch_proc_filt: assumes "N= bernoulli_stream q" and "0 < q" and "q < 1" shows "stoch_proc_filt N geom_proc borel n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" proof (rule infinite_cts_filtration.stoch_proc_filt_gen) show "infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)" unfolding infinite_cts_filtration_def proof show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "infinite_cts_filtration_axioms N (infinite_coin_toss_space.nat_filtration N)" using infinite_cts_filtration_axioms_def by blast qed show "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) geom_proc" using \infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)\ prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def using infinite_cts_filtration_def by auto qed lemma (in CRR_market) bernoulli_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof - have aeq: "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_AE_cond_exp by simp have "\w. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_real_cond_exp_AE by simp moreover have "stoch_proc_filt N geom_proc borel n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" using assms geom_proc_stoch_proc_filt by simp ultimately show ?thesis by simp qed lemma (in CRR_market) stock_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_cond_exp) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in prob_space) discount_factor_real_cond_exp: assumes "integrable M X" and "subalgebra M G" and "-1 < r" shows "AE w in M. real_cond_exp M G (\x. discount_factor r n x * X x) w = discount_factor r n w * (real_cond_exp M G X) w" proof (rule sigma_finite_subalgebra.real_cond_exp_mult) show "sigma_finite_subalgebra M G" using assms subalgebra_sigma_finite by simp show "discount_factor r n \ borel_measurable G" by (simp add: discount_factor_borel_measurable) show "random_variable borel X" using assms by simp show "integrable M (\x. discount_factor r n x * X x)" using assms discounted_integrable[of M "\n. X"] unfolding discounted_value_def by simp qed lemma (in prob_space) discounted_value_real_cond_exp: assumes "integrable M X" and "-1 < r" and "subalgebra M G" shows "AE w in M. real_cond_exp M G ((discounted_value r (\ m. X)) n) w = discounted_value r (\m. (real_cond_exp M G X)) n w" using assms unfolding discounted_value_def init_triv_filt_def filtration_def by (simp add: assms discount_factor_real_cond_exp) lemma (in CRR_market) assumes "q = (1 + r - d)/(u -d)" and "viable_market Mkt" shows gt_param: "0 < q" and lt_param: "q < 1" and risk_neutral_param: "u * q + d * (1 - q) = 1 + r" proof - show "0 < q" using down_lt_up viable_only_if_d assms by simp show "q < 1" using down_lt_up viable_only_if_u assms by simp show "u * q + d * (1 - q) = 1 + r" proof - have "1 - q = 1 - (1 + r - d) / (u - d)" using assms by simp also have "... = (u - d)/(u - d) - (1 + r - d) / (u - d)" using down_lt_up by simp also have "... = (u - d - (1 + r - d))/(u-d)" using diff_divide_distrib[of "u - d" "1 + r -d" "u -d"] by simp also have "... = (u - 1 - r)/(u-d)" by simp finally have "1 - q = (u - 1 - r)/(u -d)" . hence "u * q + d * (1 - q) = u * (1 + r - d)/(u - d) + d * (u - 1 - r)/(u - d)" using assms by simp also have "... = (u * (1 + r - d) + d * (u - 1 - r))/(u - d)" using add_divide_distrib[of "u * (1 + r - d)"] by simp also have "... = (u * (1 + r) - u * d + d * u - d * (1 + r))/(u - d)" by (simp add: diff_diff_add right_diff_distrib') also have "... = (u * (1+r) - d * (1+r))/(u - d)" by simp also have "... = ((u - d) * (1+r))/(u - d)" by (simp add: left_diff_distrib) also have "... = 1 + r" using down_lt_up by simp finally show ?thesis . qed qed lemma (in CRR_market) bernoulli_expl_cond_expect_adapt: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) f\ borel_measurable (G n)" proof - have "sets N = sets M" using assms by (simp add: bernoulli bernoulli_stream_def sets_stream_space_cong) have icf: "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) have "G n = stoch_proc_filt M geom_proc borel n" using stock_filtration by simp also have "... = fct_gen_subalgebra M (stream_space borel) (proj_stoch_proc geom_proc n)" proof (rule infinite_cts_filtration.stoch_proc_filt_gen) show "infinite_cts_filtration p M nat_filtration" using icf . show "borel_adapt_stoch_proc nat_filtration geom_proc" using geom_rand_walk_borel_adapted . qed also have "... = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" by (rule fct_gen_subalgebra_eq_sets, (simp add: \sets N = sets M\)) finally have "G n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" . moreover have "expl_cond_expect N (proj_stoch_proc geom_proc n) f \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by (simp add: expl_cond_eq_sets assms) ultimately show ?thesis by simp qed lemma (in CRR_market) real_cond_exp_discount_stock: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" proof - have qlt: "0 < q" and qgt: "q < 1" using assms by auto have "G n = (fct_gen_subalgebra M (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_cts_filtration.stoch_proc_filt_gen[of p M nat_filtration geom_proc n] geometric_process geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by simp also have "... = (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule fct_gen_subalgebra_eq_sets) show "events = sets N" using assms qlt qgt by (simp add: bernoulli bernoulli_stream_def sets_stream_space_cong) qed finally have "G n = (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" . hence "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (prices Mkt stk) (Suc n)) w" by simp moreover have "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (prices Mkt stk) (Suc n)) w = real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n)) w" proof - have "\w. (discounted_value r (prices Mkt stk) (Suc n)) w = (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n)) w" proof fix w show "discounted_value r (prices Mkt stk) (Suc n) w = discounted_value r (\m. prices Mkt stk (Suc n)) (Suc n) w" by (simp add: discounted_value_def) qed hence "(discounted_value r (prices Mkt stk) (Suc n)) = (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n))" by auto thus ?thesis by simp qed moreover have "AE w in N. (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n))) w = discounted_value r (\m. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (Suc n) w" proof (rule prob_space.discounted_value_real_cond_exp) show "-1 < r" using acceptable_rate by simp show "integrable N (prices Mkt stk (Suc n))" using stk_price geom_proc_integrable assms qlt qgt by simp show "subalgebra N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule fct_gen_subalgebra_is_subalgebra) show "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" proof - have "proj_stoch_proc geom_proc n \ measurable M (stream_space borel)" proof (rule proj_stoch_measurable_if_adapted) show "borel_adapt_stoch_proc nat_filtration geom_proc" using geometric_process geom_rand_walk_borel_adapted by simp show "filtration M nat_filtration" using CRR_infinite_cts_filtration by (simp add: nat_discrete_filtration) qed thus ?thesis using assms bernoulli_stream_equiv filt_equiv_measurable qlt qgt psgt pslt by blast qed qed show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) qed moreover have "AE w in N. discounted_value r (\m. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (Suc n) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" proof (rule discounted_AE_cong) have "AEeq N (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (\w. q * (prices Mkt stk) (Suc n) (pseudo_proj_True n w) + (1 - q) * (prices Mkt stk) (Suc n) (pseudo_proj_False n w))" proof (rule infinite_cts_filtration.f_borel_Suc_real_cond_exp) show icf: "infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)" unfolding infinite_cts_filtration_def proof show "infinite_coin_toss_space q N" using assms qlt qgt by (simp add: infinite_coin_toss_space_def) show "infinite_cts_filtration_axioms N (infinite_coin_toss_space.nat_filtration N)" using infinite_cts_filtration_axioms_def by blast qed have badapt: "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) (prices Mkt stk)" using stk_price prob_grw.geom_rand_walk_borel_adapted[of q N geom_proc] unfolding adapt_stoch_proc_def by (metis (full_types) borel_measurable_integrable geom_proc_integrable geom_rand_walk_pseudo_proj_True icf infinite_coin_toss_space.nat_filtration_borel_measurable_characterization infinite_coin_toss_space_def infinite_cts_filtration_def) show "prices Mkt stk (Suc n) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (Suc n))" using badapt unfolding adapt_stoch_proc_def by simp show "proj_stoch_proc geom_proc n \ infinite_coin_toss_space.nat_filtration N n \\<^sub>M stream_space borel" proof (rule proj_stoch_adapted_if_adapted) show "filtration N (infinite_coin_toss_space.nat_filtration N)" using icf using infinite_coin_toss_space.nat_discrete_filtration infinite_cts_filtration_def by blast show "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) geom_proc" using badapt stk_price by simp qed show "set_discriminating n (proj_stoch_proc geom_proc n) (stream_space borel)" unfolding set_discriminating_def proof (intro allI impI) fix w assume "proj_stoch_proc geom_proc n w \ proj_stoch_proc geom_proc n (pseudo_proj_True n w)" hence False using CRR_infinite_cts_filtration by (metis \proj_stoch_proc geom_proc n w \ proj_stoch_proc geom_proc n (pseudo_proj_True n w)\ geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_proj_invariant) thus "\A\sets (stream_space borel). (proj_stoch_proc geom_proc n w \ A) = (proj_stoch_proc geom_proc n (pseudo_proj_True n w) \ A)" by simp qed show "\w. proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (infinite_coin_toss_space.nat_filtration N n)" proof fix w show "proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (infinite_coin_toss_space.nat_filtration N n)" using \proj_stoch_proc geom_proc n \ infinite_coin_toss_space.nat_filtration N n \\<^sub>M stream_space borel\ using assms geom_rand_walk_borel_adapted nat_filtration_from_eq_sets qlt qgt infinite_cts_filtration.proj_stoch_singleton_set CRR_infinite_cts_filtration by blast qed show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume asm: "r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" define A where "A = infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r" have "A \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set unfolding A_def using badapt icf stk_price asm by blast moreover have "range (proj_stoch_proc geom_proc n) \ A = {r}" unfolding A_def using badapt icf stk_price infinite_cts_filtration.stream_space_single_preimage asm by blast ultimately show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" by auto qed show "\y z. proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n \ prices Mkt stk (Suc n) y = prices Mkt stk (Suc n) z" proof (intro allI impI) fix y z assume "proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n" hence "geom_proc n y = geom_proc n z" using proj_stoch_proc_component(2)[of n n] proof - show ?thesis by (metis \\w f. n \ n \ proj_stoch_proc f n w !! n = f n w\ \proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n\ order_refl) qed hence "geom_proc (Suc n) y = geom_proc (Suc n) z" using geometric_process by (simp add: \proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n\) thus "prices Mkt stk (Suc n) y = prices Mkt stk (Suc n) z" using stk_price by simp qed show "0 < q" and "q < 1" using assms by auto qed moreover have "\w. q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = (q * u + (1 - q) * d) * prices Mkt stk n w" proof fix w have "q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = q * geom_proc (Suc n) (pseudo_proj_True n w) + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" by (simp add:stk_price) also have "... = q * u * geom_proc n (pseudo_proj_True n w) + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" using geometric_process unfolding pseudo_proj_True_def by simp also have "... = q * u * geom_proc n w + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" by (metis geom_rand_walk_pseudo_proj_True o_apply) also have "... = q * u * geom_proc n w + (1-q) * d * geom_proc n (pseudo_proj_False n w)" using geometric_process unfolding pseudo_proj_False_def by simp also have "... = q * u * geom_proc n w + (1-q) * d * geom_proc n w" by (metis geom_rand_walk_pseudo_proj_False o_apply) also have "... = (q * u + (1 - q) * d) * geom_proc n w" by (simp add: distrib_right) finally show "q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = (q * u + (1 - q) * d) * prices Mkt stk n w" using stk_price by simp qed ultimately show "AEeq N (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (\w. (q * u + (1 - q) * d) * prices Mkt stk n w)" by simp qed ultimately show ?thesis by auto qed lemma (in CRR_market) risky_asset_martingale_only_if: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "martingale N G (discounted_value r (prices Mkt stk))" shows "q = (1 + r - d) / (u - d)" proof - have "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = discounted_value r (prices Mkt stk) 0 w" using assms unfolding martingale_def by simp hence "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = prices Mkt stk 0 w" by (simp add: discounted_init) moreover have "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk 0 w) (Suc 0) w" using assms real_cond_exp_discount_stock by simp ultimately have "AE w in N. discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk 0 w) (Suc 0) w = prices Mkt stk 0 w" by auto hence "AE w in N. discounted_value r (\m w. (q * u + (1 - q) * d) * init) (Suc 0) w = (\w. init) w" using stk_price geometric_process by simp hence "AE w in N. discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init = (\w. init) w" unfolding discounted_value_def by simp hence "AE w in N. (1+r) * discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init = (1+r) * (\w. init) w" by auto hence prev: "AE w in N. discount_factor r 0 w * (q * u + (1 - q) * d) * init = (1+r) * (\w. init) w" using discount_factor_times_rfr[of r 0] acceptable_rate proof - have "\s. (1 + r) * discount_factor r (Suc 0) (s::bool stream) = discount_factor r 0 s" by (metis (no_types) \\w. - 1 < r \ (1 + r) * discount_factor r (Suc 0) w = discount_factor r 0 w\ acceptable_rate) then show ?thesis using \AEeq N (\w. (1 + r) * discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init) (\w. (1 + r) * init)\ by presburger qed hence "\w. (\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init) w = (\w. (1+r) * init) w" proof - have "(\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" proof (rule borel_measurable_times)+ show "(\x. init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp show "(\x. q * u + (1 - q) * d) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp show "discount_factor r 0 \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" using discount_factor_nonrandom[of r 0 "infinite_coin_toss_space.nat_filtration N 0"] by simp qed moreover have "(\w. (1 + r) * init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp moreover have "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) ultimately show ?thesis using prev infinite_coin_toss_space.nat_filtration_AE_eq[of q N "(\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init)" "(\w. (1 + r) * init)" 0] assms by (simp add: discount_factor_init) qed hence "(q * u + (1 - q) * d) * init = (1+r) * init" by (simp add: discount_factor_init) hence "q * u + (1 - q) * d = 1+r" using S0_positive by simp hence "q * u + d - q * d = 1+r" by (simp add: left_diff_distrib) hence "q * (u - d) = 1 + r - d" by (metis (no_types, hide_lams) add.commute add.left_commute add_diff_cancel_left' add_uminus_conv_diff left_diff_distrib mult.commute) thus "q = (1 + r - d) / (u - d)" using down_lt_up by (metis add.commute add.right_neutral diff_add_cancel nonzero_eq_divide_eq order_less_irrefl) qed locale CRR_market_viable = CRR_market + assumes CRR_viable: "viable_market Mkt" lemma (in CRR_market_viable) real_cond_exp_discount_stock_q_const: assumes "N = bernoulli_stream q" and "q = (1+r-d) / (u-d)" shows "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" proof - have qlt: "0 < q" and qgt: "q < 1" using assms gt_param lt_param CRR_viable by auto have "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" using assms real_cond_exp_discount_stock[of N q] qlt qgt by simp moreover have "\w. (q * u + (1 - q) * d) * prices Mkt stk n w = (1+r) * prices Mkt stk n w" using risk_neutral_param assms CRR_viable by (simp add: mult.commute) ultimately have "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w" by simp moreover have "\w\ space N. discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w = discounted_value r (\m w. prices Mkt stk n w) n w" using acceptable_rate by (simp add:discounted_mult_times_rfr) moreover hence "\w\ space N. discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w = discounted_value r (prices Mkt stk) n w" using acceptable_rate by (simp add:discounted_value_def) ultimately show "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" by simp qed lemma (in CRR_market_viable) risky_asset_martingale_if: assumes "N = bernoulli_stream q" and "q = (1 + r - d) / (u - d)" shows "martingale N G (discounted_value r (prices Mkt stk))" proof (rule disc_martingale_charact) have qlt: "0 < q" and qgt: "q < 1" using assms gt_param lt_param CRR_viable by auto show "\n. integrable N (discounted_value r (prices Mkt stk) n)" proof fix n show "integrable N (discounted_value r (prices Mkt stk) n)" proof (rule discounted_integrable) show "space N = space M" using assms by (simp add: bernoulli bernoulli_stream_space) show "integrable N (prices Mkt stk n)" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms qlt qgt by (simp add: infinite_coin_toss_space_def) show "prices Mkt stk n \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using geom_rand_walk_borel_adapted stk_price nat_filtration_from_eq_sets unfolding adapt_stoch_proc_def by (metis \infinite_coin_toss_space q N\ borel_measurable_integrable geom_proc_integrable geom_rand_walk_pseudo_proj_True infinite_coin_toss_space.nat_filtration_borel_measurable_characterization infinite_coin_toss_space_def) qed show "-1 < r" using acceptable_rate by simp qed qed show "filtration N G" using qlt qgt by (simp add: bernoulli_gen_filtration assms) show "\n. sigma_finite_subalgebra N (G n)" using qlt qgt by (simp add: assms bernoulli_sigma_finite) show "\m. discounted_value r (prices Mkt stk) m \ borel_measurable (G m)" proof fix m have "discounted_value r (\ma. prices Mkt stk m) m \ borel_measurable (G m)" proof (rule discounted_measurable) show "prices Mkt stk m \ borel_measurable (G m)" using stock_price_borel_measurable unfolding adapt_stoch_proc_def by simp qed thus "discounted_value r (prices Mkt stk) m \ borel_measurable (G m)" by (metis (mono_tags, lifting) discounted_value_def measurable_cong) qed show "\n. AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" proof fix n show "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" using assms real_cond_exp_discount_stock_q_const by simp qed qed lemma (in CRR_market_viable) risk_neutral_iff': assumes "N = bernoulli_stream q" and "0 \ q" and "q \ 1" and "filt_equiv nat_filtration M N" shows "rfr_disc_equity_market.risk_neutral_prob G Mkt r N \ q= (1 + r - d) / (u - d)" proof have "0 < q" and "q < 1" using assms filt_equiv_sgt filt_equiv_slt psgt pslt by auto note qprops = this have dem: "rfr_disc_equity_market M G Mkt r risk_free_asset" by unfold_locales { assume "rfr_disc_equity_market.risk_neutral_prob G Mkt r N" hence "(prob_space N) \ (\ asset \ stocks Mkt. martingale N G (discounted_value r (prices Mkt asset)))" using rfr_disc_equity_market.risk_neutral_prob_def[of M G Mkt] dem by simp hence "martingale N G (discounted_value r (prices Mkt stk))" using stocks by simp thus "q = (1 + r - d) / (u - d)" using assms risky_asset_martingale_only_if[of N q] qprops by simp } { assume "q = (1 + r - d) / (u - d)" hence "martingale N G (discounted_value r (prices Mkt stk))" using risky_asset_martingale_if[of N q] assms by simp moreover have "martingale N G (discounted_value r (prices Mkt risk_free_asset))" using risk_free_asset_martingale assms qprops by simp ultimately show "rfr_disc_equity_market.risk_neutral_prob G Mkt r N" using stocks using assms(1) bernoulli_stream_def dem prob_space.prob_space_stream_space prob_space_measure_pmf rfr_disc_equity_market.risk_neutral_prob_def by fastforce } qed lemma (in CRR_market_viable) risk_neutral_iff: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "rfr_disc_equity_market.risk_neutral_prob G Mkt r N \ q= (1 + r - d) / (u - d)" using bernoulli_stream_equiv assms risk_neutral_iff' psgt pslt by auto subsection \Existence of a replicating portfolio\ fun (in CRR_market) rn_rev_price where "rn_rev_price N der matur 0 w = der w" | "rn_rev_price N der matur (Suc n) w = discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w" lemma (in CRR_market) stock_filtration_eq: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "G n = stoch_proc_filt N geom_proc borel n" proof - have "G n= stoch_proc_filt M geom_proc borel n" using stock_filtration by simp also have "... = stoch_proc_filt N geom_proc borel n" proof (rule stoch_proc_filt_filt_equiv) show "filt_equiv nat_filtration M N" using assms bernoulli_stream_equiv psgt pslt by simp qed finally show ?thesis . qed lemma (in CRR_market) real_exp_eq: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) der w = expl_cond_expect N (proj_stoch_proc geom_proc n) der w" proof - have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel n) der w = expl_cond_expect N (proj_stoch_proc geom_proc n) der w" proof (rule bernoulli_cond_exp) show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto show "integrable N der" using \integrable N der\ . qed qed lemma (in CRR_market) rn_rev_price_rev_borel_adapt: assumes "cash_flow \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "(n \ matur) \ (rn_rev_price N cash_flow matur n) \ borel_measurable (G (matur - n))" proof (induct n) case 0 thus ?case using assms by simp next case (Suc n) have "rn_rev_price N cash_flow matur (Suc n) = (\w. discount_factor r (Suc 0) w * (expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N cash_flow matur n)) w)" using rn_rev_price.simps(2) by blast also have "... \ borel_measurable (G (matur - Suc n))" proof (rule borel_measurable_times) show "discount_factor r (Suc 0) \ borel_measurable (G (matur - Suc n))" by (simp add:discount_factor_borel_measurable) show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N cash_flow matur n) \ borel_measurable (G (matur - Suc n))" using assms by (simp add: bernoulli_expl_cond_expect_adapt) qed finally show "rn_rev_price N cash_flow matur (Suc n) \ borel_measurable (G (matur - Suc n))" . qed lemma (in infinite_coin_toss_space) bernoulli_discounted_integrable: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (nat_filtration n)" and "-1 < r" shows "integrable N (discounted_value r (\m. der) m)" proof - have "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using assms filt_equiv_filtration by (simp add: assms(1) measurable_def nat_filtration_from_eq_sets nat_filtration_space) qed thus ?thesis using discounted_integrable assms by (metis \prob_space N\ prob_space.discounted_integrable) qed lemma (in CRR_market) rn_rev_expl_cond_expect: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "n \ matur \ rn_rev_price N der matur n w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) w" proof (induct n arbitrary: w) case 0 have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed have "rn_rev_price N der matur 0 w = der w" by simp also have "... = expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) w" proof (rule nat_filtration_AE_eq) show "der \ borel_measurable (nat_filtration matur)" using \der \ borel_measurable (nat_filtration matur)\ . have "(discounted_value r (\m. der) 0) = der" unfolding discounted_value_def discount_factor_def by simp moreover have "AEeq N (real_cond_exp N (G matur) der) der" proof (rule sigma_finite_subalgebra.real_cond_exp_F_meas) show "der \ borel_measurable (G matur)" using assms by simp show "integrable N der" using \integrable N der\ . show "sigma_finite_subalgebra N (G matur)" using bernoulli_sigma_finite using assms by simp qed moreover have "\w. real_cond_exp N (stoch_proc_filt N geom_proc borel matur) der w = expl_cond_expect N (proj_stoch_proc geom_proc matur) der w" using assms real_exp_eq by simp ultimately have eqn: "AEeq N der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" using stock_filtration_eq assms by auto have "stoch_proc_filt M geom_proc borel matur = stoch_proc_filt N geom_proc borel matur" using bernoulli_stream_equiv[of N q] assms psgt pslt by (simp add: stoch_proc_filt_filt_equiv) also have "stoch_proc_filt N geom_proc borel matur = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur)" using assms geom_proc_stoch_proc_filt by simp finally have "stoch_proc_filt M geom_proc borel matur = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur)" . moreover have "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur))" proof (rule expl_cond_exp_borel) show "proj_stoch_proc geom_proc matur \ space N \ space (stream_space borel)" using assms proj_stoch_proc_geom_rng by (simp add: measurable_def) show "disc_fct (proj_stoch_proc geom_proc matur)" using proj_stoch_proc_geom_disc_fct by simp show "\r\range (proj_stoch_proc geom_proc matur) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc matur) \ A = {r}" using proj_stoch_proc_geom_open_set by simp qed ultimately show ebm: "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (nat_filtration matur)" by (metis geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt) show "AEeq M der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" proof (rule filt_equiv_borel_AE_eq_iff[THEN iffD2]) show "filt_equiv nat_filtration M N" using assms bernoulli_stream_equiv psgt pslt by simp show "der \ borel_measurable (nat_filtration matur)" using \der \ borel_measurable (nat_filtration matur)\ . show "AEeq N der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" using eqn . show "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (nat_filtration matur)" using ebm . show "prob_space N" using assms by (simp add: bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) show "prob_space M" by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) qed show "0 < p" "p < 1" using psgt pslt by auto qed also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - 0)) (discounted_value r (\m. der) 0) w" by simp finally show "rn_rev_price N der matur 0 w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - 0)) (discounted_value r (\m. der) 0) w" . next case (Suc n) have "rn_rev_price N der matur (Suc n) w = discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w" by simp also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w" proof - have "expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w" proof (rule real_exp_eq[symmetric]) show "rn_rev_price N der matur n \ borel_measurable (G (matur - n))" using assms rn_rev_price_rev_borel_adapt Suc by simp show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto qed thus ?thesis by simp qed also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "rn_rev_price N der matur n \ borel_measurable N" proof - have "rn_rev_price N der matur n \ borel_measurable (G (matur - n))" by (metis (full_types) Suc.prems Suc_leD assms(1) assms(2) assms(3) assms(4) rn_rev_price_rev_borel_adapt) then show ?thesis by (metis (no_types) assms(2) bernoulli bernoulli_stream_def filtration_measurable measurable_cong_sets sets_measure_pmf sets_stream_space_cong) qed show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" using Suc.hyps Suc.prems Suc_leD \rn_rev_price N der matur n \ borel_measurable N\ by presburger show "AEeq N (rn_rev_price N der matur n) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n))" using Suc by auto qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) qed thus ?thesis by simp qed also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n))) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" by simp show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" by (metis assms(2) assms(3) assms(4) bernoulli bernoulli_expl_cond_expect_adapt bernoulli_stream_def filtration_measurable measurable_cong_sets sets_measure_pmf sets_stream_space_cong) show "AEeq N (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))" proof - have "discounted_value r (\m. der) n \ borel_measurable (G matur)" using assms discounted_measurable[of der] by simp hence "\w. (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w = (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" using real_exp_eq[of _ matur N q "matur-n"] assms by simp thus ?thesis by simp qed qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) qed thus ?thesis by simp qed also have "... = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "(\a. discount_factor r (Suc 0) a * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) a) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) thus ?thesis using discounted_measurable[of "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))"] unfolding discounted_value_def by simp qed show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)))" proof- have "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w) (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n) w)" proof - have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n))" proof (rule sigma_finite_subalgebra.real_cond_exp_nested_subalg) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce show "subalgebra (stoch_proc_filt N geom_proc borel (matur - n)) (stoch_proc_filt N geom_proc borel (matur - Suc n))" proof - have "init_triv_filt M (stoch_proc_filt M geom_proc borel)" using infinite_cts_filtration.stoch_proc_filt_triv_init using info_filtration stock_filtration by auto moreover have "matur - (Suc n) \ matur - n" by simp ultimately show ?thesis unfolding init_triv_filt_def filtration_def using assms(2) assms(3) assms(4) stock_filtration stock_filtration_eq by auto qed show "integrable N (discounted_value r (\m. der) n) " using bernoulli_discounted_integrable[of N q der matur r n] acceptable_rate assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast qed thus ?thesis by auto qed moreover have "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n) w) (\w. discount_factor r (Suc 0) w * (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n) w)" proof - have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n)) (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n)" proof (rule prob_space.discounted_value_real_cond_exp) show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "-1 < r" using acceptable_rate . show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce qed thus ?thesis by auto qed moreover have "\w. (\w. discount_factor r (Suc 0) w * (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n) w) w = (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) (Suc n)) w" unfolding discounted_value_def discount_factor_def by simp moreover have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n))) (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) (Suc n))" proof (rule prob_space.discounted_value_real_cond_exp) show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "-1 < r" using acceptable_rate . show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce qed ultimately show ?thesis by auto qed qed also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" proof (rule real_exp_eq) show "discounted_value r (\m. der) (Suc n) \ borel_measurable (G matur)" using assms discounted_measurable[of der] by simp show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto qed finally show "rn_rev_price N der matur (Suc n) w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" . qed definition (in CRR_market) rn_price where "rn_price N der matur n w = expl_cond_expect N (proj_stoch_proc geom_proc n) (discounted_value r (\m. der) (matur - n)) w" definition (in CRR_market) rn_price_ind where "rn_price_ind N der matur n w = rn_rev_price N der matur (matur - n) w" lemma (in CRR_market) rn_price_eq: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" and "n \ matur" shows "rn_price N der matur n w = rn_price_ind N der matur n w" using rn_rev_expl_cond_expect unfolding rn_price_def rn_price_ind_def by (simp add: assms) lemma (in CRR_market) geom_proc_filt_info: fixes f::"bool stream \ 'b::{t0_space}" assumes "f \ borel_measurable (G n)" shows "f w = f (pseudo_proj_True n w)" proof - have "subalgebra (nat_filtration n) (G n)" using stoch_proc_subalg_nat_filt[of geom_proc n] geometric_process stock_filtration geom_rand_walk_borel_adapted by simp hence "f\ borel_measurable (nat_filtration n)" using assms by (simp add: measurable_from_subalg) thus ?thesis using nat_filtration_info[of f n] by (metis comp_apply) qed lemma (in CRR_market) geom_proc_filt_info': fixes f::"bool stream \ 'b::{t0_space}" assumes "f \ borel_measurable (G n)" shows "f w = f (pseudo_proj_False n w)" proof - have "subalgebra (nat_filtration n) (G n)" using stoch_proc_subalg_nat_filt[of geom_proc n] geometric_process stock_filtration geom_rand_walk_borel_adapted by simp hence "f\ borel_measurable (nat_filtration n)" using assms by (simp add: measurable_from_subalg) thus ?thesis using nat_filtration_info'[of f n] by (metis comp_apply) qed lemma (in CRR_market) rn_price_borel_adapt: assumes "cash_flow \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" and "n \ matur" shows "(rn_price N cash_flow matur n) \ borel_measurable (G n)" proof - show "(rn_price N cash_flow matur n) \ borel_measurable (G n)" using assms rn_rev_price_rev_borel_adapt[of cash_flow matur N q "matur - n"] rn_price_eq rn_price_ind_def by (smt add.right_neutral cancel_comm_monoid_add_class.diff_cancel diff_commute diff_le_self increasing_measurable_info measurable_cong nat_le_linear ordered_cancel_comm_monoid_diff_class.add_diff_inverse) qed definition (in CRR_market) delta_price where "delta_price N cash_flow T = (\ n w. if (Suc n \ T) then (rn_price N cash_flow T (Suc n) (pseudo_proj_True n w) - rn_price N cash_flow T (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) else 0)" lemma (in CRR_market) delta_price_eq: assumes "Suc n \ T" shows "delta_price N cash_flow T n w = (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d))" proof - - have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = u * geom_proc n w - d * geom_proc n w" - using geometric_process - by (metis geom_rand_walk_diff_induct linordered_field_class.sign_simps(39) mult.commute) - also have "... = geom_proc n w * (u - d)" - by (simp add: linordered_field_class.sign_simps(39) mult.commute) - finally have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = geom_proc n w * (u - d)" . + have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = geom_proc n w * (u - d)" + by (simp add: geom_rand_walk_diff_induct) then show ?thesis unfolding delta_price_def using assms spick_eq_pseudo_proj_True spick_eq_pseudo_proj_False by simp qed lemma (in CRR_market) geom_proc_spick: shows "geom_proc (Suc n) (spick w n x) = (if x then u else d) * geom_proc n w" proof - have "geom_proc (Suc n) (spick w n x) = geom_rand_walk u d init (Suc n) (spick w n x)" using geometric_process by simp also have "... = (case (spick w n x) !! n of True \ u | False \ d) * geom_rand_walk u d init n (spick w n x)" by simp also have "... = (case x of True \ u | False \ d) * geom_rand_walk u d init n (spick w n x)" unfolding spick_def by simp also have "... = (if x then u else d) * geom_rand_walk u d init n (spick w n x)" by simp also have "... = (if x then u else d) * geom_rand_walk u d init n w" by (metis comp_def geom_rand_walk_pseudo_proj_True geometric_process pseudo_proj_True_stake_image spickI) finally show ?thesis using geometric_process by simp qed lemma (in CRR_market) spick_red_geom: shows "(\w. spick w n x) \ measurable (fct_gen_subalgebra M borel (geom_proc n)) (fct_gen_subalgebra M borel (geom_proc (Suc n)))" unfolding measurable_def proof (intro CollectI conjI) show "(\w. spick w n x) \ space (fct_gen_subalgebra M borel (geom_proc n)) \ space (fct_gen_subalgebra M borel (geom_proc (Suc n)))" by (simp add: bernoulli bernoulli_stream_space fct_gen_subalgebra_space) show "\y\sets (fct_gen_subalgebra M borel (geom_proc (Suc n))). (\w. spick w n x) -` y \ space (fct_gen_subalgebra M borel (geom_proc n)) \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof fix A assume A: "A \ sets (fct_gen_subalgebra M borel (geom_proc (Suc n)))" show "(\w. spick w n x) -` A \ space (fct_gen_subalgebra M borel (geom_proc n)) \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof - define sp where "sp = (\w. spick w n x)" have "A \ {(geom_proc (Suc n)) -` B \ space M |B. B \ sets borel}" using A by (simp add:fct_gen_subalgebra_sigma_sets) from this obtain C where "C\ sets borel" and "A = (geom_proc (Suc n)) -`C \ space M" by auto hence "A = (geom_proc (Suc n)) -`C" using bernoulli bernoulli_stream_space by simp hence "sp -`A = sp -` (geom_proc (Suc n)) -`C" by simp also have "... = (geom_proc (Suc n) \ sp) -` C" by auto also have "... = (\w. (if x then u else d) * geom_proc n w) -` C" using geom_proc_spick sp_def by auto also have "... \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof (cases x) case True hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. u * geom_proc n w) -` C" by simp moreover have "(\w. u * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ by (metis bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space measurable_sets) next case False hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. d * geom_proc n w) -` C" by simp moreover have "(\w. d * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ by (metis bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space measurable_sets) qed finally show ?thesis unfolding sp_def by (simp add: bernoulli bernoulli_stream_space fct_gen_subalgebra_space) qed qed qed lemma (in CRR_market) geom_spick_Suc: assumes "A \ {(geom_proc (Suc n)) -` B |B. B \ sets borel}" shows "(\w. spick w n x) -`A \ {geom_proc n -`B | B. B\ sets borel}" proof - have "sets (fct_gen_subalgebra M borel (geom_proc n)) = {geom_proc n -` B \space M |B. B \ sets borel}" by (simp add: fct_gen_subalgebra_sigma_sets) also have "... = {geom_proc n -` B |B. B \ sets borel}" using bernoulli bernoulli_stream_space by simp finally have sf: "sets (fct_gen_subalgebra M borel (geom_proc n)) = {geom_proc n -` B |B. B \ sets borel}" . define sp where "sp = (\w. spick w n x)" from assms(1) obtain C where "C\ sets borel" and "A = (geom_proc (Suc n)) -`C" by auto hence "A = (geom_proc (Suc n)) -`C" using bernoulli bernoulli_stream_space by simp hence "sp -`A = sp -` (geom_proc (Suc n)) -`C" by simp also have "... = (geom_proc (Suc n) \ sp) -` C" by auto also have "... = (\w. (if x then u else d) * geom_proc n w) -` C" using geom_proc_spick sp_def by auto also have "... \ {geom_proc n -`B | B. B\ sets borel}" proof (cases x) case True hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. u * geom_proc n w) -` C" by simp moreover have "(\w. u * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ sf by (simp add: bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space in_borel_measurable_borel) next case False hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. d * geom_proc n w) -` C" by simp moreover have "(\w. d * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ sf by (simp add: bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space in_borel_measurable_borel) qed finally show ?thesis unfolding sp_def . qed lemma (in CRR_market) geom_spick_lt: assumes "m< n" shows "geom_proc m (spick w n x) = geom_proc m w" proof - have "geom_proc m (spick w n x) = geom_proc m (pseudo_proj_True m (spick w n x))" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_True m w)" using assms by (metis less_imp_le_nat pseudo_proj_True_def pseudo_proj_True_prefix spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) finally show ?thesis . qed lemma (in CRR_market) geom_spick_eq: shows "geom_proc m (spick w m x) = geom_proc m w" proof (cases x) case True have "geom_proc m (spick w m x) = geom_proc m (pseudo_proj_True m (spick w m x))" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_True m w)" using True by (metis pseudo_proj_True_def spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) finally show ?thesis . next case False have "geom_proc m (spick w m x) = geom_proc m (pseudo_proj_False m (spick w m x))" using geom_rand_walk_pseudo_proj_False by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_False m w)" using False by (metis pseudo_proj_False_def spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_False by (metis comp_apply) finally show ?thesis . qed lemma (in CRR_market) spick_red_geom_filt: shows "(\w. spick w n x) \ measurable (G n) (G (Suc n))" unfolding measurable_def proof (intro CollectI conjI) show "(\w. spick w n x) \ space (G n) \ space (G (Suc n))" using stock_filtration by (simp add: bernoulli bernoulli_stream_space stoch_proc_filt_space) show "\y\sets (G (Suc n)). (\w. spick w n x) -` y \ space (G n) \ sets (G n)" proof fix B assume "B\ sets (G (Suc n))" hence "B\ (sigma_sets (space M) (\ i\ {m. m\ (Suc n)}. {(geom_proc i -`A) \ (space M) | A. A\ sets borel }))" using stock_filtration stoch_proc_filt_sets geometric_process proof - have "\n. sigma_sets (space M) (\n\{na. na \ n}. {geom_proc n -` R \ space M |R. R \ sets borel}) = sets (G n)" by (simp add: geom_rand_walk_borel_measurable stoch_proc_filt_sets stock_filtration) then show ?thesis using \B \ sets (G (Suc n))\ by blast qed hence "(\w. spick w n x) -` B \ sets (G n)" proof (induct rule:sigma_sets.induct) { fix C assume "C \ (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel})" hence "\m \ Suc n. C\ {geom_proc m -` A \ space M |A. A \ sets borel}" by auto from this obtain m where "m\ Suc n" and "C\ {geom_proc m -` A \ space M |A. A \ sets borel}" by auto note Cprops = this from this obtain D where "C = geom_proc m -` D\ space M" and "D\ sets borel" by auto hence "C = geom_proc m -`D" using bernoulli bernoulli_stream_space by simp have "C\ {geom_proc m -` A |A. A \ sets borel}" using bernoulli bernoulli_stream_space Cprops by simp show "(\w. spick w n x) -` C \ sets (G n)" proof (cases "m \ n") case True have "(\w. spick w n x) -` C = (\w. spick w n x) -` geom_proc m -`D" using \C = geom_proc m -`D\ by simp also have "... = (geom_proc m \ (\w. spick w n x)) -`D" by auto also have "... = geom_proc m -`D" using geom_spick_lt geom_spick_eq \m\n\ using le_eq_less_or_eq by auto also have "... \ sets (G n)" using stock_filtration geometric_process \D\ sets borel\ by (metis (no_types, lifting) True adapt_stoch_proc_def bernoulli bernoulli_stream_preimage geom_rand_walk_borel_measurable increasing_measurable_info measurable_sets stoch_proc_filt_adapt stoch_proc_filt_space) finally show "(\w. spick w n x) -` C \ sets (G n)" . next case False hence "m = Suc n" using \m \ Suc n\ by simp hence "(\w. spick w n x) -` C \ {geom_proc n -` B |B. B \ sets borel}" using \C\ {geom_proc m -` A |A. A \ sets borel}\ geom_spick_Suc by simp also have "... \ sets (G n)" proof - have "{geom_proc n -` B |B. B \ sets borel} \ {geom_proc n -` B \ space M |B. B \ sets borel}" using bernoulli bernoulli_stream_space by simp also have "... \ (\i\{m. m \ n}. {geom_proc i -` A \ space M |A. A \ sets borel})" by auto also have "... \ sigma_sets (space M) (\i\{m. m \ n}. {geom_proc i -` A \ space M |A. A \ sets borel})" by (rule sigma_sets_superset_generator) also have "... = sets (G n)" using stock_filtration geometric_process stoch_proc_filt_sets[of n geom_proc M borel] geom_rand_walk_borel_measurable by blast finally show ?thesis . qed finally show ?thesis . qed } show "(\w. spick w n x) -` {} \ sets (G n)" by simp { fix C assume "C \ sigma_sets (space M) (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel})" and "(\w. spick w n x) -` C \ sets (G n)" hence "(\w. spick w n x) -` (space M - C) = (\w. spick w n x) -` (space M) - (\w. spick w n x) -` C" by (simp add: vimage_Diff) also have "... = space M - (\w. spick w n x) -` C" using bernoulli bernoulli_stream_space by simp also have "... \ sets (G n)" using \(\w. spick w n x) -` C \ sets (G n)\ by (metis algebra.compl_sets disc_filtr_def discrete_filtration sets.sigma_algebra_axioms sigma_algebra_def subalgebra_def) finally show "(\w. spick w n x) -` (space M - C) \ sets (G n)" . } { fix C::"nat \ bool stream set" assume "(\i. C i \ sigma_sets (space M) (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel}))" and "(\i. (\w. spick w n x) -` C i \ sets (G n))" hence "(\w. spick w n x) -` \(C ` UNIV) = (\ i\ UNIV. (\w. spick w n x) -` (C i))" by blast also have "... \ sets (G n)" using \\i. (\w. spick w n x) -` C i \ sets (G n)\ by simp finally show "(\w. spick w n x) -` \(C ` UNIV) \ sets (G n)" . } qed thus "(\w. spick w n x) -` B \ space (G n) \ sets (G n)" using stock_filtration stoch_proc_filt_space bernoulli bernoulli_stream_space by simp qed qed lemma (in CRR_market) delta_price_adapted: fixes cash_flow::"bool stream \ real" assumes "cash_flow \ borel_measurable (G T)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "borel_adapt_stoch_proc G (delta_price N cash_flow T)" unfolding adapt_stoch_proc_def proof fix n show "delta_price N cash_flow T n \ borel_measurable (G n)" proof (cases "Suc n \ T") case True hence deleq: "\w. delta_price N cash_flow T n w = (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d))" using delta_price_eq by simp have "(\w. rn_price N cash_flow T (Suc n) (spick w n True)) \ borel_measurable (G n)" proof - have "rn_price N cash_flow T (Suc n) \ borel_measurable (G (Suc n))" using rn_price_borel_adapt assms using True by blast moreover have "(\w. spick w n True) \ G n \\<^sub>M G (Suc n)" using spick_red_geom_filt by simp ultimately show ?thesis by simp qed moreover have "(\w. rn_price N cash_flow T (Suc n) (spick w n False)) \ borel_measurable (G n)" proof - have "rn_price N cash_flow T (Suc n) \ borel_measurable (G (Suc n))" using rn_price_borel_adapt assms using True by blast moreover have "(\w. spick w n False) \ G n \\<^sub>M G (Suc n)" using spick_red_geom_filt by simp ultimately show ?thesis by simp qed ultimately have "(\w. rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False)) \ borel_measurable (G n)" by simp moreover have "(\w. (geom_proc n w) * (u - d)) \ borel_measurable (G n)" proof - have "geom_proc n \ borel_measurable (G n)" using stock_filtration by (metis adapt_stoch_proc_def stk_price stock_price_borel_measurable) thus ?thesis by simp qed ultimately have "(\w. (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d)))\ borel_measurable (G n)" by simp thus ?thesis using deleq by presburger next case False thus ?thesis unfolding delta_price_def by simp qed qed fun (in CRR_market) delta_predict where "delta_predict N der matur 0 = (\w. delta_price N der matur 0 w)" | "delta_predict N der matur (Suc n) = (\w. delta_price N der matur n w)" lemma (in CRR_market) delta_predict_predict: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "borel_predict_stoch_proc G (delta_predict N der matur)" unfolding predict_stoch_proc_def proof (intro conjI) show "delta_predict N der matur 0 \ borel_measurable (G 0)" using delta_price_adapted[of der matur N q] assms unfolding adapt_stoch_proc_def by force show "\n. delta_predict N der matur (Suc n) \ borel_measurable (G n)" proof fix n show "delta_predict N der matur (Suc n) \ borel_measurable (G n)" using delta_price_adapted[of der matur N q] assms unfolding adapt_stoch_proc_def by force qed qed definition (in CRR_market) delta_pf where "delta_pf N der matur = qty_single stk (delta_predict N der matur)" lemma (in CRR_market) delta_pf_support: shows "support_set (delta_pf N der matur) \ {stk}" unfolding delta_pf_def using single_comp_support[of stk "delta_predict N der matur"] by simp definition (in CRR_market) self_fin_delta_pf where "self_fin_delta_pf N der matur v0 = self_finance Mkt v0 (delta_pf N der matur) risk_free_asset" lemma (in disc_equity_market) self_finance_trading_strat: assumes "trading_strategy pf" and "portfolio pf" and "borel_adapt_stoch_proc F (prices Mkt asset)" and "support_adapt Mkt pf" shows "trading_strategy (self_finance Mkt v pf asset)" unfolding self_finance_def proof (rule sum_trading_strat) show "trading_strategy pf" using assms by simp show "trading_strategy (qty_single asset (remaining_qty Mkt v pf asset))" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio (qty_single asset (remaining_qty Mkt v pf asset))" by (simp add: self_finance_def single_comp_portfolio) show "\a. a \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) \ borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" proof (cases "support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {}") case False hence eqasset: "support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {asset}" using single_comp_support by fastforce fix a assume "a\ support_set (qty_single asset (remaining_qty Mkt v pf asset))" hence "a = asset" using eqasset by simp hence "qty_single asset (remaining_qty Mkt v pf asset) a = (remaining_qty Mkt v pf asset)" unfolding qty_single_def by simp moreover have "borel_predict_stoch_proc F (remaining_qty Mkt v pf asset)" proof (rule remaining_qty_predict) show "trading_strategy pf" using assms by simp show "borel_adapt_stoch_proc F (prices Mkt asset)" using assms by simp show "support_adapt Mkt pf" using assms by simp qed ultimately show "borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" by simp next case True thus "\a. a \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {} \ borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" by simp qed qed qed lemma (in CRR_market) self_fin_delta_pf_trad_strat: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "trading_strategy (self_fin_delta_pf N der matur v0)" unfolding self_fin_delta_pf_def proof (rule self_finance_trading_strat) show "trading_strategy (delta_pf N der matur)" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio (delta_pf N der matur)" unfolding portfolio_def using delta_pf_support by (meson finite.emptyI finite_insert infinite_super) show "\asset. asset \ support_set (delta_pf N der matur) \ borel_predict_stoch_proc G (delta_pf N der matur asset)" proof (cases "support_set (delta_pf N der matur) = {}") case False fix asset assume "asset \ support_set (delta_pf N der matur)" hence "asset = stk" using False delta_pf_support by auto hence "delta_pf N der matur asset = delta_predict N der matur" unfolding delta_pf_def qty_single_def by simp thus "borel_predict_stoch_proc G (delta_pf N der matur asset)" using delta_predict_predict assms by simp next case True thus "\asset. asset \ support_set (delta_pf N der matur) \ support_set (delta_pf N der matur) = {} \ borel_predict_stoch_proc G (delta_pf N der matur asset)" by simp qed qed show "portfolio (delta_pf N der matur)" using delta_pf_support unfolding portfolio_def by (meson finite.emptyI finite_insert infinite_super) show "borel_adapt_stoch_proc G (prices Mkt risk_free_asset)" using rf_price disc_rfr_proc_borel_adapted by simp show "support_adapt Mkt (delta_pf N der matur)" unfolding support_adapt_def proof show "\asset. asset \ support_set (delta_pf N der matur) \ borel_adapt_stoch_proc G (prices Mkt asset)" proof (cases "support_set (delta_pf N der matur) = {}") case False fix asset assume "asset \ support_set (delta_pf N der matur)" hence "asset = stk" using False delta_pf_support by auto hence "prices Mkt asset = geom_proc" using stk_price by simp thus "borel_adapt_stoch_proc G (prices Mkt asset)" using \asset = stk\ stock_price_borel_measurable by auto next case True thus "\asset. asset \ support_set (delta_pf N der matur) \ borel_adapt_stoch_proc G (prices Mkt asset)" by simp qed qed qed definition (in CRR_market) delta_hedging where "delta_hedging N der matur = self_fin_delta_pf N der matur (prob_space.expectation N (discounted_value r (\m. der) matur))" lemma (in CRR_market) geom_proc_eq_snth: shows "(\m. m \ Suc n \ geom_proc m x = geom_proc m y) \ (\m. m \ n \ snth x m = snth y m)" proof (induct n ) case 0 assume asm: "(\m. m \Suc 0 \ geom_proc m x = geom_proc m y)" and "m\ 0" hence "m = 0" by simp have "geom_proc (Suc 0) x = geom_proc (Suc 0) y" using asm by simp have "snth x 0 = snth y 0" proof (rule ccontr) assume "snth x 0 \ snth y 0" show False proof (cases "snth x 0") case True hence "\ snth y 0" using \snth x 0 \ snth y 0\ by simp have "geom_proc (Suc 0) x = u * init" using geometric_process True by simp moreover have "geom_proc (Suc 0) y = d * init" using geometric_process \\ snth y 0\ by simp ultimately have "geom_proc (Suc 0) x \ geom_proc (Suc 0) y" using S0_positive down_lt_up by simp thus ?thesis using \geom_proc (Suc 0) x = geom_proc (Suc 0) y\ by simp next case False hence "snth y 0" using \snth x 0 \ snth y 0\ by simp have "geom_proc (Suc 0) x = d * init" using geometric_process False by simp moreover have "geom_proc (Suc 0) y = u * init" using geometric_process \snth y 0\ by simp ultimately have "geom_proc (Suc 0) x \ geom_proc (Suc 0) y" using S0_positive down_lt_up by simp thus ?thesis using \geom_proc (Suc 0) x = geom_proc (Suc 0) y\ by simp qed qed thus "\m. (\m. m \ Suc 0 \ geom_proc m x = geom_proc m y) \ m \ 0 \ x !! m = y !! m" by simp next case (Suc n) assume fst: "(\m. (\m. m \ Suc n \ geom_proc m x = geom_proc m y) \ m \ n \ x !! m = y !! m)" and scd: "(\m. m \ Suc (Suc n) \ geom_proc m x = geom_proc m y)" and "m \ Suc n" show "x !! m = y !! m" proof (cases "m \ n") case True thus ?thesis using fst scd by simp next case False hence "m = Suc n" using \m\ Suc n\ by simp have "geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y" using scd by simp show ?thesis proof (rule ccontr) assume "x !! m \ y !! m" thus False proof (cases "x !! m") case True hence "\ y !! m" using \x !! m \ y !! m\ by simp have "geom_proc (Suc (Suc n)) x = u * geom_proc (Suc n) x" using geometric_process True \m = Suc n\ by simp also have "... = u * geom_proc (Suc n) y" using scd \m = Suc n\ by simp finally have "geom_proc (Suc (Suc n)) x = u * geom_proc (Suc n) y" . moreover have "geom_proc (Suc (Suc n)) y = d * geom_proc (Suc n) y" using geometric_process \m = Suc n\ \\ y !! m\ by simp ultimately have "geom_proc (Suc (Suc n)) x \ geom_proc (Suc (Suc n)) y" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl mult_cancel_right) thus ?thesis using \geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y\ by simp next case False hence "y !! m" using \x !! m \ y !! m\ by simp have "geom_proc (Suc (Suc n)) x = d * geom_proc (Suc n) x" using geometric_process False \m = Suc n\ by simp also have "... = d * geom_proc (Suc n) y" using scd \m = Suc n\ by simp finally have "geom_proc (Suc (Suc n)) x = d * geom_proc (Suc n) y" . moreover have "geom_proc (Suc (Suc n)) y = u * geom_proc (Suc n) y" using geometric_process \m = Suc n\ \y !! m\ by simp ultimately have "geom_proc (Suc (Suc n)) x \ geom_proc (Suc (Suc n)) y" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl mult_cancel_right) thus ?thesis using \geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y\ by simp qed qed qed qed lemma (in CRR_market) geom_proc_eq_pseudo_proj_True: shows "(\m. m \ n \ geom_proc m x = geom_proc m y) \ (pseudo_proj_True (n) x = pseudo_proj_True (n) y)" proof - assume a1: "\m. m \ n \ geom_proc m x = geom_proc m y" obtain nn :: "bool stream \ bool stream \ nat \ nat" where "\x1 x2 x3. (\v4 geom_proc v4 x1) = (nn x1 x2 x3 < Suc (Suc x3) \ geom_proc (nn x1 x2 x3) x2 \ geom_proc (nn x1 x2 x3) x1)" by moura then have f2: "\n s sa na. (nn sa s n < Suc (Suc n) \ geom_proc (nn sa s n) s \ geom_proc (nn sa s n) sa \ \ na < Suc n) \ s !! na = sa !! na" by (meson geom_proc_eq_snth less_Suc_eq_le) obtain nna :: "bool stream \ bool stream \ nat \ nat" where f3: "\x0 x1 x2. (\v3. Suc v3 < Suc x2 \ x1 !! v3 \ x0 !! v3) = (Suc (nna x0 x1 x2) < Suc x2 \ x1 !! nna x0 x1 x2 \ x0 !! nna x0 x1 x2)" by moura obtain nnb :: "nat \ nat" where f4: "\x0. (\v2. x0 = Suc v2) = (x0 = Suc (nnb x0))" by moura moreover { assume "\ nn y x (nnb n) < Suc (Suc (nnb n)) \ geom_proc (nn y x (nnb n)) x = geom_proc (nn y x (nnb n)) y" moreover { assume "\ nna y x n < Suc (nnb n)" then have "\ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" using f4 by (metis (no_types) Suc_le_D Suc_le_lessD less_Suc_eq_le) } ultimately have "pseudo_proj_True n x = pseudo_proj_True n y \ \ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" using f2 by meson } ultimately have "pseudo_proj_True n x = pseudo_proj_True n y \ \ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" using a1 Suc_le_D less_Suc_eq_le by presburger then show ?thesis using f3 by (meson less_Suc_eq_le pseudo_proj_True_snth') qed lemma (in CRR_market) proj_stoch_eq_pseudo_proj_True: assumes "proj_stoch_proc geom_proc m x = proj_stoch_proc geom_proc m y" shows "pseudo_proj_True m x = pseudo_proj_True m y" proof - have "\ k \ m. geom_proc k x = geom_proc k y" proof (intro allI impI) fix k assume "k \ m" thus "geom_proc k x = geom_proc k y" using proj_stoch_proc_eq_snth[of geom_proc m x y k] assms by simp qed thus ?thesis using geom_proc_eq_pseudo_proj_True[of m x y] by auto qed lemma (in CRR_market_viable) rn_rev_price_cond_expect: assumes "N = bernoulli_stream q" and "0 borel_measurable (G matur)" and "Suc n \ matur" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) (rn_rev_price N der matur (matur - Suc n)) w= (q * rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True n w) + (1 - q) * rn_rev_price N der matur (matur - Suc n) (pseudo_proj_False n w))" proof (rule infinite_cts_filtration.f_borel_Suc_expl_cond_expect) show "infinite_cts_filtration q N nat_filtration" using assms pslt psgt bernoulli_nat_filtration by simp show "rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))" using rn_rev_price_rev_borel_adapt[of der matur N q "Suc n"] assms stock_filtration stoch_proc_subalg_nat_filt[of geom_proc] geom_rand_walk_borel_adapted by (metis add_diff_cancel_right' diff_le_self measurable_from_subalg ordered_cancel_comm_monoid_diff_class.add_diff_inverse rn_rev_price_rev_borel_adapt) show "proj_stoch_proc geom_proc n \ nat_filtration n \\<^sub>M stream_space borel" using proj_stoch_adapted_if_adapted[of M nat_filtration geom_proc borel n] pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted nat_discrete_filtration by blast show "set_discriminating n (proj_stoch_proc geom_proc n) (stream_space borel)" using infinite_cts_filtration.proj_stoch_set_discriminating pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted by simp show "proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (nat_filtration n)" using infinite_cts_filtration.proj_stoch_singleton_set pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted by simp show "\y z. proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n \ rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) z" proof (intro allI impI) fix y z assume as:"proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n" hence "pseudo_proj_True n y = pseudo_proj_True n z" using proj_stoch_eq_pseudo_proj_True[of n y z] by simp moreover have "snth y n = snth z n" using as by simp ultimately have "pseudo_proj_True (Suc n) y = pseudo_proj_True (Suc n) z" proof - have f1: "\n s sa. (\na. Suc na \ n \ s !! na \ sa !! na) \ pseudo_proj_True n s = pseudo_proj_True n sa" by (meson pseudo_proj_True_snth') obtain nn :: "bool stream \ bool stream \ nat \ nat" where "\x0 x1 x2. (\v3. Suc v3 \ x2 \ x1 !! v3 \ x0 !! v3) = (Suc (nn x0 x1 x2) \ x2 \ x1 !! nn x0 x1 x2 \ x0 !! nn x0 x1 x2)" by moura then have f2: "\n s sa. Suc (nn sa s n) \ n \ s !! nn sa s n \ sa !! nn sa s n \ pseudo_proj_True n s = pseudo_proj_True n sa" using f1 by presburger have f3: "stake n y = stake n (pseudo_proj_True n z)" by (metis \pseudo_proj_True n y = pseudo_proj_True n z\ pseudo_proj_True_stake) { assume "stake (Suc n) z \ stake (Suc n) (pseudo_proj_True (Suc n) y)" then have "stake n y @ [y !! n] \ stake n z @ [z !! n]" by (metis (no_types) pseudo_proj_True_stake stake_Suc) then have "stake (Suc n) z = stake (Suc n) (pseudo_proj_True (Suc n) y)" using f3 by (simp add: \y !! n = z !! n\ pseudo_proj_True_stake) } then have "\ Suc (nn z y (Suc n)) \ Suc n \ y !! nn z y (Suc n) = z !! nn z y (Suc n)" by (metis (no_types) pseudo_proj_True_stake stake_snth) then show ?thesis using f2 by blast qed have "rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True (Suc n) y)" using nat_filtration_info[of "rn_rev_price N der matur (matur - Suc n)" "Suc n"] rn_rev_price_rev_borel_adapt[of der matur N q] by (metis \rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))\ o_apply) also have "... = rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True (Suc n) z)" using \pseudo_proj_True (Suc n) y = pseudo_proj_True (Suc n) z\ by simp also have "... = rn_rev_price N der matur (matur - Suc n) z" using nat_filtration_info[of "rn_rev_price N der matur (matur - Suc n)" "Suc n"] rn_rev_price_rev_borel_adapt[of der matur N q] by (metis \rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))\ o_apply) finally show "rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) z" . qed show "0 < q" and "q < 1" using assms by auto qed lemma (in CRR_market_viable) rn_price_eq_ind: assumes "N = bernoulli_stream q" and "n < matur" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" shows "(1+r) * rn_price N der matur n w = q * rn_price N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price N der matur (Suc n) (pseudo_proj_False n w)" proof - define V where "V = rn_price N der matur" let ?m = "matur - Suc n" have "matur -n = Suc ?m" by (simp add: assms Suc_diff_Suc Suc_le_lessD) have "(1+r) * V n w = (1+r) * rn_price_ind N der matur n w" using rn_price_eq assms unfolding V_def by simp also have "... = (1+r) * rn_rev_price N der matur (Suc ?m) w" using \matur -n = Suc ?m\ unfolding rn_price_ind_def by simp also have "... = (1+r) * discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc ?m)) (rn_rev_price N der matur ?m) w" by simp also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc ?m)) (rn_rev_price N der matur ?m) w" unfolding discount_factor_def using acceptable_rate by auto also have "... = expl_cond_expect N (proj_stoch_proc geom_proc n) (rn_rev_price N der matur ?m) w" using \matur -n = Suc ?m\ by simp also have "... = (q * rn_rev_price N der matur ?m (pseudo_proj_True n w) + (1 - q) * rn_rev_price N der matur ?m (pseudo_proj_False n w))" using rn_rev_price_cond_expect[of N q der matur n w] assms by simp also have "... = q * rn_price_ind N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price_ind N der matur (Suc n) (pseudo_proj_False n w)" unfolding rn_price_ind_def by simp also have "... = q * rn_price N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price N der matur (Suc n) (pseudo_proj_False n w)" using rn_price_eq assms by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" unfolding V_def by simp finally have "(1+r) * V n w = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" . thus ?thesis unfolding V_def by simp qed lemma self_finance_updated_suc_suc: assumes "portfolio pf" and "\n. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc n)) w = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w - val_process Mkt pf (Suc n) w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc n)) w = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * remaining_qty Mkt v pf asset (Suc (Suc n)) w" using assms by (simp add: self_finance_updated) also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((remaining_qty Mkt v pf asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) / (prices Mkt asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" using assms by (metis nonzero_mult_div_cancel_left) also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) + cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w)" using add_divide_distrib[symmetric, of "prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w" "prices Mkt asset (Suc n) w"] by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) + cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)" by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w - val_process Mkt pf (Suc n) w)" using self_finance_updated[of Mkt asset n w pf v] assms by auto finally show ?thesis . qed lemma self_finance_updated_suc_0: assumes "portfolio pf" and "\n w. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc 0) w = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * (val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc 0) w = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * remaining_qty Mkt v pf asset (Suc 0) w" using assms by (simp add: self_finance_updated) also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((remaining_qty Mkt v pf asset 0 w) + (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) / (prices Mkt asset 0 w) + (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" using assms by (metis nonzero_mult_div_cancel_left) also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w)" using add_divide_distrib[symmetric, of "prices Mkt asset 0 w * remaining_qty Mkt v pf asset 0 w" "prices Mkt asset 0 w"] by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + v - val_process Mkt pf 0 w)" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" using self_finance_init[of Mkt asset pf v w] assms by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * (val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" by simp finally show ?thesis . qed lemma self_finance_updated_ind: assumes "portfolio pf" and "\n w. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (prices Mkt asset (Suc n) w / (prices Mkt asset n w)) * (val_process Mkt (self_finance Mkt v pf asset) n w - val_process Mkt pf n w)" proof (cases "n = 0") case True thus ?thesis using assms self_finance_updated_suc_0 by simp next case False hence "\m. n = Suc m" by (simp add: not0_implies_Suc) from this obtain m where "n = Suc m" by auto hence "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc m)) w" by simp also have "... = cls_val_process Mkt pf (Suc (Suc m)) w + (prices Mkt asset (Suc (Suc m)) w / (prices Mkt asset (Suc m) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc m) w - val_process Mkt pf (Suc m) w)" using assms self_finance_updated_suc_suc[of pf] by simp also have "... = cls_val_process Mkt pf (Suc (Suc m)) w + (prices Mkt asset (Suc (Suc m)) w / (prices Mkt asset (Suc m) w)) * (val_process Mkt (self_finance Mkt v pf asset) (Suc m) w - val_process Mkt pf (Suc m) w)" using assms self_finance_charact unfolding self_financing_def by (simp add: self_finance_succ self_finance_updated) also have "... = cls_val_process Mkt pf (Suc n) w + (prices Mkt asset (Suc n) w / (prices Mkt asset n w)) * (val_process Mkt (self_finance Mkt v pf asset) n w - val_process Mkt pf n w)" using \n = Suc m\ by simp finally show ?thesis . qed lemma (in rfr_disc_equity_market) self_finance_risk_free_update_ind: assumes "portfolio pf" shows "cls_val_process Mkt (self_finance Mkt v pf risk_free_asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (1 + r) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf risk_free_asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (prices Mkt risk_free_asset (Suc n) w / (prices Mkt risk_free_asset n w)) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" proof (rule self_finance_updated_ind, (simp add: assms), intro allI) fix n w show "prices Mkt risk_free_asset n w \ 0" using positive by (metis less_irrefl) qed also have "... = cls_val_process Mkt pf (Suc n) w + (1+r) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" using rf_price positive by (metis acceptable_rate disc_rfr_proc_Suc_div) finally show ?thesis . qed lemma (in CRR_market) delta_pf_portfolio: shows "portfolio (delta_pf N der matur)" unfolding delta_pf_def by (simp add: single_comp_portfolio) lemma (in CRR_market) delta_pf_updated: shows "cls_val_process Mkt (delta_pf N der matur) (Suc n) w = geom_proc (Suc n) w * delta_price N der matur n w" unfolding delta_pf_def using stk_price qty_single_updated[of Mkt] by simp lemma (in CRR_market) delta_pf_val_process: shows "val_process Mkt (delta_pf N der matur) n w = geom_proc n w * delta_price N der matur n w" unfolding delta_pf_def using stk_price qty_single_val_process[of Mkt] by simp lemma (in CRR_market) delta_hedging_cls_val_process: shows "cls_val_process Mkt (delta_hedging N der matur) (Suc n) w = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt (delta_hedging N der matur) n w - geom_proc n w * delta_price N der matur n w)" proof - define X where "X = delta_hedging N der matur" define init where "init = integral\<^sup>L N (discounted_value r (\m. der) matur)" have "cls_val_process Mkt X (Suc n) w = cls_val_process Mkt (delta_pf N der matur) (Suc n) w + (1 + r) * (val_process Mkt X n w - val_process Mkt (delta_pf N der matur) n w)" unfolding X_def delta_hedging_def self_fin_delta_pf_def init_def proof (rule self_finance_risk_free_update_ind) show "portfolio (delta_pf N der matur)" unfolding portfolio_def using delta_pf_support by (meson finite.simps infinite_super) qed also have "... = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt X n w - val_process Mkt (delta_pf N der matur) n w)" using delta_pf_updated by simp also have "... = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt X n w - geom_proc n w * delta_price N der matur n w)" using delta_pf_val_process by simp finally show ?thesis unfolding X_def . qed lemma (in CRR_market_viable) delta_hedging_eq_derivative_price: fixes der::"bool stream \ real" and matur::nat assumes "N = bernoulli_stream ((1 + r - d) / (u - d))" and "der\ borel_measurable (G matur)" shows "\n w. n\ matur \ val_process Mkt (delta_hedging N der matur) n w = (rn_price N der matur) n w" unfolding delta_hedging_def proof - define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this define init where "init = (prob_space.expectation N (discounted_value r (\m. der) matur))" define X where "X = val_process Mkt (delta_hedging N der matur)" define V where "V = rn_price N der matur" define \ where "\ = delta_price N der matur" { fix n fix w have "n \ matur \ X n w = V n w" proof (induct n) case 0 have v0: "V 0 \ borel_measurable (G 0)" using assms rn_price_borel_adapt "0.prems" qprops unfolding V_def q_def by auto have "X 0 w= init" using self_finance_init[of Mkt risk_free_asset "delta_pf N der matur" "integral\<^sup>L N (discounted_value r (\m. der) matur)"] delta_pf_support unfolding X_def init_def delta_hedging_def self_fin_delta_pf_def init_def by (metis finite_insert infinite_imp_nonempty infinite_super less_irrefl portfolio_def positive) also have "... = V 0 w" proof - have "\x\space N. real_cond_exp N (G 0) (discounted_value r (\m. der) matur) x = integral\<^sup>L N (discounted_value r (\m. der) matur)" proof (rule prob_space.trivial_subalg_cond_expect_eq) show "prob_space N" using assms qprops unfolding q_def by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "init_triv_filt M (stoch_proc_filt M geom_proc borel)" proof (rule infinite_cts_filtration.stoch_proc_filt_triv_init) show "borel_adapt_stoch_proc nat_filtration geom_proc" using geom_rand_walk_borel_adapted by simp show "infinite_cts_filtration p M nat_filtration" using bernoulli_nat_filtration[of M p] bernoulli psgt pslt by simp qed hence "init_triv_filt N (stoch_proc_filt M geom_proc borel)" using assms qprops filt_equiv_triv_init[of nat_filtration N] stock_filtration bernoulli_stream_equiv[of N] psgt pslt unfolding q_def by simp thus "subalgebra N (G 0)" and "sets (G 0) = {{}, space N}" using stock_filtration unfolding init_triv_filt_def filtration_def bot_nat_def by auto show "integrable N (discounted_value r (\m. der) matur)" proof (rule bernoulli_discounted_integrable) show "der \ borel_measurable (nat_filtration matur)" using assms geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto qed (simp add: acceptable_rate) qed hence "integral\<^sup>L N (discounted_value r (\m. der) matur) = real_cond_exp N (G 0) (discounted_value r (\m. der) matur) w" using bernoulli_stream_space[of N q] by (simp add: assms(1) q_def) also have "... = real_cond_exp N (stoch_proc_filt M geom_proc borel 0) (discounted_value r (\m. der) matur) w" using stock_filtration by simp also have "... = real_cond_exp N (stoch_proc_filt N geom_proc borel 0) (discounted_value r (\m. der) matur) w" using stoch_proc_filt_filt_equiv[of nat_filtration M N geom_proc] bernoulli_stream_equiv[of N] q_def qprops assms pslt psgt by auto also have "... = expl_cond_expect N (proj_stoch_proc geom_proc 0) (discounted_value r (\m. der) matur) w" proof (rule bernoulli_cond_exp) show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto show "integrable N (discounted_value r (\m. der) matur)" proof (rule bernoulli_discounted_integrable) show "der \ borel_measurable (nat_filtration matur)" using assms geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto qed (simp add: acceptable_rate) qed finally show "init = V 0 w" unfolding init_def V_def rn_price_def by simp qed finally show "X 0 w = V 0 w" . next case (Suc n) hence "n < matur" by simp show ?case proof - have "X n w = V n w" using Suc by (simp add: Suc.hyps Suc.prems Suc_leD) have "0< 1+r" using acceptable_rate by simp let ?m = "matur - Suc n" have "matur -n = Suc ?m" by (simp add: Suc.prems Suc_diff_Suc Suc_le_lessD) have "(1+r) * V n w = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" using rn_price_eq_ind qprops assms Suc q_def V_def by simp show "X (Suc n) w = V (Suc n) w" proof (cases "snth w n") case True hence pseq: "pseudo_proj_True (Suc n) w = pseudo_proj_True (Suc n) (spick w n True)" by (metis (mono_tags, lifting) pseudo_proj_True_stake_image spickI stake_Suc) have "X (Suc n) w = cls_val_process Mkt (delta_hedging N der matur) (Suc n) w" unfolding X_def delta_hedging_def self_fin_delta_pf_def using delta_pf_portfolio unfolding self_financing_def by (metis less_irrefl positive self_finance_charact self_financingE) also have "... = geom_proc (Suc n) w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using delta_hedging_cls_val_process unfolding X_def \_def by simp also have "... = u * geom_proc n w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using True geometric_process by simp also have "... = u * geom_proc n w * \ n w + (1 + r) * X n w - (1+r) * geom_proc n w * \ n w" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * \ n w * u - geom_proc n w * \ n w * (1 + r)" by (simp add: mult.commute mult.left_commute) also have "... = (1+r)* X n w + geom_proc n w * \ n w * (u - (1 + r))" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * (V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) * (u - (1 + r))" using Suc V_def by (simp add: \_def delta_price_def geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + geom_proc n w * ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))) / (geom_proc n w * (u - d)) * (u - (1 + r))" proof - have "geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False) = - u * geom_proc n w - d * geom_proc n w" - using geometric_process - by (metis geom_rand_walk_diff_induct linordered_field_class.sign_simps(39) mult.commute) - also have "... = geom_proc n w * (u - d)" - by (simp add: linordered_field_class.sign_simps(39) mult.commute) - finally have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = - geom_proc n w * (u - d)" . - thus ?thesis by simp + geom_proc n w * (u - d)" + by (simp add: geom_rand_walk_diff_induct) + then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w)))* (u - (1 + r))/ (u-d)" proof - have "geom_proc n w \ 0" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl) then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))* (1 - q))" proof - have "1 - q = 1 - (1 + r - d)/(u -d)" unfolding q_def by simp also have "... = (u - d)/(u - d) - (1 + r - d)/(u -d)" using down_lt_up by simp also have "... = (u - d - (1 + r - d))/(u - d)" using diff_divide_distrib[of "u - d" "1 + r -d"] by simp also have "... = (u - (1+r))/(u-d)" by simp finally have "1 - q = (u - (1+r))/(u-d)" . thus ?thesis by simp qed also have "... = (1+r) * X n w + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" by (simp add: mult.commute right_diff_distrib) also have "... = (1+r) * V n w + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" using \X n w = V n w\ by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_False n w) + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" using assms Suc rn_price_eq_ind[of N q n matur der w] \n < matur\ qprops unfolding V_def q_def by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_True n w)" by simp also have "... = V (Suc n) (pseudo_proj_True n w)" using distrib_right[of q "1 - q" "V (Suc n) (pseudo_proj_True n w)"] by simp also have "... = V (Suc n) w" proof - have "V (Suc n) \ borel_measurable (G (Suc n))" unfolding V_def q_def proof (rule rn_price_borel_adapt) show "der \ borel_measurable (G matur)" using assms by simp show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" and "q < 1" using qprops by auto show "Suc n \ matur" using Suc by simp qed hence "V (Suc n) (pseudo_proj_True n w) = V (Suc n) (pseudo_proj_True (Suc n) (pseudo_proj_True n w))" using geom_proc_filt_info[of "V (Suc n)" "Suc n"] by simp also have "... = V (Suc n) (pseudo_proj_True (Suc n) w)" using True by (simp add: pseq spick_eq_pseudo_proj_True) also have "... = V (Suc n) w" using \V (Suc n) \ borel_measurable (G (Suc n))\ geom_proc_filt_info[of "V (Suc n)" "Suc n"] by simp finally show ?thesis . qed finally show "X (Suc n) w = V (Suc n) w" . next case False hence pseq: "pseudo_proj_True (Suc n) w = pseudo_proj_True (Suc n) (spick w n False)" using filtration by (metis (full_types) pseudo_proj_True_def spickI stake_Suc) have "X (Suc n) w = cls_val_process Mkt (delta_hedging N der matur) (Suc n) w" unfolding X_def delta_hedging_def self_fin_delta_pf_def using delta_pf_portfolio unfolding self_financing_def by (metis less_irrefl positive self_finance_charact self_financingE) also have "... = geom_proc (Suc n) w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using delta_hedging_cls_val_process unfolding X_def \_def by simp also have "... = d * geom_proc n w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using False geometric_process by simp also have "... = d * geom_proc n w * \ n w + (1 + r) * X n w - (1+r) * geom_proc n w * \ n w" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * \ n w * d - geom_proc n w * \ n w * (1 + r)" by (simp add: mult.commute mult.left_commute) also have "... = (1+r)* X n w + geom_proc n w * \ n w * (d - (1 + r))" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * (V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) * (d - (1 + r))" using Suc V_def by (simp add: \_def delta_price_def geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + geom_proc n w * ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))) / (geom_proc n w * (u - d)) * (d - (1 + r))" - proof - - have "geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False) = - u * geom_proc n w - d * geom_proc n w" - using geometric_process - by (metis geom_rand_walk_diff_induct linordered_field_class.sign_simps(39) mult.commute) - also have "... = geom_proc n w * (u - d)" - by (simp add: linordered_field_class.sign_simps(39) mult.commute) - finally have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = - geom_proc n w * (u - d)" . - thus ?thesis by simp - qed + by (simp add: geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w)))* (d - (1 + r))/ (u-d)" proof - have "geom_proc n w \ 0" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl) then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))* (-q))" proof - have "0-q = 0-(1 + r - d)/(u -d)" unfolding q_def by simp also have "... = (d - (1 + r))/(u -d)" by (simp add: minus_divide_left) finally have "0 - q = (d - (1+r))/(u-d)" . thus ?thesis by simp qed also have "... = (1+r) * X n w + (- V (Suc n) (pseudo_proj_True n w) * q + V (Suc n) (pseudo_proj_False n w)* q)" by (metis (no_types, hide_lams) add.inverse_inverse distrib_right minus_mult_commute minus_real_def mult_minus_left) also have "... = (1+r) * X n w - q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" by simp also have "... = (1+r) * V n w -q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" using \X n w = V n w\ by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_False n w) - q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" using assms Suc rn_price_eq_ind[of N q n matur der w] \n < matur\ qprops unfolding V_def q_def by simp also have "... = (1-q) * V (Suc n) (pseudo_proj_False n w) + q * V (Suc n) (pseudo_proj_False n w)" by simp also have "... = V (Suc n) (pseudo_proj_False n w)" using distrib_right[of q "1 - q" "V (Suc n) (pseudo_proj_False n w)"] by simp also have "... = V (Suc n) w" proof - have "V (Suc n) \ borel_measurable (G (Suc n))" unfolding V_def q_def proof (rule rn_price_borel_adapt) show "der \ borel_measurable (G matur)" using assms by simp show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" and "q < 1" using qprops by auto show "Suc n \ matur" using Suc by simp qed hence "V (Suc n) (pseudo_proj_False n w) = V (Suc n) (pseudo_proj_False (Suc n) (pseudo_proj_False n w))" using geom_proc_filt_info'[of "V (Suc n)" "Suc n"] by simp also have "... = V (Suc n) (pseudo_proj_False (Suc n) w)" using False spick_eq_pseudo_proj_False by (metis pseq pseudo_proj_True_imp_False) also have "... = V (Suc n) w" using \V (Suc n) \ borel_measurable (G (Suc n))\ geom_proc_filt_info'[of "V (Suc n)" "Suc n"] by simp finally show ?thesis . qed finally show "X (Suc n) w = V (Suc n) w" . qed qed qed } thus "\n w. n \ matur \ val_process Mkt (self_fin_delta_pf N der matur (integral\<^sup>L N (discounted_value r (\m. der) matur))) n w = rn_price N der matur n w" by (simp add: X_def init_def V_def delta_hedging_def) qed lemma (in CRR_market_viable) delta_hedging_same_cash_flow: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream ((1 + r - d) / (u - d))" shows "cls_val_process Mkt (delta_hedging N der matur) matur w = der w" proof - define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this have "cls_val_process Mkt (delta_hedging N der matur) matur w = val_process Mkt (delta_hedging N der matur) matur w" using self_financingE self_finance_charact unfolding delta_hedging_def self_fin_delta_pf_def by (metis delta_pf_portfolio mult_1s(1) mult_cancel_right not_real_square_gt_zero positive) also have "... = rn_price N der matur matur w" using delta_hedging_eq_derivative_price assms by simp also have "... = rn_rev_price N der matur 0 w" using rn_price_eq qprops assms unfolding rn_price_ind_def q_def by simp also have "... = der w" by simp finally show ?thesis . qed lemma (in CRR_market) delta_hedging_trading_strat: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" shows "trading_strategy (delta_hedging N der matur)" unfolding delta_hedging_def by (simp add: assms self_fin_delta_pf_trad_strat) lemma (in CRR_market) delta_hedging_self_financing: shows "self_financing Mkt (delta_hedging N der matur)" unfolding delta_hedging_def self_fin_delta_pf_def proof (rule self_finance_charact) show "\n w. prices Mkt risk_free_asset (Suc n) w \ 0" using positive by (metis less_numeral_extra(3)) show "portfolio (delta_pf N der matur)" using delta_pf_portfolio . qed lemma (in CRR_market_viable) delta_hedging_replicating: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream ((1 + r - d) / (u - d))" shows "replicating_portfolio (delta_hedging N der matur) der matur" unfolding replicating_portfolio_def proof (intro conjI) define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this let ?X = "(delta_hedging N der matur)" show "trading_strategy ?X" using delta_hedging_trading_strat qprops assms unfolding q_def by simp show "self_financing Mkt ?X" using delta_hedging_self_financing . show "stock_portfolio Mkt (delta_hedging N der matur)" unfolding delta_hedging_def self_fin_delta_pf_def stock_portfolio_def portfolio_def using stocks delta_pf_support by (smt Un_insert_right delta_pf_portfolio insert_commute portfolio_def self_finance_def self_finance_portfolio single_comp_support subset_insertI2 subset_singleton_iff sum_support_set sup_bot.right_neutral) show "AEeq M (cls_val_process Mkt (delta_hedging N der matur) matur) der" using delta_hedging_same_cash_flow assms by simp qed definition (in disc_equity_market) complete_market where "complete_market \ (\matur. \ der\ borel_measurable (F matur). (\p. replicating_portfolio p der matur))" lemma (in CRR_market_viable) CRR_market_complete: shows "complete_market" unfolding complete_market_def proof (intro allI impI) fix matur::nat show "\ der \ borel_measurable (G matur). (\p. replicating_portfolio p der matur)" proof fix der::"bool stream\real" assume "der \ borel_measurable (G matur)" define N where "N = bernoulli_stream ((1 + r - d) / (u - d))" hence "replicating_portfolio (delta_hedging N der matur) der matur" using delta_hedging_replicating \der \ borel_measurable (G matur)\ by simp thus "\pf. replicating_portfolio pf der matur" by auto qed qed lemma subalgebras_filtration: assumes "filtration M F" and "\t. subalgebra (F t) (G t)" and "\ s t. s \ t \ subalgebra (G t) (G s)" shows "filtration M G" unfolding filtration_def proof (intro conjI allI impI) { fix t have "subalgebra (F t) (G t)" using assms by simp moreover have "subalgebra M (F t)" using assms unfolding filtration_def by simp ultimately show "subalgebra M (G t)" by (metis subalgebra_def subsetCE subsetI) } { fix s t::'b assume "s \ t" thus "subalgebra (G t) (G s)" using assms by simp } qed lemma subfilt_filt_equiv: assumes "filt_equiv F M N" and "\ t. subalgebra (F t) (G t)" and "\ s t. s \ t \ subalgebra (G t) (G s)" shows "filt_equiv G M N" unfolding filt_equiv_def proof (intro conjI) show "sets M = sets N" using assms unfolding filt_equiv_def by simp show "filtration M G" using assms subalgebras_filtration[of M F G] unfolding filt_equiv_def by simp show "\t A. A \ sets (G t) \ (emeasure M A = 0) = (emeasure N A = 0)" proof (intro allI ballI impI) fix t fix A assume "A\ sets (G t)" hence "A \ sets (F t)" using assms unfolding subalgebra_def by auto thus "(emeasure M A = 0) = (emeasure N A = 0)" using assms unfolding filt_equiv_def by simp qed qed lemma (in CRR_market_viable) CRR_market_fair_price: assumes "pyf\ borel_measurable (G matur)" shows "fair_price Mkt (\ w\ range (pseudo_proj_True matur). (prod (prob_component ((1 + r - d) / (u - d)) w) {0..m. pyf) matur) w)) pyf matur" proof - define dpf where "dpf = (discounted_value r (\m. pyf) matur)" define q where "q = (1 + r - d) / (u - d)" have "\pf. replicating_portfolio pf pyf matur" using CRR_market_complete assms unfolding complete_market_def by simp from this obtain pf where "replicating_portfolio pf pyf matur" by auto note pfprop = this define N where "N = bernoulli_stream ((1 + r - d) / (u - d))" have "fair_price Mkt (integral\<^sup>L N dpf) pyf matur" unfolding dpf_def proof (rule replicating_expectation_finite) show "risk_neutral_prob N" using assms risk_neutral_iff using CRR_viable gt_param lt_param N_def by blast have "filt_equiv nat_filtration M N" using bernoulli_stream_equiv[of N "(1+r-d)/(u-d)"] assms gt_param lt_param CRR_viable psgt pslt N_def by simp thus "filt_equiv G M N" using subfilt_filt_equiv using Filtration.filtration_def filtration geom_rand_walk_borel_adapted stoch_proc_subalg_nat_filt stock_filtration by blast show "pyf \ borel_measurable (G matur)" using assms by simp show "viable_market Mkt" using CRR_viable by simp have "infinite_cts_filtration p M nat_filtration" using bernoulli_nat_filtration[of M p] bernoulli psgt pslt by simp thus "sets (G 0) = {{}, space M}" using stock_filtration infinite_cts_filtration.stoch_proc_filt_triv_init[of p M nat_filtration geom_proc] geom_rand_walk_borel_adapted bot_nat_def unfolding init_triv_filt_def by simp show "replicating_portfolio pf pyf matur" using pfprop . show "\n. \asset\support_set pf. finite (prices Mkt asset n ` space M)" proof (intro allI ballI) fix n fix asset assume "asset \ support_set pf" hence "prices Mkt asset n \ borel_measurable (G n)" using readable pfprop unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def by auto hence "prices Mkt asset n \ borel_measurable (nat_filtration n)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration n" "G n" "prices Mkt asset n" borel] unfolding adapt_stoch_proc_def by auto thus "finite (prices Mkt asset n ` space M)" using nat_filtration_vimage_finite[of "prices Mkt asset n"] by simp qed show "\n. \asset\support_set pf. finite (pf asset n ` space M)" proof (intro allI ballI) fix n fix asset assume "asset \ support_set pf" hence "pf asset n \ borel_measurable (G n)" using pfprop predict_imp_adapt[of "pf asset"] unfolding replicating_portfolio_def trading_strategy_def adapt_stoch_proc_def by auto hence "pf asset n \ borel_measurable (nat_filtration n)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration n" "G n" "pf asset n" borel] unfolding adapt_stoch_proc_def by auto thus "finite (pf asset n ` space M)" using nat_filtration_vimage_finite[of "pf asset n"] by simp qed qed moreover have "integral\<^sup>L N dpf = (\ w\ range (pseudo_proj_True matur). (prod (prob_component q w) {0.. borel_measurable (G matur)" using assms discounted_measurable[of pyf "G matur"] unfolding dpf_def by simp thus "dpf \ borel_measurable (nat_filtration matur)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration matur" "G matur" dpf] unfolding adapt_stoch_proc_def by auto qed ultimately show ?thesis unfolding dpf_def q_def by simp qed end \ No newline at end of file diff --git a/thys/Ergodic_Theory/Fekete.thy b/thys/Ergodic_Theory/Fekete.thy --- a/thys/Ergodic_Theory/Fekete.thy +++ b/thys/Ergodic_Theory/Fekete.thy @@ -1,488 +1,490 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Subadditive and submultiplicative sequences\ theory Fekete imports "HOL-Analysis.Analysis" begin text \A real sequence is subadditive if $u_{n+m} \leq u_n+u_m$. This implies the convergence of $u_n/n$ to $Inf\{u_n/n\} \in [-\infty, +\infty)$, a useful result known as Fekete lemma. We prove it below. Taking logarithms, the same result applies to submultiplicative sequences. We illustrate it with the definition of the spectral radius as the limit of $\|x^n\|^{1/n}$, the convergence following from Fekete lemma.\ subsection \Subadditive sequences\ text \We define subadditive sequences, either from the start or eventually.\ definition subadditive::"(nat\real) \ bool" where "subadditive u = (\m n. u (m+n) \ u m + u n)" lemma subadditiveI: assumes "\m n. u (m+n) \ u m + u n" shows "subadditive u" unfolding subadditive_def using assms by auto lemma subadditiveD: assumes "subadditive u" shows "u (m+n) \ u m + u n" using assms unfolding subadditive_def by auto lemma subadditive_un_le_nu1: assumes "subadditive u" "n > 0" shows "u n \ n * u 1" proof - have *: "n = 0 \ (u n \ n * u 1)" for n proof (induction n) case 0 then show ?case by auto next case (Suc n) consider "n = 0" | "n > 0" by auto then show ?case proof (cases) case 1 then show ?thesis by auto next case 2 then have "u (Suc n) \ u n + u 1" using subadditiveD[OF assms(1), of n 1] by auto then show ?thesis using Suc.IH 2 by (auto simp add: algebra_simps) qed qed show ?thesis using *[of n] \n > 0\ by auto qed definition eventually_subadditive::"(nat\real) \ nat \ bool" where "eventually_subadditive u N0 = (\m>N0. \n>N0. u (m+n) \ u m + u n)" lemma eventually_subadditiveI: assumes "\m n. m > N0 \ n > N0 \ u (m+n) \ u m + u n" shows "eventually_subadditive u N0" unfolding eventually_subadditive_def using assms by auto lemma subadditive_imp_eventually_subadditive: assumes "subadditive u" shows "eventually_subadditive u 0" using assms unfolding subadditive_def eventually_subadditive_def by auto text \The main inequality that will lead to convergence is given in the next lemma: given $n$, then eventually $u_m/m$ is bounded by $u_n/n$, up to an arbitrarily small error. This is proved by doing the euclidean division of $m$ by $n$ and using the subadditivity. (the remainder in the euclidean division will give the error term.)\ lemma eventually_subadditive_ineq: assumes "eventually_subadditive u N0" "e>0" "n>N0" shows "\N>N0. \m\N. u m/m < u n/n + e" proof - have ineq_rec: "u(a*n+r) \ a * u n + u r" if "n>N0" "r>N0" for a n r proof (induct a) case (Suc a) have "a*n+r>N0" using \r>N0\ by simp have "u((Suc a)*n+r) = u(a*n+r+n)" by (simp add: algebra_simps) also have "... \ u(a*n+r)+u n" using assms \n>N0\ \a*n+r>N0\ eventually_subadditive_def by blast also have "... \ a*u n + u r + u n" by (simp add: Suc.hyps) also have "... = (Suc a) * u n + u r" by (simp add: algebra_simps) finally show ?case by simp qed (simp) have "n>0" "real n > 0" using \n>N0\ by auto define C where "C = Max {abs(u i) |i. i\2*n}" have ineq_C: "abs(u i) \ C" if "i \ 2 * n" for i unfolding C_def by (intro Max_ge, auto simp add: that) have ineq_all_m: "u m/m \ u n/n + 3*C/m" if "m\n" for m proof - have "real m>0" using \m\n\ \0 < real n\ by linarith obtain a0 r0 where "r00 < n\ mod_div_decomp mod_less_divisor by blast define a where "a = a0-1" define r where "r = r0+n" have "r<2*n" "r\n" unfolding r_def by (auto simp add: \r0) have "a0>0" using \m = a0*n + r0\ \n \ m\ \r0 < n\ not_le by fastforce then have "m = a * n + r" using a_def r_def \m = a0*n+r0\ mult_eq_if by auto then have real_eq: "-r = real n * a - m" by simp have "r>N0" using \r\n\ \n>N0\ by simp then have "u m \ a * u n + u r" using ineq_rec \m = a*n+r\ \n>N0\ by simp then have "n * u m \ n * (a * u n + u r)" using \real n>0\ by simp then have "n * u m - m * u n \ -r * u n + n * u r" unfolding real_eq by (simp add: algebra_simps) also have "... \ r * abs(u n) + n * abs(u r)" apply (intro add_mono mult_left_mono) using real_0_le_add_iff by fastforce+ also have "... \ (2 * n) * C + n * C" apply (intro add_mono mult_mono ineq_C) using less_imp_le[OF \r < 2 * n\] by auto finally have "n * u m - m * u n \ 3*C*n" by auto then show "u m/m \ u n/n + 3*C/m" using \0 < real n\ \0 < real m\ by (simp add: divide_simps mult.commute) qed obtain M::nat where M: "M \ 3 * C / e" using real_nat_ceiling_ge by auto define N where "N = M + n + N0 + 1" have "N > 3 * C / e" "N \ n" "N > N0" unfolding N_def using M by auto have "u m/m < u n/n + e" if "m \ N" for m proof - have "3 * C / m < e" using that \N > 3 * C / e\ \e > 0\ apply (auto simp add: algebra_simps divide_simps) by (meson le_less_trans linorder_not_le mult_less_cancel_left_pos of_nat_less_iff) then show ?thesis using ineq_all_m[of m] \n \ N\ \N \ m\ by auto qed then show ?thesis using \N0 < N\ by blast qed text \From the inequality above, we deduce the convergence of $u_n/n$ to its infimum. As this infimum might be $-\infty$, we formulate this convergence in the extended reals. Then, we specialize it to the real situation, separating the cases where $u_n/n$ is bounded below or not.\ lemma subadditive_converges_ereal': assumes "eventually_subadditive u N0" shows "(\m. ereal(u m/m)) \ Inf {ereal(u n/n) | n. n>N0}" proof - define v where "v = (\m. ereal(u m/m))" define V where "V = {v n | n. n>N0}" define l where "l = Inf V" have "\t. t\V \ t\l" by (simp add: Inf_lower l_def) then have "v n \ l" if "n > N0" for n using V_def that by blast then have lower: "eventually (\n. a < v n) sequentially" if "a < l" for a by (meson that dual_order.strict_trans1 eventually_at_top_dense) have upper: "eventually (\n. a > v n) sequentially" if "a > l" for a proof - obtain t where "t\V" "ta>l\ Inf_greatest l_def not_le) then obtain e::real where "e>0" "t+e < a" by (meson ereal_le_epsilon2 leD le_less_linear) obtain n where "n>N0" "t = u n/n" using V_def v_def \t \ V\ by blast then have "u n/n + e < a" using \t+e < a\ by simp obtain N where "\m\N. u m/m < u n/n + e" using eventually_subadditive_ineq[OF assms] \0 < e\ \N0 < n\ by blast then have "u m/m < a" if "m \ N" for m using that \u n/n + e < a\ less_ereal.simps(1) less_trans by blast then have "v m< a" if "m \ N" for m using v_def that by blast then show ?thesis using eventually_at_top_linorder by auto qed show ?thesis using lower upper unfolding V_def l_def v_def by (simp add: order_tendsto_iff) qed lemma subadditive_converges_ereal: assumes "subadditive u" shows "(\m. ereal(u m/m)) \ Inf {ereal(u n/n) | n. n>0}" by (rule subadditive_converges_ereal'[OF subadditive_imp_eventually_subadditive[OF assms]]) lemma subadditive_converges_bounded': assumes "eventually_subadditive u N0" "bdd_below {u n/n | n. n>N0}" shows "(\n. u n/n) \ Inf {u n/n | n. n>N0}" proof- have *: "(\n. ereal(u n /n)) \ Inf {ereal(u n/n)|n. n > N0}" by (simp add: assms(1) subadditive_converges_ereal') define V where "V = {u n/n | n. n>N0}" have a: "bdd_below V" "V\{}" by (auto simp add: V_def assms(2)) have "Inf {ereal(t)| t. t\V} = ereal(Inf V)" by (subst ereal_Inf'[OF a], simp add: Setcompr_eq_image) moreover have "{ereal(t)| t. t\V} = {ereal(u n/n)|n. n > N0}" using V_def by blast ultimately have "Inf {ereal(u n/n)|n. n > N0} = ereal(Inf {u n/n |n. n > N0})" using V_def by auto then have "(\n. ereal(u n /n)) \ ereal(Inf {u n/n | n. n>N0})" using * by auto then show ?thesis by simp qed lemma subadditive_converges_bounded: assumes "subadditive u" "bdd_below {u n/n | n. n>0}" shows "(\n. u n/n) \ Inf {u n/n | n. n>0}" by (rule subadditive_converges_bounded'[OF subadditive_imp_eventually_subadditive[OF assms(1)] assms(2)]) text \We reformulate the previous lemma in a more directly usable form, avoiding the infimum.\ lemma subadditive_converges_bounded'': assumes "subadditive u" "\n. n > 0 \ u n \ n * (a::real)" shows "\l. (\n. u n / n) \ l \ (\n>0. u n \ n * l)" proof - have B: "bdd_below {u n/n | n. n>0}" apply (rule bdd_belowI[of _ a]) using assms(2) apply (auto simp add: divide_simps) - by (metis linordered_field_class.sign_simps(5) mult_left_le_imp_le of_nat_0_less_iff) + apply (metis mult.commute mult_left_le_imp_le of_nat_0_less_iff) + done define l where "l = Inf {u n/n | n. n>0}" have *: "u n / n \ l" if "n > 0" for n unfolding l_def using that by (auto intro!: cInf_lower[OF _ B]) show ?thesis apply (rule exI[of _ l], auto) using subadditive_converges_bounded[OF assms(1) B] apply (simp add: l_def) using * by (simp add: divide_simps algebra_simps) qed lemma subadditive_converges_unbounded': assumes "eventually_subadditive u N0" "\ (bdd_below {u n/n | n. n>N0})" shows "(\n. ereal(u n/n)) \ -\" proof - have *: "(\n. ereal(u n /n)) \ Inf {ereal(u n/n)|n. n > N0}" by (simp add: assms(1) subadditive_converges_ereal') define V where "V = {u n/n | n. n>N0}" then have "\ bdd_below V" using assms by simp have "Inf {ereal(t) | t. t\V} = -\" by (rule ereal_bot, metis (mono_tags, lifting) \\ bdd_below V\ bdd_below_def leI Inf_lower2 ereal_less_eq(3) le_less mem_Collect_eq) moreover have "{ereal(t)| t. t\V} = {ereal(u n/n)|n. n > N0}" using V_def by blast ultimately have "Inf {ereal(u n/n)|n. n > N0} = -\" by auto then show ?thesis using * by simp qed lemma subadditive_converges_unbounded: assumes "subadditive u" "\ (bdd_below {u n/n | n. n>0})" shows "(\n. ereal(u n/n)) \ -\" by (rule subadditive_converges_unbounded'[OF subadditive_imp_eventually_subadditive[OF assms(1)] assms(2)]) subsection \Superadditive sequences\ text \While most applications involve subadditive sequences, one sometimes encounters superadditive sequences. We reformulate quickly some of the above results in this setting.\ definition superadditive::"(nat\real) \ bool" where "superadditive u = (\m n. u (m+n) \ u m + u n)" lemma subadditive_of_superadditive: assumes "superadditive u" shows "subadditive (\n. -u n)" using assms unfolding superadditive_def subadditive_def by (auto simp add: algebra_simps) lemma superadditive_un_ge_nu1: assumes "superadditive u" "n > 0" shows "u n \ n * u 1" using subadditive_un_le_nu1[OF subadditive_of_superadditive[OF assms(1)] assms(2)] by auto lemma superadditive_converges_bounded'': assumes "superadditive u" "\n. n > 0 \ u n \ n * (a::real)" shows "\l. (\n. u n / n) \ l \ (\n>0. u n \ n * l)" proof - have "\l. (\n. -u n / n) \ l \ (\n>0. -u n \ n * l)" apply (rule subadditive_converges_bounded''[OF subadditive_of_superadditive[OF assms(1)], of "-a"]) using assms(2) by auto then obtain l where l: "(\n. -u n / n) \ l" "(\n>0. -u n \ n * l)" by blast have "(\n. -((-u n)/n)) \ -l" by (intro tendsto_intros l) moreover have "\n>0. u n \ n * (-l)" using l(2) by (auto simp add: algebra_simps) (metis minus_equation_iff neg_le_iff_le) ultimately show ?thesis by auto qed subsection \Almost additive sequences\ text \One often encounters sequences which are both subadditive and superadditive, but only up to an additive constant. Adding or subtracting this constant, one can make the sequence genuinely subadditive or superadditive, and thus deduce results about its convergence, as follows. Such sequences appear notably when dealing with quasimorphisms.\ lemma almost_additive_converges: fixes u::"nat \ real" assumes "\m n. abs(u(m+n) - u m - u n) \ C" shows "convergent (\n. u n/n)" "abs(u k - k * lim (\n. u n / n)) \ C" proof - have "(abs (u 0)) \ C" using assms[of 0 0] by auto then have "C \ 0" by auto define v where "v = (\n. u n + C)" have "subadditive v" unfolding subadditive_def v_def using assms by (auto simp add: algebra_simps abs_diff_le_iff) then have vle: "v n \ n * v 1" if "n > 0" for n using subadditive_un_le_nu1 that by auto define w where "w = (\n. u n - C)" have "superadditive w" unfolding superadditive_def w_def using assms by (auto simp add: algebra_simps abs_diff_le_iff) then have wge: "w n \ n * w 1" if "n > 0" for n using superadditive_un_ge_nu1 that by auto have I: "v n \ w n" for n unfolding v_def w_def using \C \ 0\ by auto then have *: "v n \ n * w 1" if "n > 0" for n using order_trans[OF wge[OF that]] by auto then obtain lv where lv: "(\n. v n/n) \ lv" "\n. n > 0 \ v n \ n * lv" using subadditive_converges_bounded''[OF \subadditive v\ *] by auto have *: "w n \ n * v 1" if "n > 0" for n using order_trans[OF _ vle[OF that]] I by auto then obtain lw where lw: "(\n. w n/n) \ lw" "\n. n > 0 \ w n \ n * lw" using superadditive_converges_bounded''[OF \superadditive w\ *] by auto have *: "v n/n = w n /n + 2*C*(1/n)" for n unfolding v_def w_def by (auto simp add: algebra_simps divide_simps) have "(\n. w n /n + 2*C*(1/n)) \ lw + 2*C*0" by (intro tendsto_add tendsto_mult lim_1_over_n lw, auto) then have "lw = lv" unfolding *[symmetric] using lv(1) LIMSEQ_unique by auto have *: "u n/n = w n /n + C*(1/n)" for n unfolding w_def by (auto simp add: algebra_simps divide_simps) have "(\n. u n /n) \ lw + C*0" unfolding * by (intro tendsto_add tendsto_mult lim_1_over_n lw, auto) then have lu: "convergent (\n. u n/n)" "lim (\n. u n/n) = lw" by (auto simp add: convergentI limI) then show "convergent (\n. u n/n)" by simp show "abs(u k - k * lim (\n. u n / n)) \ C" proof (cases "k>0") case False then show ?thesis using assms[of 0 0] by auto next case True have "u k - k * lim (\n. u n/n) = v k - C - k * lv" unfolding lu(2) \lw = lv\ v_def by auto also have "... \ -C" using lv(2)[OF True] by auto finally have A: "u k - k * lim (\n. u n/n) \ - C" by simp have "u k - k * lim (\n. u n/n) = w k + C - k * lw" unfolding lu(2) w_def by auto also have "... \ C" using lw(2)[OF True] by auto finally show ?thesis using A by auto qed qed subsection \Submultiplicative sequences, application to the spectral radius\ text \In the same way as subadditive sequences, one may define submultiplicative sequences. Essentially, a sequence is submultiplicative if its logarithm is subadditive. A difference is that we allow a submultiplicative sequence to take the value $0$, as this shows up in applications. This implies that we have to distinguish in the proofs the situations where the value $0$ is taken or not. In the latter situation, we can use directly the results from the subadditive case to deduce convergence. In the former situation, convergence to $0$ is obvious as the sequence vanishes eventually.\ lemma submultiplicative_converges: fixes u::"nat\real" assumes "\n. u n \ 0" "\m n. u (m+n) \ u m * u n" shows "(\n. root n (u n))\ Inf {root n (u n) | n. n>0}" proof - define v where "v = (\ n. root n (u n))" define V where "V = {v n | n. n>0}" then have "V \ {}" by blast have "t \ 0" if "t \ V" for t using that V_def v_def assms(1) by auto then have "Inf V \ 0" by (simp add: \V \ {}\ cInf_greatest) have "bdd_below V" by (meson \\t. t \ V \ 0 \ t\ bdd_below_def) show ?thesis proof cases assume "\n. u n = 0" then obtain n where "u n = 0" by auto then have "u m = 0" if "m \ n" for m by (metis that antisym_conv assms(1) assms(2) le_Suc_ex mult_zero_left) then have *: "v m = 0" if "m \ n" for m using v_def that by simp then have "v \ 0" using tendsto_explicit by force have "v (Suc n) \ V" using V_def by blast moreover have "v (Suc n) = 0" using * by auto ultimately have "Inf V \ 0" by (simp add: \bdd_below V\ cInf_lower) then have "Inf V = 0" using \0 \ Inf V\ by auto then show ?thesis using V_def v_def \v \ 0\ by auto next assume "\ (\n. u n = 0)" then have "u n > 0" for n by (metis assms(1) less_eq_real_def) define w where "w n = ln (u n)" for n have express_vn: "v n = exp(w n/n)" if "n>0" for n proof - have "(exp(w n/n))^n = exp(n*(w n/n))" by (metis exp_of_nat_mult) also have "... = exp(w n)" by (simp add: \0 < n\) also have "... = u n" by (simp add: \\n. 0 < u n\ w_def) finally have "exp(w n/n) = root n (u n)" by (metis \0 < n\ exp_ge_zero real_root_power_cancel) then show ?thesis unfolding v_def by simp qed have "eventually_subadditive w 0" proof (rule eventually_subadditiveI) fix m n have "w (m+n) = ln (u (m+n))" by (simp add: w_def) also have "... \ ln(u m * u n)" by (meson \\n. 0 < u n\ assms(2) zero_less_mult_iff ln_le_cancel_iff) also have "... = ln(u m) + ln(u n)" by (simp add: \\n. 0 < u n\ ln_mult) also have "... = w m + w n" by (simp add: w_def) finally show "w (m+n) \ w m + w n". qed define l where "l = Inf V" then have "v n\l" if "n > 0" for n using V_def that by (metis (mono_tags, lifting) \bdd_below V\ cInf_lower mem_Collect_eq) then have lower: "eventually (\n. a < v n) sequentially" if "a < l" for a by (meson that dual_order.strict_trans1 eventually_at_top_dense) have upper: "eventually (\n. a > v n) sequentially" if "a > l" for a proof - obtain t where "t\V" "t < a" using \V \ {}\ cInf_lessD l_def \a>l\ by blast then have "t > 0" using V_def \\n. 0 < u n\ v_def by auto then have "a/t > 1" using \t by simp define e where "e = ln(a/t)/2" have "e > 0" "e < ln(a/t)" unfolding e_def by (simp_all add: \1 < a / t\ ln_gt_zero) then have "exp(e) < a/t" by (metis \1 < a / t\ exp_less_cancel_iff exp_ln less_trans zero_less_one) obtain n where "n>0" "t = v n" using V_def v_def \t \ V\ by blast - then have "v n * exp(e) < a" using \exp(e) < a/t\ by (metis \0 < t\ linordered_field_class.sign_simps(24) pos_less_divide_eq) + with \0 < t\ have "v n * exp(e) < a" using \exp(e) < a/t\ + by (auto simp add: field_simps) obtain N where *: "N>0" "\m. m\N \ w m/m < w n/n + e" using eventually_subadditive_ineq[OF \eventually_subadditive w 0\] \0 < n\ \e>0\ by blast have "v m < a" if "m \ N" for m proof - have "m>0" using that \N>0\ by simp have "w m/m < w n/n + e" by (simp add: \N \ m\ *) then have "exp(w m/m) < exp(w n/n + e)" by simp also have "... = exp(w n/n) * exp(e)" by (simp add: mult_exp_exp) finally have "v m < v n * exp(e)" using express_vn \m>0\ \n>0\ by simp then show "v m < a" using \v n * exp(e) < a\ by simp qed then show ?thesis using eventually_at_top_linorder by auto qed show ?thesis using lower upper unfolding v_def l_def V_def by (simp add: order_tendsto_iff) qed qed text \An important application of submultiplicativity is to prove the existence of the spectral radius of a matrix, as the limit of $\|A^n\|^{1/n}$.\ definition spectral_radius::"'a::real_normed_algebra_1 \ real" where "spectral_radius x = Inf {root n (norm(x^n))| n. n>0}" lemma spectral_radius_aux: fixes x::"'a::real_normed_algebra_1" defines "V \ {root n (norm(x^n))| n. n>0}" shows "\t. t\V \ t \ spectral_radius x" "\t. t\V \ t \ 0" "bdd_below V" "V \ {}" "Inf V \ 0" proof - show "V\{}" using V_def by blast show *: "t \ 0" if "t \ V" for t using that unfolding V_def using real_root_pos_pos_le by auto then show "bdd_below V" by (meson bdd_below_def) then show "Inf V \ 0" by (simp add: \V \ {}\ * cInf_greatest) show "\t. t\V \ t \ spectral_radius x" by (metis (mono_tags, lifting) \bdd_below V\ assms cInf_lower spectral_radius_def) qed lemma spectral_radius_nonneg [simp]: "spectral_radius x \ 0" by (simp add: spectral_radius_aux(5) spectral_radius_def) lemma spectral_radius_upper_bound [simp]: "(spectral_radius x)^n \ norm(x^n)" proof (cases) assume "\(n = 0)" have "root n (norm(x^n)) \ spectral_radius x" using spectral_radius_aux \n \ 0\ by auto then show ?thesis by (metis \n \ 0\ spectral_radius_nonneg norm_ge_zero not_gr0 power_mono real_root_pow_pos2) qed (simp) lemma spectral_radius_limit: "(\n. root n (norm(x^n))) \ spectral_radius x" proof - have "norm(x^(m+n)) \ norm(x^m) * norm(x^n)" for m n by (simp add: power_add norm_mult_ineq) then show ?thesis unfolding spectral_radius_def using submultiplicative_converges by auto qed end (*of Fekete.thy*) diff --git a/thys/Ergodic_Theory/Gouezel_Karlsson.thy b/thys/Ergodic_Theory/Gouezel_Karlsson.thy --- a/thys/Ergodic_Theory/Gouezel_Karlsson.thy +++ b/thys/Ergodic_Theory/Gouezel_Karlsson.thy @@ -1,1809 +1,1810 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Gouezel-Karlsson\ theory Gouezel_Karlsson imports Asymptotic_Density Kingman begin text \This section is devoted to the proof of the main ergodic result of the article "Subadditive and multiplicative ergodic theorems" by Gouezel and Karlsson~\cite{gouezel_karlsson}. It is a version of Kingman theorem ensuring that, for subadditive cocycles, there are almost surely many times $n$ where the cocycle is nearly additive at \emph{all} times between $0$ and $n$. This theorem is then used in this article to construct horofunctions characterizing the behavior at infinity of compositions of semi-contractions. This requires too many further notions to be implemented in current Isabelle/HOL, but the main ergodic result is completely proved below, in Theorem~\verb+Gouezel_Karlsson+, following the arguments in the paper (but in a slightly more general setting here as we are not making any ergodicity assumption). To simplify the exposition, the theorem is proved assuming that the limit of the subcocycle vanishes almost everywhere, in the locale \verb+Gouezel_Karlsson_Kingman+. The final result is proved by an easy reduction to this case. The main steps of the proof are as follows: \begin{itemize} \item assume first that the map is invertible, and consider the inverse map and the corresponding inverse subcocycle. With combinatorial arguments that only work for this inverse subcocycle, we control the density of bad times given some allowed error $d>0$, in a precise quantitative way, in Lemmas~\verb+upper_density_all_times+ and~\verb+upper_density_large_k+. We put these estimates together in Lemma~\verb+upper_density_delta+. \item These estimates are then transfered to the original time direction and the original subcocycle in Lemma~\verb+upper_density_good_direction_invertible+. The fact that we have quantitative estimates in terms of asymptotic densities is central here, just having some infiniteness statement would not be enough. \item The invertibility assumption is removed in Lemma~\verb+upper_density_good_direction+ by using the result in the natural extension. \item Finally, the main result is deduced in Lemma~\verb+infinite_AE+ (still assuming that the asymptotic limit vanishes almost everywhere), and in full generality in Theorem~\verb+Gouezel_Karlsson_Kingman+. \end{itemize} \ lemma upper_density_eventually_measure: fixes a::real assumes [measurable]: "\n. {x \ space M. P x n} \ sets M" and "emeasure M {x \ space M. upper_asymptotic_density {n. P x n} < a} > b" shows "\N. emeasure M {x \ space M. \n \ N. card ({n. P x n} \ {.. b" proof - define G where "G = {x \ space M. upper_asymptotic_density {n. P x n} < a}" define H where "H = (\N. {x \ space M. \n \ N. card ({n. P x n} \ {.. sets M" "\N. H N \ sets M" unfolding G_def H_def by auto have "G \ (\N. H N)" proof fix x assume "x \ G" then have "x \ space M" unfolding G_def by simp have "eventually (\n. card({n. P x n} \ {..x \ G\ unfolding G_def using upper_asymptotic_densityD by auto then obtain N where "\n. n \ N \ card({n. P x n} \ {.. H N" unfolding H_def using \x \ space M\ by auto then show "x \ (\N. H N)" by blast qed have "b < emeasure M G" using assms(2) unfolding G_def by simp also have "... \ emeasure M (\N. H N)" apply (rule emeasure_mono) using \G \ (\N. H N)\ by auto finally have "emeasure M (\N. H N) > b" by simp moreover have "(\N. emeasure M (H N)) \ emeasure M (\N. H N)" apply (rule Lim_emeasure_incseq) unfolding H_def incseq_def by auto ultimately have "eventually (\N. emeasure M (H N) > b) sequentially" by (simp add: order_tendsto_iff) then obtain N where "emeasure M (H N) > b" using eventually_False_sequentially eventually_mono by blast then show ?thesis unfolding H_def by blast qed locale Gouezel_Karlsson_Kingman = pmpt + fixes u::"nat \ 'a \ real" assumes subu: "subcocycle u" and subu_fin: "subcocycle_avg_ereal u > -\" and subu_0: "AE x in M. subcocycle_lim u x = 0" begin lemma int_u [measurable]: "integrable M (u n)" using subu unfolding subcocycle_def by auto text \Next lemma is Lemma 2.1 in~\cite{gouezel_karlsson}.\ lemma upper_density_all_times: assumes "d > (0::real)" shows "\c> (0::real). emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) x \ - c * l} < d} > 1 - d" proof - define f where "f = (\x. abs (u 1 x))" have [measurable]: "f \ borel_measurable M" unfolding f_def by auto define G where "G = {x \ space M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x \ (\n. u n x / n) \ 0}" have [measurable]: "G \ sets M" unfolding G_def by auto have "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" apply (rule birkhoff_theorem_AE_nonergodic) using subu unfolding f_def subcocycle_def by auto moreover have "AE x in M. (\n. u n x / n) \ 0" using subu_0 kingman_theorem_nonergodic(1)[OF subu subu_fin] by auto ultimately have "AE x in M. x \ G" unfolding G_def by auto then have "emeasure M G = 1" by (simp add: emeasure_eq_1_AE) define V where "V = (\c x. {n. \l \ {1..n}. u n x - u (n-l) x \ - c * l})" define Good where "Good = (\c. {x \ G. upper_asymptotic_density (V c x) < d})" have [measurable]: "Good c \ sets M" for c unfolding Good_def V_def by auto have I: "upper_asymptotic_density (V c x) \ real_cond_exp M Invariants f x / c" if "c>0" "x \ G" for c x proof - have [simp]: "c>0" "c \ 0" "c \ 0" using \c>0\ by auto define U where "U = (\n. abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n}))" have main: "u n x \ U n" for n proof (rule nat_less_induct) fix n assume H: "\m U m" consider "n = 0" | "n\1 \ n \ V c x" | "n\1 \ n \ V c x" by linarith then show "u n x \ U n" proof (cases) assume "n = 0" then show ?thesis unfolding U_def by auto next assume A: "n\1 \ n \ V c x" then have "n \ 1" by simp then have "n-1 {n}" using \1 \ n\ atLeastLessThanSuc by auto then have *: "card (V c x \ {1..n}) = card (V c x \ {1..n-1})" using A by auto have "u n x \ u (n-1) x + u 1 ((T^^(n-1)) x)" using \n\1\ subu unfolding subcocycle_def by (metis le_add_diff_inverse2) also have "... \ U (n-1) + f ((T^^(n-1)) x)" unfolding f_def using H \n-1 by auto also have "... = abs(u 0 x) + birkhoff_sum f (n-1) x + f ((T^^(n-1)) x) - c * card (V c x \ {1..n-1})" unfolding U_def by auto also have "... = abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n})" using * birkhoff_sum_cocycle[of f "n-1" 1 x] \1 \ n\ by auto also have "... = U n" unfolding U_def by simp finally show ?thesis by auto next assume B: "n\1 \ n \ V c x" then obtain l where l: "l\{1..n}" "u n x - u (n-l) x \ - c * l" unfolding V_def by blast then have "n-l < n" by simp have m: "- (r * ra) - r * rb = - (r * (rb + ra))" for r ra rb::real by (simp add: algebra_simps) have "card(V c x \ {1..n}) \ card ((V c x \ {1..n-l}) \ {n-l+1..n})" by (rule card_mono, auto) also have "... \ card (V c x \ {1..n-l}) + card {n-l+1..n}" by (rule card_Un_le) also have "... \ card (V c x \ {1..n-l}) + l" by auto finally have "card(V c x \ {1..n}) \ card (V c x \ {1..n-l}) + real l" by auto then have *: "-c * card (V c x \ {1..n-l}) - c * l \ -c * card(V c x \ {1..n})" using m by auto have "birkhoff_sum f ((n-l) + l) x = birkhoff_sum f (n-l) x + birkhoff_sum f l ((T^^(n-l))x)" by (rule birkhoff_sum_cocycle) moreover have "birkhoff_sum f l ((T^^(n-l))x) \ 0" unfolding f_def birkhoff_sum_def using sum_nonneg by auto ultimately have **: "birkhoff_sum f (n-l) x \ birkhoff_sum f n x" using l(1) by auto have "u n x \ u (n-l) x - c * l" using l by simp also have "... \ U (n-l) - c* l" using H \n-l < n\ by auto also have "... = abs(u 0 x) + birkhoff_sum f (n-l) x - c * card (V c x \ {1..n-l}) - c*l" unfolding U_def by auto also have "... \ abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n})" using * ** by simp finally show ?thesis unfolding U_def by auto qed qed have "(\n. abs(u 0 x) * (1/n) + birkhoff_sum f n x / n - u n x / n) \ abs(u 0 x) * 0 + real_cond_exp M Invariants f x - 0" apply (intro tendsto_intros) using \x \ G\ unfolding G_def by auto moreover have "(abs(u 0 x) + birkhoff_sum f n x - u n x)/n = abs(u 0 x) * (1/n) + birkhoff_sum f n x / n - u n x / n" for n by (auto simp add: add_divide_distrib diff_divide_distrib) ultimately have "(\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n) \ real_cond_exp M Invariants f x" by auto then have a: "limsup (\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n) = real_cond_exp M Invariants f x" by (simp add: assms lim_imp_Limsup) have "c * card (V c x \ {1..n})/n \ (abs(u 0 x) + birkhoff_sum f n x - u n x)/n" for n using main[of n] unfolding U_def by (simp add: divide_right_mono) then have "limsup (\n. c * card (V c x \ {1..n})/n) \ limsup (\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n)" by (simp add: Limsup_mono) then have b: "limsup (\n. c * card (V c x \ {1..n})/n) \ real_cond_exp M Invariants f x" using a by simp have "ereal(upper_asymptotic_density (V c x)) = limsup (\n. card (V c x \ {1..n})/n)" using upper_asymptotic_density_shift[of "V c x" 1 0] by auto also have "... = limsup (\n. ereal(1/c) * ereal(c * card (V c x \ {1..n})/n))" by auto also have "... = (1/c) * limsup (\n. c * card (V c x \ {1..n})/n)" by (rule limsup_ereal_mult_left, auto) also have "... \ ereal (1/c) * real_cond_exp M Invariants f x" by (rule ereal_mult_left_mono[OF b], auto) finally show "upper_asymptotic_density (V c x) \ real_cond_exp M Invariants f x / c" by auto qed { fix r::real obtain c::nat where "r / d < c" using reals_Archimedean2 by auto then have "r/d < real c+1" by auto then have "r / (real c+1) < d" using \d>0\ by (simp add: divide_less_eq mult.commute) then have "\c::nat. r / (real c+1) < d" by auto } then have unG: "(\c::nat. {x \ G. real_cond_exp M Invariants f x / (c+1) < d}) = G" by auto have *: "r < d * (real n + 1)" if "m \ n" "r < d * (real m + 1)" for m n r proof - have "d * (real m + 1) \ d * (real n + 1)" using \d>0\ \m \ n\ by auto then show ?thesis using \r < d * (real m + 1)\ by auto qed have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ emeasure M (\c::nat. {x \ G. real_cond_exp M Invariants f x / (c+1) < d})" apply (rule Lim_emeasure_incseq) unfolding incseq_def by (auto simp add: divide_simps *) then have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ emeasure M G" using unG by auto then have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ 1" using \emeasure M G = 1\ by simp then have "eventually (\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d} > 1 - d) sequentially" apply (rule order_tendstoD) apply (insert \0, auto simp add: ennreal_1[symmetric] ennreal_lessI simp del: ennreal_1) done then obtain c0 where c0: "emeasure M {x \ G. real_cond_exp M Invariants f x / (real c0+1) < d} > 1 - d" using eventually_sequentially by auto define c where "c = real c0 + 1" then have "c > 0" by auto have *: "emeasure M {x \ G. real_cond_exp M Invariants f x / c < d} > 1 - d" unfolding c_def using c0 by auto have "{x \ G. real_cond_exp M Invariants f x / c < d} \ {x \ space M. upper_asymptotic_density (V c x) < d}" apply auto using G_def apply blast using I[OF \c>0\] by fastforce then have "emeasure M {x \ G. real_cond_exp M Invariants f x / c < d} \ emeasure M {x \ space M. upper_asymptotic_density (V c x) < d}" apply (rule emeasure_mono) unfolding V_def by auto then have "emeasure M {x \ space M. upper_asymptotic_density (V c x) < d} > 1 - d" using * by auto then show ?thesis unfolding V_def using \c>0\ by auto qed text \Next lemma is Lemma 2.2 in~\cite{gouezel_karlsson}.\ lemma upper_density_large_k: assumes "d > (0::real)" "d \ 1" shows "\k::nat. emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d} > 1 - d" proof - have [simp]: "d>0" "d \ 0" "d \ 0" using \d>0\ by auto define rho where "rho = d * d * d / 4" have [simp]: "rho > 0" "rho \ 0" "rho \ 0" unfolding rho_def using assms by auto text \First step: choose a time scale $s$ at which all the computations will be done. the integral of $u_s$ should be suitably small -- how small precisely is given by $\rho$.\ have "ennreal(\x. abs(u n x / n) \M) = (\\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M)" for n proof - have "ennreal(\x. abs(u n x / n) \M) = (\\<^sup>+x. abs(u n x /n) \M)" apply (rule nn_integral_eq_integral[symmetric]) using int_u by auto also have "... = (\\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M)" apply (rule nn_integral_cong_AE) using subu_0 by auto finally show ?thesis by simp qed moreover have "(\n. \\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M) \ 0" by (rule kingman_theorem_nonergodic(3)[OF subu subu_fin]) ultimately have "(\n. ennreal(\x. abs(u n x / n) \M)) \ 0" by auto then have "(\n. (\x. abs(u n x / n) \M)) \ 0" by (simp add: ennreal_0[symmetric] del: ennreal_0) then have "eventually (\n. (\x. abs(u n x / n) \M) < rho) sequentially" by (rule order_tendstoD(2), auto) then obtain s::nat where [simp]: "s>0" and s_int: "(\x. abs(u s x / s) \M) < rho" by (metis (mono_tags, lifting) neq0_conv eventually_sequentially gr_implies_not0 linorder_not_le of_nat_0_eq_iff order_refl zero_neq_one) text \Second step: a truncation argument, to decompose $|u_1|$ as a sum of a constant (its contribution will be small if $k$ is large at the end of the argument) and of a function with small integral).\ have "(\n. (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x \M)) \ (\x. 0 \M)" proof (rule integral_dominated_convergence[where ?w = "\x. abs(u 1 x)"]) show "AE x in M. norm (abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ abs(u 1 x)" for n unfolding indicator_def by auto { fix x have "abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x = (0::real)" if "n > abs(u 1 x)" for n::nat unfolding indicator_def using that by auto then have "eventually (\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x = 0) sequentially" by (metis (mono_tags, lifting) eventually_at_top_linorder reals_Archimedean2 less_le_trans of_nat_le_iff) then have "(\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ 0" by (rule Lim_eventually) } then show "AE x in M. (\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ 0" by simp qed (auto simp add: int_u) then have "eventually (\n. (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x \M) < rho) sequentially" by (rule order_tendstoD(2), auto) then obtain Knat::nat where Knat: "Knat > 0" "(\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ Knat} x \M) < rho" by (metis (mono_tags, lifting) eventually_sequentially gr_implies_not0 neq0_conv linorder_not_le of_nat_0_eq_iff order_refl zero_neq_one) define K where "K = real Knat" then have [simp]: "K \ 0" "K>0" and K: "(\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ K} x \M) < rho" using Knat by auto define F where "F = (\x. abs(u 1 x) * indicator {x. abs(u 1 x) \ K} x)" have int_F [measurable]: "integrable M F" unfolding F_def apply (rule Bochner_Integration.integrable_bound[where ?f = "\x. abs(u 1 x)"]) unfolding indicator_def by (auto simp add: int_u) have "(\x. F x \M) = (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ K} x \M)" apply (rule integral_cong_AE) unfolding F_def by (auto simp add: indicator_def) then have F_int: "(\x. F x \M) < rho" using K by auto have F_pos: "F x \ 0" for x unfolding F_def by auto have u1_bound: "abs(u 1 x) \ K + F x" for x unfolding F_def indicator_def apply (cases "x \ {x \ space M. K \ \u 1 x\}") by auto define F2 where "F2 = (\x. F x + abs(u s x/s))" have int_F2 [measurable]: "integrable M F2" unfolding F2_def using int_F int_u[of s] by auto have F2_pos: "F2 x \ 0" for x unfolding F2_def using F_pos by auto have "(\x. F2 x \M) = (\x. F x \M) + (\x. abs(u s x/s) \M)" unfolding F2_def apply (rule Bochner_Integration.integral_add) using int_F int_u by auto then have F2_int: "(\x. F2 x \M) < 2 * rho" using F_int s_int by auto text \We can now choose $k$, large enough. The reason for our choice will only appear at the end of the proof.\ define k where "k = max (2 * s + 1) (nat(ceiling((2 * d * s + 2 * K * s)/(d/2))))" have "k > 2 * s" unfolding k_def by auto have "k \ (2 * d * s + 2 * K * s)/(d/2)" unfolding k_def by linarith then have "(2 * d * s + 2 * K * s)/k \ d/2" using \k > 2 * s\ by (simp add: divide_simps mult.commute) text \Third step: definition of a good set $G$ where everything goes well.\ define G where "G = {x \ space M. (\n. u n x / n) \ 0 \ (\n. birkhoff_sum (\x. abs(u s x / s)) n x / n) \ real_cond_exp M Invariants (\x. abs(u s x / s)) x \ (\n. birkhoff_sum F n x / n) \ real_cond_exp M Invariants F x \ real_cond_exp M Invariants F x + real_cond_exp M Invariants (\x. abs(u s x / s)) x = real_cond_exp M Invariants F2 x}" have [measurable]: "G \ sets M" unfolding G_def by auto have "AE x in M. (\n. u n x / n) \ 0" using kingman_theorem_nonergodic(1)[OF subu subu_fin] subu_0 by auto moreover have "AE x in M.(\n. birkhoff_sum (\x. abs(u s x / s)) n x / n) \ real_cond_exp M Invariants (\x. abs(u s x / s)) x" apply (rule birkhoff_theorem_AE_nonergodic) using int_u[of s] by auto moreover have "AE x in M. (\n. birkhoff_sum F n x / n) \ real_cond_exp M Invariants F x" by (rule birkhoff_theorem_AE_nonergodic[OF int_F]) moreover have "AE x in M. real_cond_exp M Invariants F x + real_cond_exp M Invariants (\x. abs(u s x / s)) x = real_cond_exp M Invariants F2 x" unfolding F2_def apply (rule AE_symmetric[OF real_cond_exp_add]) using int_u[of s] int_F int_u[of s] by auto ultimately have "AE x in M. x \ G" unfolding G_def by auto then have "emeasure M G = 1" by (simp add: emeasure_eq_1_AE) text \Estimation of asymptotic densities of bad times, for points in $G$. There is a trivial part, named $U$ below, that has to be treated separately because it creates problematic boundary effects.\ define U where "U = (\x. {n. \l \ {n-s<..n}. u n x - u (n-l) x \ - d * l})" define V where "V = (\x. {n. \l \ {k..n-s}. u n x - u (n-l) x \ - d * l})" text \Trivial estimate for $U(x)$: this set is finite for $x\in G$.\ have densU: "upper_asymptotic_density (U x) = 0" if "x \ G" for x proof - define C where "C = Max {abs(u m x) |m. m {n. u n x \ C - d * n}" proof (auto) fix n assume "n \ U x" then obtain l where l: "l\ {n-s <..n}" "u n x - u (n-l) x \ - d * l" unfolding U_def by auto define m where "m = n-l" have "m < s" unfolding m_def using l by auto have "u n x \ u m x - d * l" using l m_def by auto also have "... \ abs(u m x) - d * n + d * m" unfolding m_def using l - by (simp add: linordered_field_class.sign_simps(38) of_nat_diff) + by (simp add: algebra_simps of_nat_diff) also have "... \ Max {abs(u m x) |m. mm < s\ apply (auto) by (rule Max_ge, auto) also have "... \ Max {abs(u m x) |m. mm < s\ \d>0\ by auto finally show "u n x \ C - d * n" unfolding C_def by auto qed have "eventually (\n. u n x / n > -d/2) sequentially" apply (rule order_tendstoD(1)) using \x \ G\ \d>0\ unfolding G_def by auto then obtain N where N: "\n. n \ N \ u n x / n > - d/2" using eventually_sequentially by auto { fix n assume *: "u n x \ C - d * n" "n > N" then have "n \ N" "n > 0" by auto have "2 * u n x \ 2 * C - 2 * d * n" using * by auto moreover have "2 * u n x \ - d * n" using N[OF \n \ N\] \n > 0\ by (simp add: divide_simps) ultimately have "- d * n \ 2 * C - 2 * d * n" by auto then have "d * n \ 2 * C" by auto then have "n \ 2 * C / d" using \d>0\ by (simp add: mult.commute divide_simps) } then have "{n. u n x \ C - d * n} \ {..max (nat (floor(2*C/d))) N}" by (auto, meson le_max_iff_disj le_nat_floor not_le) then have "finite {n. u n x \ C - d * n}" using finite_subset by blast then have "finite (U x)" using * finite_subset by blast then show ?thesis using upper_asymptotic_density_finite by auto qed text \Main step: control of $u$ along the sequence $ns+t$, with a term $-(d/2) Card(V(x) \cap [1,ns+t])$ on the right. Then, after averaging in $t$, Birkhoff theorem will imply that $Card(V(x) \cap [1,n])$ is suitably small.\ define Z where "Z = (\t n x. Max {u i x|i. i< s} + (\i {1..n * s + t}))" have Main: "u (n * s + t) x \ Z t n x" if "t < s" for n x t proof (rule nat_less_induct[where ?n = n]) fix n assume H: "\m Z t m x" consider "n = 0"|"n>0 \ V x \ {(n-1) * s+t<..n * s+t} = {}"|"n>0 \ V x \ {(n-1) * s+t<..n * s+t} \ {}" by auto then show "u (n * s+t) x \ Z t n x" proof (cases) assume "n = 0" then have "V x \ {1..n * s + t} = {}" unfolding V_def using \t \k>2* s\ by auto then have *: "card(V x \ {1..n * s+t}) = 0" by simp have **: "0 \ (\in = 0\ by auto also have "... \ Max {u i x|i. i< s}" by (rule Max_ge, auto simp add: \t) also have "... \ Z t n x" unfolding Z_def birkhoff_sum_def using \n = 0\ * ** by auto finally show ?thesis by simp next assume A: "n>0 \ V x \ {(n-1) * s+t<..n * s+t} = {}" then have "n\1" by simp have "n * s + t = (n-1) * s + t + s" using \n\1\ by (simp add: add.commute add.left_commute mult_eq_if) have "V x \ {1..n * s + t} = V x \ {1..(n-1) * s + t} \ V x \ {(n-1) * s + t<..n * s + t}" using \n\1\ by (auto, simp add: mult_eq_if) then have *: "card(V x \ {1..n * s+t}) = card(V x \ {1..(n-1) * s+t})" using A by auto have **: "birkhoff_sum F ((n-1) * s + t) x \ birkhoff_sum F (n * s + t) x" unfolding birkhoff_sum_def apply (rule sum_mono2) using \n * s+t = (n-1) * s+t + s\ F_pos by auto have "(\i (\i (\in\1\ lessThan_Suc_atMost sum.lessThan_Suc[of "\i. abs(u s ((T^^(i* s+t))x))" "n-1", symmetric] by auto finally have ***: "(\i (\in\1\ by (simp add: add.commute add.left_commute mult_eq_if) also have "... \ u ((n-1) * s+t) x + u s ((T^^((n-1) * s+t)) x)" using subcocycle_ineq[OF subu, of "(n-1) * s+t" s x] by simp also have "... \ Max {u i x|i. i< s} + (\i {1..(n-1) * s+t}) + u s ((T^^((n-1) * s+t)) x)" using H \n\1\ unfolding Z_def by auto also have "... \ Max {u i x|i. i< s} + (\i {1..n * s+t})" using * ** *** by auto also have "... \ Z t n x" unfolding Z_def by (auto simp add: divide_simps) finally show ?thesis by simp next assume B: "n>0 \ V x \ {(n-1) * s+t<..n * s+t} \ {}" then have [simp]: "n>0" "n\1" "n \ 0" by auto obtain m where m: "m \ V x \ {(n-1) * s + t<..n * s + t}" using B by blast then obtain l where l: "l \ {k..m-s}" "u m x - u (m-l) x \ - d * l" unfolding V_def by auto then have "m-s>0" using \k>2* s\ by auto then have "m-l \ s" using l by auto define p where "p = (m-l-t) div s" have p1: "m-l \ p * s + t" unfolding p_def using \m-l \ s\ \s>t\ minus_mod_eq_div_mult [symmetric, of "m - l - t" s] by simp have p2: "m-l < p* s + t + s" unfolding p_def using \m-l \ s\ \s>t\ div_mult_mod_eq[of "m-l-t" s] mod_less_divisor[OF \s>0\, of "m-l-t"] by linarith then have "l \ m - p * s - t -s" by auto then have "l \ (n-1) * s + t -p * s - t- s" using m by auto then have "l + 2 * s\ (n * s + t) - (p * s+t)" by (simp add: diff_mult_distrib) have "(p+1) * s + t \ (n-1) * s + t" using \k> 2 * s\ m l(1) p1 by (auto simp add: algebra_simps) then have "p+1 \ n-1" using \s>0\ by (meson add_le_cancel_right mult_le_cancel2) then have "p \ n-1" "p (n * s + t)" using m l(1) p1 by (auto simp add: algebra_simps) then have "(1::real) \ ((n * s + t) - (p* s + t)) / k" using \k > 2* s\ by auto have In: "u (n * s + t) x \ u m x + (\i \ {(n-1) * s + t..i \ {(n-1) * s+t.. 0" by (rule sum_nonneg, auto) then show ?thesis using True by auto next case False then have m2: "n * s + t - m >0" "(n-1) * s+t \ m" using m by auto have "birkhoff_sum (u 1) (n * s+t-m) ((T^^m) x) = (\iij \ {m.. (\j \ {m.. (\j \ {(n-1) * s+t..j \ {m..j \ {(n-1) * s+t.. (\j \ {(n-1) * s+t.. u m x + u (n * s+t-m) ((T^^m) x)" using subcocycle_ineq[OF subu, of m "n * s+t-m"] m2 by auto also have "... \ u m x + birkhoff_sum (u 1) (n * s+t-m) ((T^^m) x)" using subcocycle_bounded_by_birkhoff1[OF subu \n * s+t - m >0\, of "(T^^m)x"] by simp finally show ?thesis using * by auto qed have Ip: "u (m-l) x \ u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x)))" proof (cases "m-l = p* s+t") case True have "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) \ 0" by (rule sum_nonneg, auto) then show ?thesis using True by auto next case False then have "m-l - (p* s+t) > 0" using p1 by auto have I: "p * s + t + (m - l - (p * s + t)) = m - l" using p1 by auto have "birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x) = (\iij \ {p* s+t.. (\j \ {p* s+t.. (\j \ {p* s+t..j \ {m-l..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" by auto also have "... = (\j \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" apply (rule sum.atLeastLessThan_concat) using p1 p2 by auto finally have *: "birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x) \ (\j \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" by auto have "u (m-l) x \ u (p* s+t) x + u (m-l - (p* s+t)) ((T^^(p* s+t)) x)" using subcocycle_ineq[OF subu, of "p* s+t" "m-l - (p* s+t)" x] I by auto also have "... \ u (p* s+t) x + birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x)" using subcocycle_bounded_by_birkhoff1[OF subu \m-l - (p* s+t) > 0\, of "(T^^(p* s+t)) x"] by simp finally show ?thesis using * by auto qed have "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) \ (\i \ {p* s+t..<(p+1)* s+t}. K + F ((T^^i) x))" apply (rule sum_mono) using u1_bound by auto moreover have "(\i \ {(n-1) * s+t.. (\i \ {(n-1) * s+t..i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. (\i \ {p* s+t..<(p+1)* s+t}. K + F ((T^^i) x)) + (\i \ {(n-1) * s+t..i \ {p* s+t..<(p+1)* s+t}. F ((T^^i) x)) + (\i \{(n-1) * s+t.. 2* K * s + (\i \ {p* s+t..<(n-1) * s+t}. F ((T^^i) x)) + (\i \{(n-1) * s+t..(p+1)* s+t\(n-1) * s+t\ F_pos by auto also have "... = 2* K * s + (\i \ {p* s+t..p\n-1\ by auto finally have A0: "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. 2* K * s + (\i \ {p* s+t.. {p * s + t<.. n * s+t}) \ card {p * s + t<.. n * s+t}" by (rule card_mono, auto) have "2 * d * s + 2 * K * s > 0" using \K>0\ \s>0\ \d>0\ by (metis add_pos_pos mult_2 mult_zero_left of_nat_0_less_iff pos_divide_less_eq times_divide_eq_right) then have "2 * d * s + 2 * K * s \ ((n * s + t) - (p* s + t)) * ((2 * d * s + 2 * K * s) / k)" using \1 \ ((n * s + t) - (p* s + t)) / k\ by (simp add: le_divide_eq_1 pos_le_divide_eq) also have "... \ ((n * s + t) - (p* s + t)) * (d/2)" apply (rule mult_left_mono) using \(2 * d * s + 2 * K * s)/k \ d/2\ by auto finally have "2 * d * s + 2 * K * s \ ((n * s + t) - (p* s + t)) * (d/2)" by auto then have "-d * ((n * s+t) - (p* s+t)) + 2 * d * s + 2 * K * s \ -d * ((n * s+t) - (p* s+t)) + ((n * s + t) - (p* s + t)) * (d/2)" by linarith also have "... = (-d/2) * card {p * s + t<.. n * s+t}" by auto also have "... \ (-d/2) * card(V x \ {p * s + t<.. n * s+t})" using \card(V x \ {p * s + t<.. n * s+t}) \ card {p * s + t<.. n * s+t}\ by auto finally have A1: "-d * ((n * s+t) - (p* s+t)) + 2 * d * s + 2 * K * s \ (-d/2) * card(V x \ {p * s + t<.. n * s+t})" by simp have "V x \ {1.. n * s+t} = V x \ {1..p * s + t} \ V x \ {p * s + t<.. n * s+t}" using \p * s + t + k \ n * s + t\ by auto then have "card (V x \ {1.. n * s+t}) = card(V x \ {1..p * s + t} \ V x \ {p * s + t<.. n * s+t})" by auto also have "... = card (V x \ {1..p * s + t}) + card (V x \ {p * s + t<.. n * s+t})" by (rule card_Un_disjoint, auto) finally have A2: "card (V x \ {1..p * s + t}) + card (V x \ {p * s + t<.. n * s+t}) = card (V x \ {1.. n * s+t})" by simp have A3: "(\i (\ip\n-1\ by auto have A4: "birkhoff_sum F (p * s + t) x + (\i \ {p* s+t..p\n-1\ by auto have "u (n * s+t) x \ u m x + (\i \ {(n-1) * s+t.. (u m x - u (m-l) x) + u (m-l) x + (\i \ {(n-1) * s+t.. - d * l + u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. - d * ((n * s+t) - (p* s+t)) + 2*d* s + u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t..l + 2* s\ (n * s+t) - (p* s+t)\ apply (auto simp add: algebra_simps) by (metis assms(1) distrib_left mult.commute mult_2 of_nat_add of_nat_le_iff real_mult_le_cancel_iff2) also have "... \ -d * ((n * s+t) - (p* s+t)) + 2*d* s + Z t p x + 2* K * s + (\i \ {p* s+t..p by auto also have "... \ Z t p x - d/2 * card(V x \ {p * s + t<.. n * s+t}) + (\i \ {p* s+t..i {1..p * s + t}) - d/2 * card(V x \ {p * s + t<.. n * s+t}) + (\i \ {p* s+t.. Max {u i x |i. i < s} + (\ii \ {p* s+t.. {1..p * s + t}) - d/2 * card(V x \ {p * s + t<.. n * s+t})" using A3 by auto also have "... = Z t n x" unfolding Z_def using A2 A4 by (auto simp add: algebra_simps, metis distrib_left of_nat_add) finally show ?thesis by simp qed qed have Main2: "(d/2) * card(V x \ {1..n}) \ Max {u i x|i. i< s} + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x + birkhoff_sum F (n + 2 * s) x + (1/s) * (\i< 2 * s. abs(u (n+i) x))" for n x proof - define N where "N = (n div s) + 1" then have "n \ N * s" using \s > 0\ dividend_less_div_times less_or_eq_imp_le by auto have "N * s \ n + s" by (auto simp add: N_def) have eq_t: "(d/2) * card(V x \ {1..n}) \ abs(u(N* s+t) x) + (Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x) + (\i birkhoff_sum F (n+ 2* s) x" unfolding birkhoff_sum_def apply (rule sum_mono2) using F_pos \N * s \ n + s\ \t by auto have "card(V x \ {1..n}) \ card(V x \ {1..N* s+t})" apply (rule card_mono) using \n \ N * s\ by auto then have "(d/2) * card(V x \ {1..n}) \ (d/2) * card(V x \ {1..N* s+t})" by auto also have "... \ - u (N* s+t) x + Max {u i x|i. i< s} + (\it < s\, of N x] unfolding Z_def by auto also have "... \ abs(u(N* s+t) x) + Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x + (\iti\{N* s.. (\i\{n..n \ N * s\ \N * s \ n + s\ by auto also have "... = (\i<2* s. abs (u (n+i) x))" by (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - n"], auto) finally have **: "(\t (\i<2* s. abs (u (n+i) x))" by simp have "(\tii (\iN * s \ n + s\ by auto finally have ***: "(\ti s * birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x" unfolding birkhoff_sum_def using \s>0\ by (auto simp add: sum_divide_distrib[symmetric]) have ****: "s * (\ii {1..n}) = (\t {1..n}))" by auto also have "... \ (\tittti (\i<2* s. abs (u (n+i) x)) + s * (Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x) + s * birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x" using ** *** by auto also have "... = s * ((1/s) * (\i<2* s. abs (u (n+i) x)) + Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x)" by (auto simp add: divide_simps mult.commute distrib_left) finally show ?thesis by auto qed have densV: "upper_asymptotic_density (V x) \ (2/d) * real_cond_exp M Invariants F2 x" if "x \ G" for x proof - have *: "(\n. abs(u n x/n)) \ 0" apply (rule tendsto_rabs_zero) using \x\G\ unfolding G_def by auto define Bound where "Bound = (\n. (Max {u i x|i. i< s}*(1/n) + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x / n + birkhoff_sum F (n + 2* s) x / n + (1/s) * (\i<2* s. abs(u (n+i) x) / n)))" have "Bound \ (Max {u i x|i. i< s} * 0 + real_cond_exp M Invariants (\x. abs(u s x/s)) x + real_cond_exp M Invariants F x + (1/s) * (\i < 2 * s. 0))" unfolding Bound_def apply (intro tendsto_intros) using \x\G\ * unfolding G_def by auto moreover have "real_cond_exp M Invariants (\x. abs(u s x/s)) x + real_cond_exp M Invariants F x = real_cond_exp M Invariants F2 x" using \x \ G\ unfolding G_def by auto ultimately have B_conv: "Bound \ real_cond_exp M Invariants F2 x" by simp have *: "(d/2) * card(V x \ {1..n}) / n \ Bound n" for n proof - have "(d/2) * card(V x \ {1..n}) / n \ (Max {u i x|i. i< s} + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x + birkhoff_sum F (n + 2* s) x + (1/s) * (\i<2* s. abs(u (n+i) x)))/n" using Main2[of x n] using divide_right_mono of_nat_0_le_iff by blast also have "... = Bound n" unfolding Bound_def by (auto simp add: add_divide_distrib sum_divide_distrib[symmetric]) finally show ?thesis by simp qed have "ereal(d/2 * upper_asymptotic_density (V x)) = ereal(d/2) * ereal(upper_asymptotic_density (V x))" by auto also have "... = ereal (d/2) * limsup(\n. card(V x \ {1..n}) / n)" using upper_asymptotic_density_shift[of "V x" 1 0] by auto also have "... = limsup(\n. ereal (d/2) * (card(V x \ {1..n}) / n))" by (rule limsup_ereal_mult_left[symmetric], auto) also have "... \ limsup Bound" apply (rule Limsup_mono) using * not_eventuallyD by auto also have "... = ereal(real_cond_exp M Invariants F2 x)" using B_conv convergent_limsup_cl convergent_def convergent_real_imp_convergent_ereal limI by force finally have "d/2 * upper_asymptotic_density (V x) \ real_cond_exp M Invariants F2 x" by auto then show ?thesis by (simp add: divide_simps mult.commute) qed define epsilon where "epsilon = 2 * rho / d" have [simp]: "epsilon > 0" "epsilon \ 0" "epsilon \ 0" unfolding epsilon_def by auto have "emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon} \ (1/epsilon) * (\x. real_cond_exp M Invariants F2 x \M)" apply (intro integral_Markov_inequality real_cond_exp_pos real_cond_exp_int(1)) by (auto simp add: int_F2 F2_pos) also have "... = (1/epsilon) * (\x. F2 x \M)" apply (intro arg_cong[where f = ennreal]) by (simp, rule real_cond_exp_int(2), simp add: int_F2) also have "... < (1/epsilon) * 2 * rho" using F2_int by (intro ennreal_lessI) (auto simp add: divide_simps) also have "... = d" unfolding epsilon_def by auto finally have *: "emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon} < d" by simp define G2 where "G2 = {x \ G. real_cond_exp M Invariants F2 x < epsilon}" have [measurable]: "G2 \ sets M" unfolding G2_def by simp have "1 = emeasure M G" using \emeasure M G = 1\ by simp also have "... \ emeasure M (G2 \ {x\space M. real_cond_exp M Invariants F2 x \ epsilon})" apply (rule emeasure_mono) unfolding G2_def using sets.sets_into_space[OF \G \ sets M\] by auto also have "... \ emeasure M G2 + emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon}" by (rule emeasure_subadditive, auto) also have "... < emeasure M G2 + d" using * by auto finally have "1 - d < emeasure M G2" using emeasure_eq_measure \d \ 1\ by (auto intro!: ennreal_less_iff[THEN iffD2] simp del: ennreal_plus simp add: ennreal_plus[symmetric]) have "upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d" if "x \ G2" for x proof - have "x \ G" using \x \ G2\ unfolding G2_def by auto have "{n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} \ U x \ V x" unfolding U_def V_def by fastforce then have "upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} \ upper_asymptotic_density (U x \ V x)" by (rule upper_asymptotic_density_subset) also have "... \ upper_asymptotic_density (U x) + upper_asymptotic_density (V x)" by (rule upper_asymptotic_density_union) also have "... \ (2/d) * real_cond_exp M Invariants F2 x" using densU[OF \x \ G\] densV[OF \x \ G\] by auto also have "... < (2/d) * epsilon" using \x \ G2\ unfolding G2_def by (simp add: divide_simps) text \This is where the choice of $\rho$ at the beginning of the proof is relevant: we choose it so that the above term is at most $d$.\ also have "... = d" unfolding epsilon_def rho_def by auto finally show ?thesis by simp qed then have "G2 \ {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d}" using sets.sets_into_space[OF \G2 \ sets M\] by blast then have "emeasure M G2 \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d}" by (rule emeasure_mono, auto) then have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d} > 1 -d" using \emeasure M G2 > 1 - d\ by auto then show ?thesis by blast qed text \The two previous lemmas are put together in the following lemma, corresponding to Lemma 2.3 in~\cite{gouezel_karlsson}.\ lemma upper_density_delta: fixes d::real assumes "d > 0" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. \(N::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * N} > 1 - d" proof - define d2 where "d2 = d/2" have [simp]: "d2 > 0" unfolding d2_def using assms by simp + then have "\ d2 < 0" using not_less [of d2 0] by (simp add: less_le) have "d2/2 > 0" by simp obtain c0 where c0: "c0> (0::real)" "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} < d2/2} > 1 - (d2/2)" using upper_density_all_times[OF \d2/2 > 0\] by blast have "\N. emeasure M {x \ space M. \n \ N. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. 1 - (d2/2)" apply (rule upper_density_eventually_measure) using c0(2) by auto then obtain N1 where N1: "emeasure M {x \ space M. \B \ N1. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. 1 - (d2/2)" by blast define O1 where "O1 = {x \ space M. \B \ N1. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. sets M" unfolding O1_def by auto have "emeasure M O1 > 1 - (d2/2)" unfolding O1_def using N1 by auto have *: "\N. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" if "e>0" "e \ 1" for e::real proof - obtain k where k: "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} < e} > 1 - e" using upper_density_large_k[OF \e>0\ \e \ 1\] by blast then obtain N0 where N0: "emeasure M {x \ space M. \B \ N0. card({n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" using upper_density_eventually_measure[OF _ k] by auto define N where "N = max k N0" have "emeasure M {x \ space M. \B \ N0. card({n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} \ {.. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. space M" "\B\N0. card ({n. \l\{k..n}. u n x - u (n - l) x \ - (e * real l)} \ {.. B" have "card({n. \l \ {N..n}. u n x - u (n-l) x \ - (e * real l)} \ {.. card({n. \l \ {k..n}. u n x - u (n-l) x \ -(e * real l)} \ {..l \ {N..n}. u n x - u (n-l) x \ - (e * real l)} \ {.. card({n. \l \ {k..n}. u n x - u (n-l) x \ -(e * real l)} \ {..B\N\ unfolding N_def by auto finally show "card ({n. \l\{N..n}. u n x - u (n - l) x \ - (e * real l)} \ {.. space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" using N0 by simp then show ?thesis by auto qed define Ne where "Ne = (\(e::real). SOME N. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e)" have Ne: "emeasure M {x \ space M. \B \ Ne e. card({n. \l \ {Ne e..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" if "e>0" "e \ 1" for e::real unfolding Ne_def by (rule someI_ex[OF *[OF that]]) define eps where "eps = (\(n::nat). d2 * (1/2)^n)" have [simp]: "eps n > 0" for n unfolding eps_def by auto then have [simp]: "eps n \ 0" for n by (rule less_imp_le) have "eps n \ (1 / 2) * 1" for n unfolding eps_def d2_def using \d \ 1\ by (intro mult_mono power_le_one) auto also have "\ < 1" by auto finally have [simp]: "eps n < 1" for n by simp then have [simp]: "eps n \ 1" for n by (rule less_imp_le) have "(\n. d2 * (1/2)^n) \ d2 * 0" by (rule tendsto_mult, auto simp add: LIMSEQ_realpow_zero) then have "eps \ 0" unfolding eps_def by auto define Nf where "Nf = (\N. (if (N = 0) then 0 else if (N = 1) then N1 + 1 else max (N1+1) (Max {Ne(eps n)|n. n \ N}) + N))" have "Nf N < Nf (N+1)" for N proof - consider "N = 0" | "N = 1" | "N>1" by fastforce then show ?thesis proof (cases) assume "N>1" have "Max {Ne (eps n) |n. n \ N} \ Max {Ne (eps n) |n. n \ Suc N}" by (rule Max_mono, auto) then show ?thesis unfolding Nf_def by auto qed (auto simp add: Nf_def) qed then have "strict_mono Nf" using strict_mono_Suc_iff by auto define On where "On = (\(N::nat). (if (N = 1) then O1 else {x \ space M. \B \ Nf N. card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. sets M" for N unfolding On_def by auto have "emeasure M (On N) > 1 - eps N" if "N>0" for N proof - consider "N = 1" | "N>1" using \N>0\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis unfolding On_def eps_def using \emeasure M O1 > 1 - (d2/2)\ by auto next case 2 have "Ne (eps N) \ Max {Ne(eps n)|n. n \ N}" by (rule Max.coboundedI, auto) also have "... \ Nf N" unfolding Nf_def using \N>1\ by auto finally have "Ne (eps N) \ Nf N" by simp have "1 - eps N < emeasure M {x \ space M. \B \ Ne(eps N). card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. emeasure M {x \ space M. \B \ Nf N. card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. space M" "\n\Ne (eps N). card ({n. \l\{Ne (eps N)..n}. u n x - u (n - l) x \ - (eps N * l)} \ {.. n" have "card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ -(eps N) * l} \ {..Ne (eps N) \ Nf N\ by auto then have "real(card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ -(eps N) * l} \ {..n \ Nf N\ \Ne (eps N) \ Nf N\ by auto finally show "real (card ({n. \l\{Nf N..n}. u n x - u (n - l) x \ - (eps N * l)} \ {..N>1\ by auto finally show ?thesis by simp qed qed then have *: "emeasure M (On (N+1)) > 1 - eps (N+1)" for N by simp define Ogood where "Ogood = (\N. On (N+1))" have [measurable]: "Ogood \ sets M" unfolding Ogood_def by auto have "emeasure M Ogood \ 1 - (\N. eps(N+1))" unfolding Ogood_def apply (intro emeasure_intersection, auto) using * by (auto simp add: eps_def summable_mult summable_divide summable_geometric less_imp_le) moreover have "(\N. eps(N+1)) = d2" unfolding eps_def apply (subst suminf_mult) using sums_unique[OF power_half_series, symmetric] by (auto intro!: summable_divide summable_geometric) finally have "emeasure M Ogood \ 1 - d2" by simp then have "emeasure M Ogood > 1 - d" unfolding d2_def using \d>0\ \d \ 1\ by (simp add: emeasure_eq_measure field_sum_of_halves ennreal_less_iff) have Ogood_union: "Ogood = (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})" apply auto using sets.sets_into_space[OF \Ogood \ sets M\] apply blast proof - fix x define M where "M = Max {abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}}" obtain N::nat where "N > M" using reals_Archimedean2 by blast have "finite { (n, l) | n l. n \ {1..Nf 1} \ l \ {1..n}}" by (rule finite_subset[where ?B = "{1.. Nf 1} \ {1..Nf 1}"], auto) moreover have "{abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}} = (\ (n, l). abs(u n x - u (n-l) x)/l)` { (n, l) | n l. n \ {1..Nf 1} \ l \ {1..n}}" by auto ultimately have fin: "finite {abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}}" by auto { fix n l assume nl: "n \ {1..Nf 1} \ l \ {1..n}" then have "real l>0" by simp have "abs(u n x - u (n-l) x)/l \ M" unfolding M_def apply (rule Max_ge) using fin nl by auto then have "abs(u n x - u (n-l) x)/l < real N" using \N>M\ by simp then have "abs(u n x - u (n-l) x)< real N * l" using \0 < real l\ pos_divide_less_eq by blast then have "u n x - u (n-l) x > - (real N * l)" by simp } then have "\n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real N * real l) < u n x - u (n - l) x" by auto then show "\N. \n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real N * real l) < u n x - u (n - l) x" by auto qed have "(\K. emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})) \ emeasure M (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})" apply (rule Lim_emeasure_incseq, auto) unfolding incseq_def apply auto proof - fix m n x na l assume "m \ (n::nat)" "\n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real m * real l) < u n x - u (n - l) x" "Suc 0 \ l" "l \ na" "na \ Nf (Suc 0)" then have "- (real m * real l) < u na x - u (na - l) x" by auto moreover have "- (real n * real l) \ - (real m * real l)" using \m \ n\ by (simp add: mult_mono) ultimately show "- (real n * real l) < u na x - u (na - l) x" by auto qed moreover have "emeasure M (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d" using Ogood_union \emeasure M Ogood > 1 - d\ by auto ultimately have a: "eventually (\K. emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d) sequentially" by (rule order_tendstoD(1)) have b: "eventually (\K. K \ max c0 d2) sequentially" using eventually_at_top_linorder nat_ceiling_le_eq by blast have "eventually (\K. K \ max c0 d2 \ emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d) sequentially" by (rule eventually_elim2[OF a b], auto) then obtain K where K: "K\max c0 d2" "emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d" using eventually_False_sequentially eventually_elim2 by blast define Og where "Og = Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}" have [measurable]: "Og \ sets M" unfolding Og_def by simp have "emeasure M Og > 1 - d" unfolding Og_def using K by simp have fin: "finite {N. Nf N \ n}" for n using pseudo_inverse_finite_set[OF filterlim_subseq[OF \strict_mono Nf\]] by auto define prev_N where "prev_N = (\n. Max {N. Nf N \ n})" define delta where "delta = (\l. if (prev_N l \ 1) then K else eps (prev_N l))" have "\l. delta l > 0" unfolding delta_def using \K\max c0 d2\ \c0>0\ by auto have "LIM n sequentially. prev_N n :> at_top" unfolding prev_N_def apply (rule tendsto_at_top_pseudo_inverse2) using \strict_mono Nf\ by (simp add: filterlim_subseq) then have "eventually (\l. prev_N l > 1) sequentially" by (simp add: filterlim_iff) then have "eventually (\l. delta l = eps(prev_N l)) sequentially" unfolding delta_def by (simp add: eventually_mono) moreover have "(\l. eps(prev_N l)) \ 0" by (rule filterlim_compose[OF \eps \ 0\ \LIM n sequentially. prev_N n :> at_top\]) ultimately have "delta \ 0" by (simp add: filterlim_cong) have "delta n \ K" for n proof - have *: "d2 * (1 / 2) ^ prev_N n \ real K * 1" apply (rule mult_mono') using \K \ max c0 d2\ \d2>0\ by (auto simp add: power_le_one less_imp_le) then show ?thesis unfolding delta_def apply auto unfolding eps_def using * by auto qed define bad_times where "bad_times = (\x. {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ (\N\{2..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)}))" have card_bad_times: "card (bad_times x \ {.. d2 * B" if "x \ Og" for x B proof - have "(\N\{.. (\N. (1/2)^N)" by (rule sum_le_suminf, auto simp add: summable_geometric) also have "... = 2" using suminf_geometric[of "1/(2::real)"] by auto finally have *: "(\N\{.. 2" by simp have "(\ N \ {2.. (\ N \ {2..N\{2..N\{..N\{..N\{.. (d2 * (1/4) * B) * 2" apply (rule mult_left_mono) using * \d2 > 0\ apply auto by (metis \0 < d2\ mult_eq_0_iff mult_le_0_iff not_le of_nat_eq_0_iff of_nat_le_0_iff) finally have I: "(\ N \ {2.. d2/2 * B" by auto have "x \ On 1" using \x \ Og\ unfolding Og_def Ogood_def by auto then have "x \ O1" unfolding On_def by auto have B1: "real(card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" for B proof (cases "B \ N1") case True have "card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. card({n. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" using \x \ O1\ unfolding O1_def using True by auto finally show ?thesis by auto next case False then have "B < Nf 1" unfolding Nf_def by auto then have "{n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" - by (metis \0 < d2 / 2\ divide_le_eq div_0 linordered_field_class.sign_simps(24) of_nat_0 of_nat_0_le_iff) + using \\ d2 < 0\ by simp finally show ?thesis by simp qed have BN: "real(card ({n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" if "N \ 2" for N B proof - have "x \ On ((N-1) + 1)" using \x \ Og\ unfolding Og_def Ogood_def by auto then have "x \ On N" using \N\2\ by auto show ?thesis proof (cases "B \ Nf N") case True have "card ({n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card ({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" using \x \ On N\ \N\2\ True unfolding On_def by auto finally show ?thesis by simp next case False then have "{n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" by (metis \\n. 0 < eps n\ le_less mult_eq_0_iff mult_pos_pos of_nat_0 of_nat_0_le_iff) finally show ?thesis by simp qed qed { fix N assume "N \ B" have "Nf N \ B" using seq_suble[OF \strict_mono Nf\, of N] \N \ B\ by simp then have "{Nf N..} \ {.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {..N\{B..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {2.. {B..}" by auto then have "(\N\{2..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (\N\{B..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {..N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {..N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. real(card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (d2/2 * B) + (\ N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (d2/2 * B) + (\ N \ {2.. (d2/2 * B) + (d2/2*B)" using I by auto finally show ?thesis by simp qed have ineq_on_Og: "u n x - u (n-l) x > - delta l * l" if "l \ {1..n}" "n \ bad_times x" "x \ Og" for n x l proof - consider "n < Nf 1" | "n \ Nf 1 \ prev_N l \ 1" | "n \ Nf 1 \ prev_N l \ 2" by linarith then show ?thesis proof (cases) assume "n < Nf 1" then have "{N. Nf N \ n} = {0}" apply auto using \strict_mono Nf\ unfolding strict_mono_def apply (metis le_trans less_Suc0 less_imp_le_nat linorder_neqE_nat not_less) unfolding Nf_def by auto then have "prev_N n = 0" unfolding prev_N_def by auto moreover have "prev_N l \ prev_N n" unfolding prev_N_def apply (rule Max_mono) using \l \ {1..n}\ fin apply auto unfolding Nf_def by auto ultimately have "prev_N l = 0" using \prev_N l \ prev_N n\ by auto then have "delta l = K" unfolding delta_def by auto have "1 \ {N. Nf N \ n}" using fin[of n] by (metis (full_types) Max_ge \prev_N n = 0\ fin not_one_le_zero prev_N_def) then have "n < Nf 1" by auto moreover have "n \ 1" using \l \ {1..n}\ by auto ultimately have "n \ {1..Nf 1}" by auto then have "u n x - u (n-l) x > - (real K * l)" using \x \ Og\ unfolding Og_def using \l \ {1..n}\ by auto then show ?thesis using \delta l = K\ by auto next assume H: "n \ Nf 1 \ prev_N l \ 1" then have "delta l = K" unfolding delta_def by auto have "n \ {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)}" using \n \ bad_times x\ unfolding bad_times_def by auto then have "u n x - u (n-l) x > - (c0 * l)" using H \l \ {1..n}\ by force moreover have "- (c0 * l) \ - (real K * l)" using K(1) by (simp add: mult_mono) ultimately show ?thesis using \delta l = K\ by auto next assume H: "n \ Nf 1 \ prev_N l \ 2" define N where "N = prev_N l" have "N \ 2" unfolding N_def using H by auto have "prev_N l \ {N. Nf N \ l}" unfolding prev_N_def apply (rule Max_in, auto simp add: fin) unfolding Nf_def by auto then have "Nf N \ l" unfolding N_def by auto then have "Nf N \ n" using \l \ {1..n}\ by auto have "n \ {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)}" using \n \ bad_times x\ \N\2\ unfolding bad_times_def by auto then have "u n x - u (n-l) x > - (eps N * l)" using \Nf N \ n\ \Nf N \ l\ \l \ {1..n}\ by force moreover have "eps N = delta l" unfolding delta_def N_def using H by auto ultimately show ?thesis by auto qed qed have "Og \ {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B}" proof (auto) fix x assume "x \ Og" then show "x \ space M" unfolding Og_def by auto next fix x B assume "x \ Og" have *: "{n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ bad_times x \ {..x\Og\ by force have "card {n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ card (bad_times x \ {.. d2 * B" using card_bad_times \x \ Og\ by auto also have "... \ d * B" unfolding d2_def using \d>0\ by auto finally show "card {n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ d * B" by simp qed then have "emeasure M Og \ emeasure M {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B}" by (rule emeasure_mono, auto) then have "emeasure M {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B} > 1-d" using \emeasure M Og > 1 - d\ by auto then show ?thesis using \delta \ 0\ \\l. delta l > 0\ by auto qed text \We go back to the natural time direction, by using the previous result for the inverse map and the inverse subcocycle, and a change of variables argument. The price to pay is that the estimates we get are weaker: we have a control on a set of upper asymptotic density close to $1$, while having a set of lower asymptotic density close to $1$ as before would be stronger. This will nevertheless be sufficient for our purposes below.\ lemma upper_density_good_direction_invertible: assumes "invertible_qmpt" "d>(0::real)" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ ennreal(1-d)" proof - interpret I: Gouezel_Karlsson_Kingman M Tinv "(\n x. u n ((Tinv^^n) x))" proof show "Tinv \ quasi_measure_preserving M M" using Tinv_qmpt[OF \invertible_qmpt\] unfolding qmpt_def qmpt_axioms_def by simp show "Tinv \ measure_preserving M M" using Tinv_mpt[OF \invertible_qmpt\] unfolding mpt_def mpt_axioms_def by simp show "mpt.subcocycle M Tinv (\n x. u n ((Tinv ^^ n) x))" using subcocycle_u_Tinv[OF subu \invertible_qmpt\] by simp show "- \ < subcocycle_avg_ereal (\n x. u n ((Tinv ^^ n) x))" using subcocycle_avg_ereal_Tinv[OF subu \invertible_qmpt\] subu_fin by simp show "AE x in M. fmpt.subcocycle_lim M Tinv (\n x. u n ((Tinv ^^ n) x)) x = 0" using subcocycle_lim_Tinv[OF subu \invertible_qmpt\] subu_0 by auto qed have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by simp define e where "e = d * d / 2" have "e>0" "e\1" unfolding e_def using \d>0\ \d \ 1\ by (auto, meson less_imp_le mult_left_le one_le_numeral order_trans) obtain delta::"nat \ real" where d: "\l. delta l > 0" "delta \ 0" "emeasure M {x \ space M. \N. card {n \ {..l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - delta l * real l} \ e * real N} > 1-e" using I.upper_density_delta[OF \e>0\ \e\1\] by blast define S where "S = {x \ space M. \N. card {n \ {..l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - delta l * real l} \ e * real N}" have [measurable]: "S \ sets M" unfolding S_def by auto have "emeasure M S > 1 - e" unfolding S_def using d(3) by simp define Og where "Og = (\n. {x \ space M. \l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) > - delta l * real l})" have [measurable]: "Og n \ sets M" for n unfolding Og_def by auto define Pg where "Pg = (\n. {x \ space M. \l\{1..n}. u n x - u (n - l) ((T^^l) x) > - delta l * real l})" have [measurable]: "Pg n \ sets M" for n unfolding Pg_def by auto define Bad where "Bad = (\i::nat. {x \ space M. \N\i. card {n \ {.. Pg n} \ (1-d) * real N})" have [measurable]: "Bad i \ sets M" for i unfolding Bad_def by auto then have "range Bad \ sets M" by auto have "incseq Bad" unfolding Bad_def incseq_def by auto have inc: "{x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ (\i. Bad i)" proof fix x assume H: "x \ {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" then have "x \ space M" by simp define A where "A = {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l}" have "upper_asymptotic_density A < 1-d" using H unfolding A_def by simp then have "\i. \N\i. card (A \ {.. (1-d)* real N" using upper_asymptotic_densityD[of A "1-d"] by (metis (no_types, lifting) eventually_sequentially less_imp_le) then obtain i where "card (A \ {.. (1-d)* real N" if "N\i" for N by blast moreover have "A \ {.. (\l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l)}" for N unfolding A_def by auto ultimately have "x \ Bad i" unfolding Bad_def Pg_def using \x \ space M\ by auto then show "x \ (\i. Bad i)" by blast qed have "emeasure M (Og n) \ emeasure M (Pg n)" for n proof - have *: "(T^^n)-`(Og n) \ space M \ Pg n" for n proof fix x assume x: "x \ (T^^n)-`(Og n) \ space M" define y where "y = (T^^n) x" then have "y \ Og n" using x by auto have "y \ space M" using sets.sets_into_space[OF \Og n \ sets M\] \y \ Og n\ by auto have "x = (Tinv^^n) y" unfolding y_def Tinv_def using inv_fn_o_fn_is_id[OF bij, of n] by (metis comp_apply) { fix l assume "l \ {1..n}" have "(T^^l) x = (T^^l) ((Tinv^^l) ((Tinv^^(n-l))y))" apply (subst \x = (Tinv^^n) y\) using funpow_add[of l "n-l" Tinv] \l \ {1..n}\ by fastforce then have *: "(T^^l) x = (Tinv^^(n-l)) y" unfolding Tinv_def using fn_o_inv_fn_is_id[OF bij] by (metis comp_apply) then have "u n x - u (n-l) ((T^^l) x) = u n ((Tinv^^n) y) - u (n-l) ((Tinv^^(n-l)) y)" using \x = (Tinv^^n) y\ by auto also have "... > - delta l * real l" using \y \ Og n\ \l \ {1..n}\ unfolding Og_def by auto finally have "u n x - u (n-l) ((T^^l) x) > - delta l * real l" by simp } then show "x \ Pg n" unfolding Pg_def using x by auto qed have "emeasure M (Og n) = emeasure M ((T^^n)-`(Og n) \ space M)" using T_vrestr_same_emeasure(2) unfolding vimage_restr_def by auto also have "... \ emeasure M (Pg n)" apply (rule emeasure_mono) using * by auto finally show ?thesis by simp qed { fix N::nat assume "N \ 1" have I: "card {n\{.. Og n} \ (1-e) * real N" if "x \ S" for x proof - have "x \ space M" using \x \ S\ sets.sets_into_space[OF \S \ sets M\] by auto have a: "real (card {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))}) \ e * real N" using \x \ S\ unfolding S_def by auto have *: "{n. n < N} = {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))} \ {n. n < N \ x \ Og n}" unfolding Og_def using \x \ space M\ by (auto, meson atLeastAtMost_iff linorder_not_le) have "N = card {n. n < N}" by auto also have "... = card {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))} + card {n. n < N \ x \ Og n}" apply (subst *, rule card_Un_disjoint) unfolding Og_def by auto ultimately have "real N \ e * real N + card {n. n < N \ x \ Og n}" using a by auto then show ?thesis by (auto simp add: algebra_simps) qed define m where "m = measure M (Bad N)" have "m \ 0" "1-m \ 0" unfolding m_def by auto have *: "1-e \ emeasure M S" using \emeasure M S > 1 - e\ by auto have "ennreal((1-e) * ((1-e) * real N)) = ennreal(1-e) * ennreal((1-e) * real N)" apply (rule ennreal_mult) using \e \ 1\ by auto also have "... \ emeasure M S * ennreal((1-e) * real N)" using mult_right_mono[OF *] by simp also have "... = (\\<^sup>+ x\S. ((1-e) * real N) \M)" by (metis \S \ events\ mult.commute nn_integral_cmult_indicator) also have "... \ (\\<^sup>+x \ S. ennreal(card {n\{.. Og n}) \M)" apply (rule nn_integral_mono) using I unfolding indicator_def by (simp) also have "... \ (\\<^sup>+x \ space M. ennreal(card {n\{.. Og n}) \M)" by (rule nn_set_integral_set_mono, simp only: sets.sets_into_space[OF \S \ sets M\]) also have "... = (\\<^sup>+x. ennreal(card {n\{.. Og n}) \M)" by (rule nn_set_integral_space) also have "... = (\\<^sup>+x. ennreal (\n\{..M)" apply (rule nn_integral_cong) using sum_indicator_eq_card2[of "{..\<^sup>+x. (\n\{..M)" apply (rule nn_integral_cong, auto, simp only: sum_ennreal[symmetric]) by (metis ennreal_0 ennreal_eq_1 indicator_eq_1_iff indicator_simps(2) real_of_nat_indicator) also have "... = (\n \{..\<^sup>+x. (indicator (Og n) x) \M))" by (rule nn_integral_sum, simp) also have "... = (\n \{.. (\n \{..\n. emeasure M (Og n) \ emeasure M (Pg n)\ by simp also have "... = (\n \{..\<^sup>+x. (indicator (Pg n) x) \M))" by simp also have "... = (\\<^sup>+x. (\n\{..M)" by (rule nn_integral_sum[symmetric], simp) also have "... = (\\<^sup>+x. ennreal (\n\{..M)" apply (rule nn_integral_cong, auto, simp only: sum_ennreal[symmetric]) by (metis ennreal_0 ennreal_eq_1 indicator_eq_1_iff indicator_simps(2) real_of_nat_indicator) also have "... = (\\<^sup>+x. ennreal(card {n\{.. Pg n}) \M)" apply (rule nn_integral_cong) using sum_indicator_eq_card2[of "{..\<^sup>+x \ space M. ennreal(card {n\{.. Pg n}) \M)" by (rule nn_set_integral_space[symmetric]) also have "... = (\\<^sup>+x \ Bad N \ (space M - Bad N). ennreal(card {n\{.. Pg n}) \M)" apply (rule nn_integral_cong) unfolding indicator_def by auto also have "... = (\\<^sup>+x \ Bad N. ennreal(card {n\{.. Pg n}) \M) + (\\<^sup>+x \ space M - Bad N. ennreal(card {n\{.. Pg n}) \M)" by (rule nn_integral_disjoint_pair, auto) also have "... \ (\\<^sup>+x \ Bad N. ennreal((1-d) * real N) \M) + (\\<^sup>+x \ space M - Bad N. ennreal(real N) \M)" apply (rule add_mono) apply (rule nn_integral_mono, simp add: Bad_def indicator_def, auto) apply (rule nn_integral_mono, simp add: indicator_def, auto) using card_Collect_less_nat[of N] card_mono[of "{n. n < N}"] by (simp add: Collect_mono_iff) also have "... = ennreal((1-d) * real N) * emeasure M (Bad N) + ennreal(real N) * emeasure M (space M - Bad N)" by (simp add: nn_integral_cmult_indicator) also have "... = ennreal((1-d) * real N) * ennreal(m) + ennreal(real N) * ennreal(1-m)" unfolding m_def by (simp add: emeasure_eq_measure prob_compl) also have "... = ennreal((1-d) * real N * m + real N * (1-m))" using \m \ 0\ \1-m \ 0\ \d \ 1\ ennreal_plus ennreal_mult by auto finally have "ennreal((1-e) * ((1-e) * real N)) \ ennreal((1-d) * real N * m + real N * (1-m))" by simp moreover have "(1-d) * real N * m + real N * (1-m) \ 0" using \m \ 0\ \1-m \ 0\ \d \ 1\ by auto ultimately have "(1-e) * ((1-e) * real N) \ (1-d) * real N * m + real N * (1-m)" using ennreal_le_iff by auto then have "0 \ (e * 2 - d * m - e * e) * real N" by (auto simp add: algebra_simps) then have "0 \ e * 2 - d * m - e * e" using \N \ 1\ by (simp add: zero_le_mult_iff) also have "... \ e * 2 - d * m" using \e > 0\ by auto finally have "m \ e * 2 / d" using \d>0\ by (auto simp add: algebra_simps divide_simps) then have "m \ d" unfolding e_def using \d>0\ by (auto simp add: divide_simps) then have "emeasure M (Bad N) \ d" unfolding m_def by (simp add: emeasure_eq_measure ennreal_leI) } then have "emeasure M (\i. Bad i) \ d" using LIMSEQ_le_const2[OF Lim_emeasure_incseq[OF \range Bad \ sets M\ \incseq Bad\]] by auto then have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ d" using emeasure_mono[OF inc, of M] by auto then have *: "measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ d" using emeasure_eq_measure \d>0\ by fastforce have "{x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} = space M - {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" by auto then have "measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} = measure M (space M - {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d})" by simp also have "... = measure M (space M) - measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" by (rule measure_Diff, auto) also have "... \ 1 - d" using * prob_space by linarith finally have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ 1 - d" using emeasure_eq_measure by auto then show ?thesis using d(1) d(2) by blast qed text \Now, we want to remove the invertibility assumption in the previous lemma. The idea is to go to the natural extension of the system, use the result there and project it back. However, if the system is not defined on a polish space, there is no reason why it should have a natural extension, so we have first to project the original system on a polish space on which the subcocycle is defined. This system is obtained by considering the joint distribution of the subcocycle and all its iterates (this is indeed a polish system, as a space of functions from $\mathbb{N}^2$ to $\mathbb{R}$).\ lemma upper_density_good_direction: assumes "d>(0::real)" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ ennreal(1-d)" proof - define U where "U = (\x. (\n. u n x))" define projJ where "projJ = (\x. (\n. U ((T^^n)x)))" define MJ where "MJ = (distr M borel (\x. (\n. U ((T^^n)x))))" define TJ::"(nat \ nat \ real) \ (nat \ nat \ real)" where "TJ = nat_left_shift" have *: "mpt_factor projJ MJ TJ" unfolding projJ_def MJ_def TJ_def apply (rule fmpt_factor_projection) unfolding U_def by (rule measurable_coordinatewise_then_product, simp) interpret J: polish_pmpt MJ TJ unfolding projJ_def polish_pmpt_def apply (auto) apply (rule pmpt_factor) using * apply simp unfolding polish_pmpt_axioms_def MJ_def by auto have [simp]: "projJ \ measure_preserving M MJ" using mpt_factorE(1)[OF *] by simp then have [measurable]: "projJ \ measurable M MJ" by (simp add: measure_preservingE(1)) text \We define a subcocycle $uJ$ in the projection corresponding to the original subcocycle $u$ above. (With the natural definition, it is only a subcocycle almost everywhere.) We check that it shares most properties of $u$.\ define uJ::"nat \ (nat \ nat \ real) \ real" where "uJ = (\n x. x 0 n)" have [measurable]: "uJ n \ borel_measurable borel" for n unfolding uJ_def by (metis measurable_product_coordinates measurable_product_then_coordinatewise) moreover have "measurable borel borel = measurable MJ borel" apply (rule measurable_cong_sets) unfolding MJ_def by auto ultimately have [measurable]: "uJ n \ borel_measurable MJ" for n by fast have uJ_proj: "u n x = uJ n (projJ x)" for n x unfolding uJ_def projJ_def U_def by auto have uJ_int: "integrable MJ (uJ n)" for n apply (rule measure_preserving_preserves_integral'(1)[OF \projJ \ measure_preserving M MJ\]) apply (subst uJ_proj[of n, symmetric]) using int_u[of n] by auto have uJ_int2: "(\x. uJ n x \MJ) = (\x. u n x \M)" for n unfolding uJ_proj apply (rule measure_preserving_preserves_integral'(2)[OF \projJ \ measure_preserving M MJ\]) apply (subst uJ_proj[of n, symmetric]) using int_u[of n] by auto have uJ_AE: "AE x in MJ. uJ (n+m) x \ uJ n x + uJ m ((TJ^^n) x)" for m n proof - have "AE x in M. uJ (n+m) (projJ x) \ uJ n (projJ x) + uJ m (projJ ((T^^n) x))" unfolding uJ_proj[symmetric] using subcocycle_ineq[OF subu] by auto moreover have "AE x in M. projJ ((T^^n) x) = (TJ^^n) (projJ x)" using qmpt_factor_iterates[OF mpt_factor_is_qmpt_factor[OF *]] by auto ultimately have *: "AE x in M. uJ (n+m) (projJ x) \ uJ n (projJ x) + uJ m ((TJ^^n) (projJ x))" by auto show ?thesis apply (rule quasi_measure_preserving_AE'[OF measure_preserving_is_quasi_measure_preserving[OF \projJ \ measure_preserving M MJ\], OF *]) by auto qed have uJ_0: "AE x in MJ. (\n. uJ n x / n) \ 0" proof - have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic(1)[OF subu subu_fin]) moreover have "AE x in M. subcocycle_lim u x = 0" using subu_0 by simp ultimately have *: "AE x in M. (\n. uJ n (projJ x) / n) \ 0" unfolding uJ_proj by auto show ?thesis apply (rule quasi_measure_preserving_AE'[OF measure_preserving_is_quasi_measure_preserving[OF \projJ \ measure_preserving M MJ\], OF *]) by auto qed text \Then, we go to the natural extension of $TJ$, to have an invertible system.\ define MI where "MI = J.natural_extension_measure" define TI where "TI = J.natural_extension_map" define projI where "projI = J.natural_extension_proj" interpret I: pmpt MI TI unfolding MI_def TI_def by (rule J.natural_extension(1)) have "I.mpt_factor projI MJ TJ" unfolding projI_def using I.mpt_factorE(1) J.natural_extension(3) MI_def TI_def by auto then have [simp]: "projI \ measure_preserving MI MJ" using I.mpt_factorE(1) by auto then have [measurable]: "projI \ measurable MI MJ" by (simp add: measure_preservingE(1)) have "I.invertible_qmpt" using J.natural_extension(2) MI_def TI_def by auto text \We define a natural subcocycle $uI$ there, and check its properties.\ define uI where uI_proj: "uI = (\n x. uJ n (projI x))" have [measurable]: "uI n \ borel_measurable MI" for n unfolding uI_proj by auto have uI_int: "integrable MI (uI n)" for n unfolding uI_proj by (rule measure_preserving_preserves_integral(1)[OF \projI \ measure_preserving MI MJ\ uJ_int]) have "(\x. uJ n x \MJ) = (\x. uI n x \MI)" for n unfolding uI_proj by (rule measure_preserving_preserves_integral(2)[OF \projI \ measure_preserving MI MJ\ uJ_int]) then have uI_int2: "(\x. uI n x \MI) = (\x. u n x \M)" for n using uJ_int2 by simp have uI_AE: "AE x in MI. uI (n+m) x \ uI n x + uI m (((TI)^^n) x)" for m n proof - have "AE x in MI. uJ (n+m) (projI x) \ uJ n (projI x) + uJ m (((TJ)^^n) (projI x))" apply (rule quasi_measure_preserving_AE[OF measure_preserving_is_quasi_measure_preserving[OF \projI \ measure_preserving MI MJ\]]) using uJ_AE by auto moreover have "AE x in MI. ((TJ)^^n) (projI x) = projI (((TI)^^n) x)" using I.qmpt_factor_iterates[OF I.mpt_factor_is_qmpt_factor[OF \I.mpt_factor projI MJ TJ\]] by auto ultimately show ?thesis unfolding uI_proj by auto qed have uI_0: "AE x in MI. (\n. uI n x / n) \ 0" unfolding uI_proj apply (rule quasi_measure_preserving_AE[OF measure_preserving_is_quasi_measure_preserving[OF \projI \ measure_preserving MI MJ\]]) using uJ_0 by simp text \As $uI$ is only a subcocycle almost everywhere, we correct it to get a genuine subcocycle, to which we will apply Lemma \verb+upper_density_good_direction_invertible+.\ obtain vI where H: "I.subcocycle vI" "AE x in MI. \n. vI n x = uI n x" using I.subcocycle_AE[OF uI_AE uI_int] by blast have [measurable]: "\n. vI n \ borel_measurable MI" "\n. integrable MI (vI n)" using I.subcocycle_integrable[OF H(1)] by auto have "(\x. vI n x \MI) = (\x. uI n x \MI)" for n apply (rule integral_cong_AE) using H(2) by auto then have "(\x. vI n x \MI) = (\x. u n x \M)" for n using uI_int2 by simp then have "I.subcocycle_avg_ereal vI = subcocycle_avg_ereal u" unfolding I.subcocycle_avg_ereal_def subcocycle_avg_ereal_def by auto then have vI_fin: "I.subcocycle_avg_ereal vI > -\" using subu_fin by simp have "AE x in MI. (\n. vI n x / n) \ 0" using uI_0 H(2) by auto moreover have "AE x in MI. (\n. vI n x / n) \ I.subcocycle_lim vI x" by (rule I.kingman_theorem_nonergodic(1)[OF H(1) vI_fin]) ultimately have vI_0: "AE x in MI. I.subcocycle_lim vI x = 0" using LIMSEQ_unique by auto interpret GKK: Gouezel_Karlsson_Kingman MI TI vI apply standard using H(1) vI_fin vI_0 by auto obtain delta where delta: "\l. delta l > 0" "delta \ 0" "emeasure MI {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d } \ 1 - d" using GKK.upper_density_good_direction_invertible[OF \I.invertible_qmpt\ \d>0\ \d\1\] by blast text \Then, we need to go back to the original system, showing that the estimates for $TI$ carry over. First, we go to $TJ$.\ have BJ: "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d } \ 1 - d" proof - have *: "AE x in MI. uJ n (projI x) = vI n x" for n using uI_proj H(2) by auto have **: "AE x in MI. \n. uJ n (projI x) = vI n x" by (subst AE_all_countable, auto intro: *) have "AE x in MI. \m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)" by (rule I.T_AE_iterates[OF **]) then have "AE x in MI. (\m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)) \ (\n. projI ((TI^^n) x) = (TJ^^n) (projI x))" using I.qmpt_factor_iterates[OF I.mpt_factor_is_qmpt_factor[OF \I.mpt_factor projI MJ TJ\]] by auto then obtain ZI where ZI: "\x. x \ space MI - ZI \ (\m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)) \ (\n. projI ((TI^^n) x) = (TJ^^n) (projI x))" "ZI \ null_sets MI" using AE_E3 by blast have *: "uJ n (projI x) - uJ (n - l) ((TJ ^^ l) (projI x)) = vI n x - vI (n - l) ((TI ^^ l) x)" if "x \ space MI - ZI" for x n l proof - have "(TI^^0) x = x" "(TJ^^0) (projI x) = (projI x)" by auto then show ?thesis using ZI(1)[OF that] by metis qed have "projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI - ZI = {x \ space MI - ZI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n (projI x) - uJ (n - l) ((TJ ^^ l) (projI x))} \ 1 - d}" by (auto simp add: measurable_space[OF \projI \ measurable MI MJ\]) also have "... = {x \ space MI - ZI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d}" using * by auto also have "... = {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI" by auto finally have *: "projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI - ZI = {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI" by simp have "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} = emeasure MI (projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI)" by (rule measure_preservingE(2)[symmetric], auto) also have "... = emeasure MI ((projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI) - ZI)" by (rule emeasure_Diff_null_set[OF \ZI \ null_sets MI\, symmetric], measurable) also have "... = emeasure MI ({x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI)" using * by simp also have "... = emeasure MI {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d}" by (rule emeasure_Diff_null_set[OF \ZI \ null_sets MI\], measurable) also have "... \ 1-d" using delta(3) by simp finally show ?thesis by simp qed text \Then, we go back to $T$ with virtually the same argument.\ have "emeasure M {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d } \ 1 - d" proof - obtain Z where Z: "\x. x \ space M - Z \ (\n. projJ ((T^^n) x) = (TJ^^n) (projJ x))" "Z \ null_sets M" using AE_E3[OF qmpt_factor_iterates[OF mpt_factor_is_qmpt_factor[OF \mpt_factor projJ MJ TJ\]]] by blast have *: "uJ n (projJ x) - uJ (n - l) ((TJ ^^ l) (projJ x)) = u n x - u (n - l) ((T^^ l) x)" if "x \ space M - Z" for x n l proof - have "(T^^0) x = x" "(TJ^^0) (projJ x) = (projJ x)" by auto then show ?thesis using Z(1)[OF that] uJ_proj by metis qed have "projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M - Z = {x \ space M - Z. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n (projJ x) - uJ (n - l) ((TJ ^^ l) (projJ x))} \ 1 - d}" by (auto simp add: measurable_space[OF \projJ \ measurable M MJ\]) also have "... = {x \ space M - Z. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d}" using * by auto also have "... = {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z" by auto finally have *: "projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M - Z = {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z" by simp have "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} = emeasure M (projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M)" by (rule measure_preservingE(2)[symmetric], auto) also have "... = emeasure M ((projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M) - Z)" by (rule emeasure_Diff_null_set[OF \Z \ null_sets M\, symmetric], measurable) also have "... = emeasure M ({x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z)" using * by simp also have "... = emeasure M {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d}" by (rule emeasure_Diff_null_set[OF \Z \ null_sets M\], measurable) finally show ?thesis using BJ by simp qed then show ?thesis using delta(1) delta(2) by auto qed text \From the quantitative lemma above, we deduce the qualitative statement we are after, still in the setting of the locale.\ lemma infinite_AE: shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l})" proof - have "\deltaf::real \ nat \ real. \d. ((d > 0 \ d \ 1) \ ((\l. deltaf d l > 0) \ (deltaf d \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d} \ ennreal(1-d)))" apply (subst choice_iff'[symmetric]) using upper_density_good_direction by auto then obtain deltaf::"real \ nat \ real" where H: "\d. d > 0 \ d \1 \ (\l. deltaf d l > 0) \ (deltaf d \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d} \ ennreal(1-d)" by blast define U where "U = (\d. {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d})" have [measurable]: "U d \ sets M" for d unfolding U_def by auto have *: "emeasure M (U d) \ 1 - d" if "d>0 \ d\ 1" for d unfolding U_def using H that by auto define V where "V = (\n::nat. U (1/(n+2)))" have [measurable]: "V \ sets M" unfolding V_def by auto have a: "emeasure M V \ 1 - 1 / (n + 2)" for n::nat proof - have "1 - 1 / (n + 2) = 1 - 1 / (real n + 2)" by auto also have "... \ emeasure M (U (1/(real n+2)))" using *[of "1 / (real n + 2)"] by auto also have "... \ emeasure M V" apply (rule Measure_Space.emeasure_mono) unfolding V_def by auto finally show ?thesis by simp qed have b: "(\n::nat. 1 - 1 / (n + 2)) \ ennreal(1 - 0)" by (intro tendsto_intros LIMSEQ_ignore_initial_segment) have "emeasure M V \ 1 - 0" apply (rule Lim_bounded) using a b by auto then have "emeasure M V = 1" by (simp add: emeasure_ge_1_iff) then have "AE x in M. x \ V" by (simp add: emeasure_eq_measure prob_eq_1) moreover { fix x assume "x \ V" then obtain n::nat where "x \ U (1/(real n+2))" unfolding V_def by blast define d where "d = 1/(real n + 2)" have "0 < d" "d\1" unfolding d_def by auto have "0 < 1-d" unfolding d_def by auto also have "... \ upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l}" using \x \ U (1/(real n+2))\ unfolding U_def d_def by auto finally have "infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l}" using upper_asymptotic_density_finite by force then have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l})" using H \0 < d\ \d\1\ by auto } ultimately show ?thesis by auto qed end text \Finally, we obtain the full statement, by reducing to the previous situation where the asymptotic average vanishes.\ theorem (in pmpt) Gouezel_Karlsson_Kingman: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" proof - have [measurable]: "integrable M (u n)" "u n \ borel_measurable M" for n using subcocycle_integrable[OF assms(1)] by auto define v where "v = birkhoff_sum (\x. -subcocycle_lim u x)" have int [measurable]: "integrable M (\x. -subcocycle_lim u x)" using kingman_theorem_nonergodic(2)[OF assms] by auto have "subcocycle v" unfolding v_def apply (rule subcocycle_birkhoff) using assms \integrable M (\x. -subcocycle_lim u x)\ unfolding subcocycle_def by auto have "subcocycle_avg_ereal v > - \" unfolding v_def using subcocycle_avg_ereal_birkhoff[OF int] kingman_theorem_nonergodic(2)[OF assms] by auto have "AE x in M. subcocycle_lim v x = real_cond_exp M Invariants (\x. -subcocycle_lim u x) x" unfolding v_def by (rule subcocycle_lim_birkhoff[OF int]) moreover have "AE x in M. real_cond_exp M Invariants (\x. - subcocycle_lim u x) x = - subcocycle_lim u x" by (rule real_cond_exp_F_meas[OF int], auto) ultimately have AEv: "AE x in M. subcocycle_lim v x = - subcocycle_lim u x" by auto define w where "w = (\n x. u n x + v n x)" have a: "subcocycle w" unfolding w_def by (rule subcocycle_add[OF assms(1) \subcocycle v\]) have b: "subcocycle_avg_ereal w > -\" unfolding w_def by (rule subcocycle_avg_add(1)[OF assms(1) \subcocycle v\ assms(2) \subcocycle_avg_ereal v > - \\]) have "AE x in M. subcocycle_lim w x = subcocycle_lim u x + subcocycle_lim v x" unfolding w_def by (rule subcocycle_lim_add[OF assms(1) \subcocycle v\ assms(2) \subcocycle_avg_ereal v > - \\]) then have c: "AE x in M. subcocycle_lim w x = 0" using AEv by auto interpret Gouezel_Karlsson_Kingman M T w apply standard using a b c by auto have "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. w n x - w (n-l) ((T^^l) x) > - delta l * l})" using infinite_AE by auto moreover { fix x assume H: "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. w n x - w (n-l) ((T^^l) x) > - delta l * l})" "x \ space M" have *: "v n x = - n * subcocycle_lim u x" for n unfolding v_def using birkhoff_sum_of_invariants[OF _ \x \ space M\] by auto have **: "v n ((T^^l) x) = - n * subcocycle_lim u x" for n l proof - have "v n ((T^^l) x) = - n * subcocycle_lim u ((T^^l) x)" unfolding v_def using birkhoff_sum_of_invariants[OF _ T_spaceM_stable(2)[OF \x \ space M\]] by auto also have "... = - n * subcocycle_lim u x" using Invariants_func_is_invariant_n[OF subcocycle_lim_meas_Inv(2) \x \ space M\] by auto finally show ?thesis by simp qed have "w n x - w (n-l) ((T^^l) x) = u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x" if "l \ {1..n}" for n l unfolding w_def using *[of n] **[of "n-l" l] that apply (auto simp add: algebra_simps) by (metis comm_semiring_class.distrib diff_add_inverse nat_le_iff_add of_nat_add) then have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" using H(1) by auto } ultimately show ?thesis by auto qed text \The previous theorem only contains a lower bound. The corresponding upper bound follows readily from Kingman's theorem. The next statement combines both upper and lower bounds.\ theorem (in pmpt) Gouezel_Karlsson_Kingman': assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l})" proof - { fix x assume x: "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" "(\l. u l x/l) \ subcocycle_lim u x" then obtain alpha::"nat \ real" where a: "\l. alpha l > 0" "alpha \ 0" "infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l}" by auto define delta::"nat \ real" where "delta = (\l. alpha l + norm(u l x / l - subcocycle_lim u x))" { fix n assume *: "\l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l" have H: "x > -a \ x < a \ abs x < a" for a x::real by simp have "abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l" if "l\{1..n}" for l proof (rule H) have "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x \ u l x - l * subcocycle_lim u x" using assms(1) subcocycle_ineq[OF assms(1), of l "n-l" x] that by auto also have "... \ l * norm(u l x/l - subcocycle_lim u x)" using that by (auto simp add: algebra_simps divide_simps) also have "... < delta l * l" unfolding delta_def using a(1)[of l] that by auto finally show "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x < delta l * l" by simp have "- (delta l * l) \ -alpha l * l" unfolding delta_def by (auto simp add: algebra_simps) also have "... < u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x" using * that by auto finally show "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > -(delta l * l)" by simp qed then have "\l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l" by auto } then have "{n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l} \ {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l}" by auto then have "infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l}" using a(3) finite_subset by blast moreover have "delta \ 0 + 0" unfolding delta_def using x(2) by (intro tendsto_intros a(2) tendsto_norm_zero LIM_zero) moreover have "delta l > 0" for l unfolding delta_def using a(1)[of l] by auto ultimately have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l})" by auto } then show ?thesis using Gouezel_Karlsson_Kingman[OF assms] kingman_theorem_nonergodic(1)[OF assms] by auto qed end (*of Gouezel_Karlsson.thy*) diff --git a/thys/Ergodic_Theory/Invariants.thy b/thys/Ergodic_Theory/Invariants.thy --- a/thys/Ergodic_Theory/Invariants.thy +++ b/thys/Ergodic_Theory/Invariants.thy @@ -1,1795 +1,1795 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \The invariant sigma-algebra, Birkhoff theorem\ theory Invariants imports Recurrence "HOL-Probability.Conditional_Expectation" begin subsection \The sigma-algebra of invariant subsets\ text \The invariant sigma-algebra of a qmpt is made of those sets that are invariant under the dynamics. When the transformation is ergodic, it is made of sets of zero or full measure. In general, the Birkhoff theorem is expressed in terms of the conditional expectation of an integrable function with respect to the invariant sigma-algebra.\ context qmpt begin text \We define the invariant sigma-algebra, as the sigma algebra of sets which are invariant under the dynamics, i.e., they coincide with their preimage under $T$.\ definition Invariants where "Invariants = sigma (space M) {A \ sets M. T-`A \ space M = A}" text \For this definition to make sense, we need to check that it really defines a sigma-algebra: otherwise, the \verb+sigma+ operation would make garbage out of it. This is the content of the next lemma.\ lemma Invariants_sets: "sets Invariants = {A \ sets M. T-`A \ space M = A}" proof - have "sigma_algebra (space M) {A \ sets M. T-`A \ space M = A}" proof - define I where "I = {A. T-`A \ space M = A}" have i: "\A. A \ I \ A \ space M" unfolding I_def by auto have "algebra (space M) I" proof (subst algebra_iff_Un) have a: "I \ Pow (space M)" using i by auto have b: "{} \ I" unfolding I_def by auto { fix A assume *: "A \ I" then have "T-`(space M - A) = T-`(space M) - T-`A" by auto then have "T-`(space M - A) \ space M = T-`(space M) \ (space M) - T-`A \ (space M)" by auto also have "... = space M - A" using * I_def by (simp add: inf_absorb2 subsetI) finally have "space M - A \ I" unfolding I_def by simp } then have c: "(\a\I. space M - a \ I)" by simp have d: "(\a\I. \b\I. a \ b \ I)" unfolding I_def by auto show "I \ Pow (space M) \ {} \ I \ (\a\I. space M - a \ I) \ (\a\I. \b\I. a \ b \ I)" using a b c d by blast qed moreover have "(\F. range F \ I \ (\i::nat. F i) \ I)" unfolding I_def by auto ultimately have "sigma_algebra (space M) I" using sigma_algebra_iff by auto moreover have "sigma_algebra (space M) (sets M)" using measure_space measure_space_def by auto ultimately have "sigma_algebra (space M) (I \ (sets M))" using sigma_algebra_intersection by auto moreover have "I \ sets M = {A \ sets M. T-`A \ space M = A}" unfolding I_def by auto ultimately show ?thesis by simp qed then show ?thesis unfolding Invariants_def using sigma_algebra.sets_measure_of_eq by blast qed text \By definition, the invariant subalgebra is a subalgebra of the original algebra. This is expressed in the following lemmas.\ lemma Invariants_is_subalg: "subalgebra M Invariants" unfolding subalgebra_def using Invariants_sets Invariants_def by (simp add: space_measure_of_conv) lemma Invariants_in_sets: assumes "A \ sets Invariants" shows "A \ sets M" using Invariants_sets assms by blast lemma Invariants_measurable_func: assumes "f \ measurable Invariants N" shows "f \ measurable M N" using Invariants_is_subalg measurable_from_subalg assms by auto text \We give several trivial characterizations of invariant sets or functions.\ lemma Invariants_vrestr: assumes "A \ sets Invariants" shows "T--`A = A" using assms Invariants_sets Invariants_in_sets[OF assms] by auto lemma Invariants_points: assumes "A \ sets Invariants" "x \ A" shows "T x \ A" using assms Invariants_sets by auto lemma Invariants_func_is_invariant: fixes f::"_ \ 'b::t2_space" assumes "f \ borel_measurable Invariants" "x \ space M" shows "f (T x) = f x" proof - have "{f x} \ sets borel" by simp then have "f-`({f x}) \ space M \ Invariants" using assms(1) by (metis (no_types, lifting) Invariants_def measurable_sets space_measure_of_conv) moreover have "x \ f-`({f x}) \ space M" using assms(2) by blast ultimately have "T x \ f-`({f x}) \ space M" by (rule Invariants_points) then show ?thesis by simp qed lemma Invariants_func_is_invariant_n: fixes f::"_ \ 'b::t2_space" assumes "f \ borel_measurable Invariants" "x \ space M" shows "f ((T^^n) x) = f x" by (induction n, auto simp add: assms Invariants_func_is_invariant) lemma Invariants_func_charac: assumes [measurable]: "f \ measurable M N" and "\x. x \ space M \ f(T x) = f x" shows "f \ measurable Invariants N" proof (rule measurableI) fix A assume "A \ sets N" have "space Invariants = space M" using Invariants_is_subalg subalgebra_def by force show "f -` A \ space Invariants \ sets Invariants" apply (subst Invariants_sets) apply (auto simp add: assms \A \ sets N\ \space Invariants = space M\) using \A \ sets N\ assms(1) measurable_sets by blast next fix x assume "x \ space Invariants" have "space Invariants = space M" using Invariants_is_subalg subalgebra_def by force then show "f x \ space N" using assms(1) \x \ space Invariants\ by (metis measurable_space) qed lemma birkhoff_sum_of_invariants: fixes f::" _ \ real" assumes "f \ borel_measurable Invariants" "x \ space M" shows "birkhoff_sum f n x = n * f x" unfolding birkhoff_sum_def using Invariants_func_is_invariant_n[OF assms] by auto text \There are two possible definitions of the invariant sigma-algebra, competing in the literature: one could also use the sets such that $T^{-1}A$ coincides with $A$ up to a measure $0$ set. It turns out that this is equivalent to being invariant (in our sense) up to a measure $0$ set. Therefore, for all interesting purposes, the two definitions would give the same results. For the proof, we start from an almost invariant set, and build a genuinely invariant set that coincides with it by adding or throwing away null parts. \ proposition Invariants_quasi_Invariants_sets: assumes [measurable]: "A \ sets M" shows "(\B \ sets Invariants. A \ B \ null_sets M) \ (T--`A \ A \ null_sets M)" proof assume "\B \ sets Invariants. A \ B \ null_sets M" then obtain B where "B \ sets Invariants" "A \ B \ null_sets M" by auto then have [measurable]: "B \ sets M" using Invariants_in_sets by simp have "B = T--` B" using Invariants_vrestr \B \ sets Invariants\ by simp then have "T--`A \ B = T--`(A \ B)" by simp moreover have "T--`(A \ B) \ null_sets M" by (rule T_quasi_preserves_null2(1)[OF \A \ B \ null_sets M\]) ultimately have "T--`A \ B \ null_sets M" by simp then show "T--`A \ A \ null_sets M" by (rule null_sym_diff_transitive) (auto simp add: \A \ B \ null_sets M\ Un_commute) next assume H: "T --` A \ A \ null_sets M" have [measurable]: "\n. (T^^n)--`A \ sets M" by simp { fix K assume [measurable]: "K \ sets M" and "T--`K \ K \ null_sets M" fix n::nat have "(T^^n)--`K \ K \ null_sets M" proof (induction n) case 0 have "(T^^0)--` K = K" using T_vrestr_0 by simp then show ?case using Diff_cancel sup.idem by (metis null_sets.empty_sets) next case (Suc n) have "T--`((T^^n)--`K \ K) \ null_sets M" using Suc.IH T_quasi_preserves_null(1)[of "((T^^n)--`K \ K)"] by simp then have *: "(T^^(Suc n))--`K \ T--`K \ null_sets M" using T_vrestr_composed(2)[OF \K \ sets M\] by simp then show ?case by (rule null_sym_diff_transitive, simp add: \T--`K \ K \ null_sets M\ \K \ sets M\, measurable) qed } note * = this define C where "C = (\n. (T^^n)--`A)" have [measurable]: "C \ sets M" unfolding C_def by simp have "C \ A \ (\n. (T^^n)--`A \ A)" unfolding C_def by auto moreover have "(\n. (T^^n)--`A \ A) \ null_sets M" using * null_sets_UN assms \T --` A \ A \ null_sets M\ by auto ultimately have CA: "C \ A \ null_sets M" by (meson \C \ sets M\ assms sets.Diff sets.Un null_sets_subset) then have "T--`(C \ A) \ null_sets M" by (rule T_quasi_preserves_null2(1)) then have "T--`C \ T--`A \ null_sets M" by simp then have "T--`C \ A \ null_sets M" by (rule null_sym_diff_transitive, auto simp add: H) then have TCC: "T--`C \ C \ null_sets M" apply (rule null_sym_diff_transitive) using CA by (auto simp add: Un_commute) have "C \ (\n\{1..}. (T^^n)--`A)" unfolding C_def by auto moreover have "T--`C = (\n\{1..}. (T^^n)--`A)" using T_vrestr_composed(2)[OF assms] by (simp add: C_def atLeast_Suc_greaterThan greaterThan_0) ultimately have "C \ T--`C" by blast then have "(T^^0)--`C \ (T^^1)--`C" using T_vrestr_0 by auto moreover have "(T^^1)--`C \ (\n\{1..}. (T^^n)--`C)" by auto ultimately have "(T^^0)--`C \ (\n\{1..}. (T^^n)--`C)" by auto then have "(T^^0)--`C \ (\n\{1..}. (T^^n)--`C) = (\n\{1..}. (T^^n)--`C)" by auto moreover have "(\n. (T^^n)--`C) = (T^^0)--`C \ (\n\{1..}. (T^^n)--`C)" by (rule union_insert_0) ultimately have C2: "(\n. (T^^n)--`C) = (\n\{1..}. (T^^n)--`C)" by simp define B where "B = (\n. (T^^n)--`C)" have [measurable]: "B \ sets M" unfolding B_def by simp have "B \ C \ (\n. (T^^n)--`C \ C)" unfolding B_def by auto moreover have "(\n. (T^^n)--`C \ C) \ null_sets M" using * null_sets_UN assms TCC by auto ultimately have "B \ C \ null_sets M" by (meson \B \ sets M\ \C \ sets M\ assms sets.Diff sets.Un null_sets_subset) then have "B \ A \ null_sets M" by (rule null_sym_diff_transitive, auto simp add: CA) then have a: "A \ B \ null_sets M" by (simp add: Un_commute) have "T--`B = (\n\{1..}. (T^^n)--`C)" using T_vrestr_composed(2)[OF \C \ sets M\] by (simp add: B_def atLeast_Suc_greaterThan greaterThan_0) then have "T--`B = B" unfolding B_def using C2 by simp then have "B \ sets Invariants" using Invariants_sets vimage_restr_def by auto then show "\B \ sets Invariants. A \ B \ null_sets M" using a by blast qed text \In a conservative setting, it is enough to be included in its image or its preimage to be almost invariant: otherwise, since the difference set has disjoint preimages, and is therefore null by conservativity.\ lemma (in conservative) preimage_included_then_almost_invariant: assumes [measurable]: "A \ sets M" and "T--`A \ A" shows "A \ (T--`A) \ null_sets M" proof - define B where "B = A - T--`A" then have [measurable]: "B \ sets M" by simp have "(T^^(Suc n))--`A \ (T^^n)--`A" for n using T_vrestr_composed(3)[OF assms(1)] vrestr_inclusion[OF assms(2)] by auto then have "disjoint_family (\n. (T^^n)--`A - (T^^(Suc n))--`A)" by (rule disjoint_family_Suc2[where ?A = "\n. (T^^n)--`A"]) moreover have "(T^^n)--`A - (T^^(Suc n))--`A = (T^^n)--`B" for n unfolding B_def Suc_eq_plus1 using T_vrestr_composed(3)[OF assms(1)] by auto ultimately have "disjoint_family (\n. (T^^n)--` B)" by simp then have "\n. n \ 0 \ ((T^^n)--`B) \ B = {}" unfolding disjoint_family_on_def by (metis UNIV_I T_vrestr_0(1)[OF \B \ sets M\]) then have "\n. n > 0 \ (T^^n)-`B \ B = {}" unfolding vimage_restr_def by (simp add: Int_assoc) then have "B \ null_sets M" using disjoint_then_null[OF \B \ sets M\] Int_commute by auto then show ?thesis unfolding B_def using assms(2) by (simp add: Diff_mono Un_absorb2) qed lemma (in conservative) preimage_includes_then_almost_invariant: assumes [measurable]: "A \ sets M" and "A \ T--`A" shows "A \ (T--`A) \ null_sets M" proof - define B where "B = T--`A - A" then have [measurable]: "B \ sets M" by simp have "\n. (T^^(Suc n))--`A \ (T^^n)--`A" using T_vrestr_composed(3)[OF assms(1)] vrestr_inclusion[OF assms(2)] by auto then have "disjoint_family (\n. (T^^(Suc n))--`A - (T^^n)--`A)" by (rule disjoint_family_Suc[where ?A = "\n. (T^^n)--`A"]) moreover have "\n. (T^^(Suc n))--`A - (T^^n)--`A = (T^^n)--`B" unfolding B_def Suc_eq_plus1 using T_vrestr_composed(3)[OF assms(1)] by auto ultimately have "disjoint_family (\n. (T^^n)--` B)" by simp then have "\n. n \ 0 \ ((T^^n)--`B) \ B = {}" unfolding disjoint_family_on_def by (metis UNIV_I T_vrestr_0(1)[OF \B \ sets M\]) then have "\n. n > 0 \ (T^^n)-`B \ B = {}" unfolding vimage_restr_def by (simp add: Int_assoc) then have "B \ null_sets M" using disjoint_then_null[OF \B \ sets M\] Int_commute by auto then show ?thesis unfolding B_def using assms(2) by (simp add: Diff_mono Un_absorb1) qed text \The above properties for sets are also true for functions: if $f$ and $f \circ T$ coincide almost everywhere, i.e., $f$ is almost invariant, then $f$ coincides almost everywhere with a true invariant function. The idea of the proof is straightforward: throw away the orbits on which $f$ is not really invariant (say this is the complement of the good set), and replace it by $0$ there. However, this does not work directly: the good set is not invariant, some points may have a non-constant value of $f$ on their orbit but reach the good set eventually. One can however define $g$ to be equal to the eventual value of $f$ along the orbit, if the orbit reaches the good set, and $0$ elsewhere.\ proposition Invariants_quasi_Invariants_functions: fixes f::"_ \ 'b::{second_countable_topology, t2_space}" assumes f_meas [measurable]: "f \ borel_measurable M" shows "(\g \ borel_measurable Invariants. AE x in M. f x = g x) \ (AE x in M. f(T x) = f x)" proof assume "\g\borel_measurable Invariants. AE x in M. f x = g x" then obtain g where g:"g\borel_measurable Invariants" "AE x in M. f x = g x" by blast then have [measurable]: "g \ borel_measurable M" using Invariants_measurable_func by auto define A where "A = {x \ space M. f x = g x}" have [measurable]: "A \ sets M" unfolding A_def by simp define B where "B = space M - A" have [measurable]: "B \ sets M" unfolding B_def by simp moreover have "AE x in M. x \ B" unfolding B_def A_def using g(2) by auto ultimately have "B \ null_sets M" using AE_iff_null_sets by blast then have "T--`B \ null_sets M" by (rule T_quasi_preserves_null2(1)) then have "B \ T--`B \ null_sets M" using \B \ null_sets M\ by auto then have "AE x in M. x \ (B \ T--`B)" using AE_iff_null_sets null_setsD2 by blast then have i: "AE x in M. x \ space M - (B \ T--`B)" by auto { fix x assume *: "x \ space M - (B \ T--`B)" then have "x \ A" unfolding B_def by blast then have "f x = g x" unfolding A_def by blast have "T x \ A" using * B_def by auto then have "f(T x) = g(T x)" unfolding A_def by blast moreover have "g(T x) = g x" apply (rule Invariants_func_is_invariant) using * by (auto simp add: assms \g\borel_measurable Invariants\) ultimately have "f(T x) = f x" using \f x = g x\ by simp } then show "AE x in M. f(T x) = f x" using i by auto next assume *: "AE x in M. f (T x) = f x" text \\verb+good_set+ is the set of points for which $f$ is constant on their orbit. Here, we define $g = f$. If a point ever enters \verb+good_set+, then we take $g$ to be the value of $f$ there. Otherwise, $g$ takes an arbitrary value, say $y_0$.\ define good_set where "good_set = {x \ space M. \n. f((T^^(Suc n)) x) = f((T^^n) x)}" define good_time where "good_time = (\x. Inf {n. (T^^n) x \ good_set})" have "AE x in M. x \ good_set" using T_AE_iterates[OF *] by (simp add: good_set_def) have [measurable]: "good_set \ sets M" unfolding good_set_def by auto obtain y0::'b where True by auto define g where "g = (\x. if (\n. (T^^n) x \ good_set) then f((T^^(good_time x)) x) else y0)" have [measurable]: "good_time \ measurable M (count_space UNIV)" unfolding good_time_def by measurable have [measurable]: "g \ borel_measurable M" unfolding g_def by measurable have "f x = g x" if "x \ good_set" for x proof - have a: "0 \ {n. (T^^n) x \ good_set}" using that by simp have "good_time x = 0" unfolding good_time_def apply (intro cInf_eq_non_empty) using a by blast+ moreover have "{n. (T^^n) x \ good_set} \ {}" using a by blast ultimately show "f x = g x" unfolding g_def by auto qed then have "AE x in M. f x = g x" using \AE x in M. x \ good_set\ by auto have *: "f((T^^(Suc 0)) x) = f((T^^0) x)" if "x \ good_set" for x using that unfolding good_set_def by blast have good_1: "T x \ good_set \ f(T x) = f x" if "x \ good_set" for x using *[OF that] that unfolding good_set_def apply (auto) unfolding T_Tn_T_compose by blast then have good_k: "\x. x \ good_set \ (T^^k) x \ good_set \ f((T^^k) x) = f x" for k by (induction k, auto) have "g(T x) = g x" if "x \ space M" for x proof (cases) assume *: "\n. (T^^n) (T x) \ good_set" define n where "n = Inf {n. (T^^n) (T x) \ good_set}" have "(T^^n)(T x) \ good_set" using * Inf_nat_def1 by (metis empty_iff mem_Collect_eq n_def) then have a: "(T^^(n+1)) x \ good_set" by (metis Suc_eq_plus1 comp_eq_dest_lhs funpow.simps(2) funpow_swap1) then have **: "\m. (T^^m) x \ good_set" by blast define m where "m = Inf {m. (T^^m) x \ good_set}" then have "(T^^m) x \ good_set" using ** Inf_nat_def1 by (metis empty_iff mem_Collect_eq) have "n+1 \ {m. (T^^m) x \ good_set}" using a by simp then have "m \ n+1" using m_def by (simp add: Inf_nat_def Least_le) then obtain k where "n+1 = m + k" using le_iff_add by blast have "g x = f((T^^m) x)" unfolding g_def good_time_def using ** m_def by simp also have "... = f((T^^k) ((T^^m) x))" using \(T^^m) x \ good_set\ good_k by simp also have "... = f((T^^(n+1))x)" using \n+1 = m + k\[symmetric] funpow_add by (metis add.commute comp_apply) also have "... = f((T^^n) (T x))" using funpow_Suc_right by (metis Suc_eq_plus1 comp_apply) also have "... = g(T x)" unfolding g_def good_time_def using * n_def by simp finally show "g(T x) = g x" by simp next assume *: "\(\n. (T^^n) (T x) \ good_set)" then have "g(T x) = y0" unfolding g_def by simp have **: "\(\n. (T^^(Suc n)) x \ good_set)" using funpow_Suc_right * by (metis comp_apply) have "T x \ good_set" using good_k * by blast then have "x \ good_set" using good_1 by auto then have "\(\n. (T^^n) x \ good_set)" using ** using good_1 by fastforce then have "g x = y0" unfolding g_def by simp then show "g(T x) = g x" using \g(T x) = y0\ by simp qed then have "g \ borel_measurable Invariants" by (rule Invariants_func_charac[OF \g \ borel_measurable M\]) then show "\g\borel_measurable Invariants. AE x in M. f x = g x" using \AE x in M. f x = g x\ by blast qed text \In a conservative setting, it suffices to have an almost everywhere inequality to get an almost everywhere equality, as the set where there is strict inequality has $0$ measure as its iterates are disjoint, by conservativity.\ proposition (in conservative) AE_decreasing_then_invariant: fixes f::"_ \ 'b::{linorder_topology, second_countable_topology}" assumes "AE x in M. f(T x) \ f x" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f(T x) = f x" proof - obtain D::"'b set" where D: "countable D" "(\x y. x < y \ (\d \ D. x \ d \ d < y))" using countable_separating_set_linorder2 by blast define A where "A = {x \ space M. f(T x) \ f x}" then have [measurable]: "A \ sets M" by simp define B where "B = {x \ space M. \n. f((T^^(n+1)) x) \ f((T^^n)x)}" then have [measurable]: "B \ sets M" by simp have "space M - A \ null_sets M" unfolding A_def using assms by (simp add: assms(1) AE_iff_null_sets) then have "(\n. (T^^n)--`(space M - A)) \ null_sets M" by (metis null_sets_UN T_quasi_preserves_null2(2)) moreover have "space M - B = (\n. (T^^n)--`(space M - A))" unfolding B_def A_def by auto ultimately have "space M - B \ null_sets M" by simp have *: "B = (\n. (T^^n)--`A)" unfolding B_def A_def by auto then have "T--`B = (\n. T--` (T^^n)--`A)" by auto also have "... = (\n. (T^^(n+1))--`A)" using T_vrestr_composed(2)[OF \A \ sets M\] by simp also have "... \ (\n. (T^^n)--`A)" by blast finally have B1: "B \ T--`B" using * by simp have "B \ A" using * T_vrestr_0[OF \A \ sets M\] by blast then have B2: "\x. x \ B \ f(T x) \ f x" unfolding A_def by auto define C where "C = (\t. {x \ B. f x \ t})" { fix t have "C t = B \ f-`{..t} \ space M" unfolding C_def using sets.sets_into_space[OF \B \ sets M\] by auto then have [measurable]: "C t \ sets M" using assms(2) by simp have "C t \ T--`(C t)" using B1 unfolding C_def vimage_restr_def apply auto using B2 order_trans by blast then have "T--`(C t) - C t \ null_sets M" by (metis Diff_mono Un_absorb1 preimage_includes_then_almost_invariant[OF \C t \ sets M\]) } then have "(\d\D. T--`(C d) - C d) \ null_sets M" using \countable D\ by (simp add: null_sets_UN') then have "(space M - B) \ (\d\D. T--`(C d) - C d) \ null_sets M" using \space M - B \ null_sets M\ by auto then have "AE x in M. x \ (space M - B) \ (\d\D. T--`(C d) - C d)" using AE_not_in by blast moreover { fix x assume x: "x \ space M" "x \ (space M - B) \ (\d\D. T--`(C d) - C d)" then have "x \ B" by simp then have "T x \ B" using B1 by auto have "f(T x) = f x" proof (rule ccontr) assume "f(T x) \ f x" then have "f(T x) < f x" using B2[OF \x \ B\] by simp then obtain d where d: "d \ D" "f(T x) \ d \ d < f x" using D by auto then have "T x \ C d" using \T x \ B\ unfolding C_def by simp then have "x \ T--`(C d)" using \x \ space M\ by simp then have "x \ C d" using x \d \ D\ by simp then have "f x \ d" unfolding C_def by simp then show False using d by auto qed } ultimately show ?thesis by auto qed proposition (in conservative) AE_increasing_then_invariant: fixes f::"_ \ 'b::{linorder_topology, second_countable_topology}" assumes "AE x in M. f(T x) \ f x" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f(T x) = f x" proof - obtain D::"'b set" where D: "countable D" "(\x y. x < y \ (\d \ D. x < d \ d \ y))" using countable_separating_set_linorder1 by blast define A where "A = {x \ space M. f(T x) \ f x}" then have [measurable]: "A \ sets M" by simp define B where "B = {x \ space M. \n. f((T^^(n+1)) x) \ f((T^^n)x)}" then have [measurable]: "B \ sets M" by simp have "space M - A \ null_sets M" unfolding A_def using assms by (simp add: assms(1) AE_iff_null_sets) then have "(\n. (T^^n)--`(space M - A)) \ null_sets M" by (metis null_sets_UN T_quasi_preserves_null2(2)) moreover have "space M - B = (\n. (T^^n)--`(space M - A))" unfolding B_def A_def by auto ultimately have "space M - B \ null_sets M" by simp have *: "B = (\n. (T^^n)--`A)" unfolding B_def A_def by auto then have "T--`B = (\n. T--` (T^^n)--`A)" by auto also have "... = (\n. (T^^(n+1))--`A)" using T_vrestr_composed(2)[OF \A \ sets M\] by simp also have "... \ (\n. (T^^n)--`A)" by blast finally have B1: "B \ T--`B" using * by simp have "B \ A" using * T_vrestr_0[OF \A \ sets M\] by blast then have B2: "\x. x \ B \ f(T x) \ f x" unfolding A_def by auto define C where "C = (\t. {x \ B. f x \ t})" { fix t have "C t = B \ f-`{t..} \ space M" unfolding C_def using sets.sets_into_space[OF \B \ sets M\] by auto then have [measurable]: "C t \ sets M" using assms(2) by simp have "C t \ T--`(C t)" using B1 unfolding C_def vimage_restr_def apply auto using B2 order_trans by blast then have "T--`(C t) - C t \ null_sets M" by (metis Diff_mono Un_absorb1 preimage_includes_then_almost_invariant[OF \C t \ sets M\]) } then have "(\d\D. T--`(C d) - C d) \ null_sets M" using \countable D\ by (simp add: null_sets_UN') then have "(space M - B) \ (\d\D. T--`(C d) - C d) \ null_sets M" using \space M - B \ null_sets M\ by auto then have "AE x in M. x \ (space M - B) \ (\d\D. T--`(C d) - C d)" using AE_not_in by blast moreover { fix x assume x: "x \ space M" "x \ (space M - B) \ (\d\D. T--`(C d) - C d)" then have "x \ B" by simp then have "T x \ B" using B1 by auto have "f(T x) = f x" proof (rule ccontr) assume "f(T x) \ f x" then have "f(T x) > f x" using B2[OF \x \ B\] by simp then obtain d where d: "d \ D" "f(T x) \ d \ d > f x" using D by auto then have "T x \ C d" using \T x \ B\ unfolding C_def by simp then have "x \ T--`(C d)" using \x \ space M\ by simp then have "x \ C d" using x \d \ D\ by simp then have "f x \ d" unfolding C_def by simp then show False using d by auto qed } ultimately show ?thesis by auto qed text \For an invertible map, the invariants of $T$ and $T^{-1}$ are the same.\ lemma Invariants_Tinv: assumes "invertible_qmpt" shows "qmpt.Invariants M Tinv = Invariants" proof - interpret I: qmpt M Tinv using Tinv_qmpt[OF assms] by auto have "(T -` A \ space M = A) \ (Tinv -` A \ space M = A)" if "A \ sets M" for A proof assume "T -` A \ space M = A" then show "Tinv -` A \ space M = A" using assms that unfolding Tinv_def invertible_qmpt_def apply auto apply (metis IntE UNIV_I bij_def imageE inv_f_f vimageE) apply (metis I.T_spaceM_stable(1) Int_iff Tinv_def bij_inv_eq_iff vimageI) done next assume "Tinv -` A \ space M = A" then show "T -` A \ space M = A" using assms that unfolding Tinv_def invertible_qmpt_def apply auto apply (metis IntE bij_def inv_f_f vimageE) apply (metis T_Tinv_of_set T_meas Tinv_def assms qmpt.vrestr_of_set qmpt_axioms vrestr_image(3)) done qed then have "{A \ sets M. Tinv -` A \ space M = A} = {A \ sets M. T -` A \ space M = A}" by blast then show ?thesis unfolding Invariants_def I.Invariants_def by auto qed end sublocale fmpt \ finite_measure_subalgebra M Invariants unfolding finite_measure_subalgebra_def finite_measure_subalgebra_axioms_def using Invariants_is_subalg by (simp add: finite_measureI) context fmpt begin text \The conditional expectation with respect to the invariant sigma-algebra is the same for $f$ or $f \circ T$, essentially by definition.\ lemma Invariants_of_foTn: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants (f o (T^^n)) x = real_cond_exp M Invariants f x" proof (rule real_cond_exp_charact) fix A assume [measurable]: "A \ sets Invariants" then have [measurable]: "A \ sets M" using Invariants_in_sets by blast then have ind_meas [measurable]: "((indicator A)::('a \ real)) \ borel_measurable Invariants" by auto have "set_lebesgue_integral M A (f \ (T^^n)) = (\x. indicator A x * f((T^^n) x) \M)" by (auto simp: comp_def set_lebesgue_integral_def) also have "... = (\x. indicator A ((T^^n) x) * f ((T^^n) x) \M)" by (rule Bochner_Integration.integral_cong, auto simp add: Invariants_func_is_invariant_n[OF ind_meas]) also have "... = (\x. indicator A x * f x \M)" apply (rule Tn_integral_preserving(2)) using integrable_mult_indicator[OF \A \ sets M\ assms] by auto also have "... = (\x. indicator A x * real_cond_exp M Invariants f x \M)" apply (rule real_cond_exp_intg(2)[symmetric]) using integrable_mult_indicator[OF \A \ sets M\ assms] by auto also have "... = set_lebesgue_integral M A (real_cond_exp M Invariants f)" by (auto simp: set_lebesgue_integral_def) finally show "set_lebesgue_integral M A (f \ (T^^n)) = set_lebesgue_integral M A (real_cond_exp M Invariants f)" by simp qed (auto simp add: assms real_cond_exp_int Tn_integral_preserving(1)[OF assms] comp_def) lemma Invariants_of_foT: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants f x = real_cond_exp M Invariants (f o T) x" using Invariants_of_foTn[OF assms, where ?n = 1] by auto lemma birkhoff_sum_Invariants: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants (birkhoff_sum f n) x = n * real_cond_exp M Invariants f x" proof - define F where "F = (\i. f o (T^^i))" have [measurable]: "\i. F i \ borel_measurable M" unfolding F_def by auto have *: "integrable M (F i)" for i unfolding F_def by (subst comp_def, rule Tn_integral_preserving(1)[OF assms, of i]) have "AE x in M. n * real_cond_exp M Invariants f x = (\i\{..i\{..i\{..i\{..x. \i\{..x. \i\{..Birkhoff theorem\ subsubsection \Almost everywhere version of Birkhoff theorem\ text \This paragraph is devoted to the proof of Birkhoff theorem, arguably the most fundamental result of ergodic theory. This theorem asserts that Birkhoff averages of an integrable function $f$ converge almost surely, to the conditional expectation of $f$ with respect to the invariant sigma algebra. This result implies for instance the strong law of large numbers (in probability theory). There are numerous proofs of this statement, but none is really easy. We follow the very efficient argument given in Katok-Hasselblatt. To help the reader, here is the same proof informally. The first part of the proof is formalized in \verb+birkhoff_lemma1+, the second one in \verb+birkhoff_lemma+, and the conclusion in \verb+birkhoff_theorem+. Start with an integrable function $g$. let $G_n(x) = \max_{k\leq n} S_k g(x)$. Then $\limsup S_n g/n \leq 0$ outside of $A$, the set where $G_n$ tends to infinity. Moreover, $G_{n+1} - G_n \circ T$ is bounded by $g$, and tends to $g$ on $A$. It follows from the dominated convergence theorem that $\int_A G_{n+1} - G_n \circ T \to \int_A g$. As $\int_A G_{n+1} - G_n \circ T = \int_A G_{n+1} - G_n \geq 0$, we obtain $\int_A g \geq 0$. Apply now this result to the function $g = f - E(f | I) - \epsilon$, where $\epsilon>0$ is fixed. Then $\int_A g = -\epsilon \mu(A)$, then have $\mu(A) = 0$. Thus, almost surely, $\limsup S_n g/n\leq 0$, i.e., $\limsup S_n f/n \leq E(f|I)+\epsilon$. Letting $\epsilon$ tend to $0$ gives $\limsup S_n f/n \leq E(f|I)$. Applying the same result to $-f$ gives $S_n f/n \to E(f|I)$. \ context fmpt begin lemma birkhoff_aux1: fixes f::"'a \ real" assumes [measurable]: "integrable M f" defines "A \ {x \ space M. limsup (\n. ereal(birkhoff_sum f n x)) = \}" shows "A \ sets Invariants" "(\x. f x * indicator A x \M) \ 0" proof - let ?bsf = "birkhoff_sum f" have [measurable]: "A \ sets M" unfolding A_def by simp have Ainv: "x \ A \ T x \ A" if "x \ space M" for x proof - have "ereal(?bsf (1 + n) x) = ereal(f x) + ereal(?bsf n (T x))" for n unfolding birkhoff_sum_cocycle birkhoff_sum_1 by simp moreover have "limsup (\n. ereal(f x) + ereal(?bsf n (T x))) = ereal(f x) + limsup(\n. ereal(?bsf n (T x)))" by (rule ereal_limsup_lim_add, auto) moreover have "limsup (\n. ereal(?bsf (n+1) x)) = limsup (\n. ereal(?bsf n x))" using limsup_shift by simp ultimately have "limsup (\n. ereal(birkhoff_sum f n x)) = ereal(f x) + limsup (\n. ereal(?bsf n (T x)))" by simp then have "limsup (\n. ereal(?bsf n x)) = \ \ limsup (\n. ereal(?bsf n (T x))) = \" by simp then show "x \ A \ T x \ A" using \x \ space M\ A_def by simp qed then show "A \ sets Invariants" using assms(2) Invariants_sets by auto define F where "F = (\n x. MAX k \{0..n}. ?bsf k x)" have [measurable]: "\n. F n \ borel_measurable M" unfolding F_def by measurable have intFn: "integrable M (F n)" for n unfolding F_def by (rule integrable_MAX, auto simp add: birkhoff_sum_integral(1)[OF assms(1)]) have Frec: "F (n+1) x - F n (T x) = max (-F n (T x)) (f x)" for n x proof - have "{0..n+1} = {0} \ {1..n+1}" by auto then have "(\k. ?bsf k x) ` {0..n+1} = (\k. ?bsf k x) ` {0} \ (\k. ?bsf k x) ` {1..n+1}" by blast then have *: "(\k. ?bsf k x) ` {0..n+1} = {0} \ (\k. ?bsf k x) ` {1..n+1}" using birkhoff_sum_1(1) by simp have b: "F (n+1) x = max (Max {0}) (MAX k \{1..n+1}. ?bsf k x)" by (subst F_def, subst *, rule Max.union, auto) have "(\k. ?bsf k x) ` {1..n+1} = (\k. ?bsf (1+k) x) ` {0..n}" using Suc_le_D by fastforce also have "... = (\k. f x + ?bsf k (T x)) ` {0..n}" by (subst birkhoff_sum_cocycle, subst birkhoff_sum_1(2), auto) finally have c: "F (n+1) x = max 0 (MAX k \{0..n}. ?bsf k (T x) + f x)" using b by (simp add: add_ac) have "{f x + birkhoff_sum f k (T x) |k. k \{0..n}} = (+) (f x) ` {birkhoff_sum f k (T x) |k. k \{0..n}}" by blast have "(MAX k \{0..n}. ?bsf k (T x) + f x) = (MAX k \{0..n}. ?bsf k (T x)) + f x" by (rule Max_add_commute) auto also have "... = F n (T x) + f x" unfolding F_def by simp finally have "(MAX k \{0..n}. ?bsf k (T x) + f x) = f x + F n (T x)" by simp then have "F (n+1) x = max 0 (f x + F n (T x))" using c by simp then show "F (n+1) x - F n (T x) = max (-F n (T x)) (f x)" by auto qed have a: "abs((F (n+1) x - F n (T x)) * indicator A x) \ abs(f x)" for n x proof - have "F (n+1) x -F n (T x) \ f x" using Frec by simp then have *: "F (n+1) x -F n (T x) \ - abs(f x)" by simp have "F n (T x) \ birkhoff_sum f 0 (T x)" unfolding F_def apply (rule Max_ge, simp) using atLeastAtMost_iff by blast then have "F n (T x) \ 0" using birkhoff_sum_1(1) by simp then have "-F n (T x) \ abs (f x)" by simp moreover have "f x \ abs(f x)" by simp ultimately have "F (n+1) x -F n (T x) \ abs(f x)" using Frec by simp then have "abs(F (n+1) x - F n (T x)) \ abs(f x)" using * by simp then show "abs((F (n+1) x - F n (T x)) * indicator A x) \ abs(f x)" unfolding indicator_def by auto qed have b: "(\n. (F (n+1) x - F n (T x)) * indicator A x) \ f x * indicator A x" for x proof (rule Lim_eventually, cases) assume "x \ A" then have "T x \ A" using Ainv A_def by auto then have "limsup (\n. ereal(birkhoff_sum f n (T x))) > ereal(-f x)" unfolding A_def by simp then obtain N where "ereal(?bsf N (T x)) > ereal(-f x)" using Limsup_obtain by blast then have *: "?bsf N (T x) > -f x" by simp { fix n assume "n\N" then have "?bsf N (T x) \ (\k. ?bsf k (T x)) ` {0..n}" by auto then have "F n (T x) \ ?bsf N (T x)" unfolding F_def by simp then have "F n (T x) \ -f x" using * by simp then have "max (-F n (T x)) (f x) = f x" by simp then have "F (n+1) x - F n (T x) = f x" using Frec by simp then have "(F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x" by simp } then show "eventually (\n. (F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x) sequentially" using eventually_sequentially by blast next assume "\(x \ A)" then have "indicator A x = (0::real)" by simp then show "eventually (\n. (F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x) sequentially" by auto qed have lim: "(\n. (\x. (F (n+1) x - F n (T x)) * indicator A x \M)) \ (\x. f x * indicator A x \M)" proof (rule integral_dominated_convergence[where ?w = "(\x. abs(f x))"]) show "integrable M (\x. \f x\)" using assms(1) by auto show "AE x in M. (\n. (F (n + 1) x - F n (T x)) * indicator A x) \ f x * indicator A x" using b by auto show "\n. AE x in M. norm ((F (n + 1) x - F n (T x)) * indicator A x) \ \f x\" using a by auto qed (simp_all) have "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) \ 0" for n proof - have "(\x. F n (T x) * indicator A x \M) = (\x. (\x. F n x * indicator A x) (T x) \M)" by (rule Bochner_Integration.integral_cong, auto simp add: Ainv indicator_def) also have "... = (\x. F n x * indicator A x \M)" by (rule T_integral_preserving, auto simp add: intFn integrable_real_mult_indicator) finally have i: "(\x. F n (T x) * indicator A x \M) = (\x. F n x * indicator A x \M)" by simp have "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) = (\x. F (n+1) x * indicator A x - F n (T x) * indicator A x \M)" by (simp add: mult.commute right_diff_distrib) also have "... = (\x. F (n+1) x * indicator A x \M) - (\x. F n (T x) * indicator A x \M)" by (rule Bochner_Integration.integral_diff, auto simp add: intFn integrable_real_mult_indicator T_meas T_integral_preserving(1)) also have "... = (\x. F (n+1) x * indicator A x \M) - (\x. F n x * indicator A x \M)" using i by simp also have "... = (\x. F (n+1) x * indicator A x - F n x * indicator A x \M)" by (rule Bochner_Integration.integral_diff[symmetric], auto simp add: intFn integrable_real_mult_indicator) also have "... = (\x. (F (n+1) x - F n x) * indicator A x \M)" by (simp add: mult.commute right_diff_distrib) finally have *: "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) = (\x. (F (n+1) x - F n x) * indicator A x \M)" by simp have "F n x \ F (n+1) x" for x unfolding F_def by (rule Max_mono, auto) then have "(F (n+1) x - F n x) * indicator A x \ 0" for x by simp then have "integral\<^sup>L M (\x. 0) \ integral\<^sup>L M (\x. (F (n+1) x - F n x) * indicator A x)" by (auto simp add: intFn integrable_real_mult_indicator intro: integral_mono) then have "(\x. (F (n+1) x - F n x) * indicator A x \M) \ 0" by simp then show "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) \ 0" using * by simp qed then show "(\x. f x * indicator A x \M) \ 0" using lim by (simp add: LIMSEQ_le_const) qed lemma birkhoff_aux2: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x" proof - { fix \ assume "\ > (0::real)" define g where "g = (\x. f x - real_cond_exp M Invariants f x - \)" then have intg: "integrable M g" using assms real_cond_exp_int(1) assms by auto define A where "A = {x \ space M. limsup (\n. ereal(birkhoff_sum g n x)) = \}" have Ag: "A \ sets Invariants" "(\x. g x * indicator A x \M) \ 0" unfolding A_def by (rule birkhoff_aux1[where ?f = g, OF intg])+ then have [measurable]: "A \ sets M" by (simp add: Invariants_in_sets) have eq: "(\x. indicator A x * real_cond_exp M Invariants f x \M) = (\x. indicator A x * f x \M)" proof (rule real_cond_exp_intg[where ?f = "\x. (indicator A x)::real" and ?g = f]) have "(\x. indicator A x * f x) = (\x. f x * indicator A x)" by auto then show "integrable M (\x. indicator A x * f x)" using integrable_real_mult_indicator[OF \A \ sets M\ assms] by simp show "indicator A \ borel_measurable Invariants" using \A \ sets Invariants\ by measurable qed (simp) have "0 \ (\x. g x * indicator A x \M)" using Ag by simp also have "... = (\x. f x * indicator A x - real_cond_exp M Invariants f x * indicator A x - \ * indicator A x \M)" unfolding g_def by (simp add: left_diff_distrib) also have "... = (\x. f x * indicator A x \M) - (\x. real_cond_exp M Invariants f x * indicator A x \M) - (\x. \ * indicator A x \M)" using assms real_cond_exp_int(1)[OF assms] integrable_real_mult_indicator[OF \A \ sets M\] by auto also have "... = - (\x. \ * indicator A x \M)" by (auto simp add: eq mult.commute) also have "... = - \ * measure M A" by auto finally have "0 \ - \ * measure M A" by simp then have "measure M A = 0" using \\ > 0\ by (simp add: measure_le_0_iff mult_le_0_iff) then have "A \ null_sets M" by (simp add: emeasure_eq_measure null_setsI) then have "AE x in M. x \ space M - A" by (metis (no_types, lifting) AE_cong Diff_iff AE_not_in) moreover { fix x assume "x \ space M - A" then have "limsup (\n. ereal(birkhoff_sum g n x)) < \" unfolding A_def by auto then obtain C where C: "\n. birkhoff_sum g n x \ C" using limsup_finite_then_bounded by presburger { fix n::nat assume "n > 0" have "birkhoff_sum g n x = birkhoff_sum f n x - birkhoff_sum (real_cond_exp M Invariants f) n x - birkhoff_sum (\x. \) n x" unfolding g_def using birkhoff_sum_add birkhoff_sum_diff by auto moreover have "birkhoff_sum (real_cond_exp M Invariants f) n x = n * real_cond_exp M Invariants f x" using birkhoff_sum_of_invariants using \x \ space M - A\ by auto moreover have "birkhoff_sum (\x. \) n x = n * \" unfolding birkhoff_sum_def by auto ultimately have "birkhoff_sum g n x = birkhoff_sum f n x - n * real_cond_exp M Invariants f x - n * \" by simp then have "birkhoff_sum f n x = birkhoff_sum g n x + n * real_cond_exp M Invariants f x + n * \" by simp then have "birkhoff_sum f n x / n = birkhoff_sum g n x / n + real_cond_exp M Invariants f x + \" - using \n > 0\ by (simp add: add_divide_eq_iff linordered_field_class.sign_simps(27)) + using \n > 0\ by (simp add: field_simps) then have "birkhoff_sum f n x / n \ C/n + real_cond_exp M Invariants f x + \" using C[of n] \n > 0\ by (simp add: divide_right_mono) then have "ereal(birkhoff_sum f n x / n) \ ereal(C/n + real_cond_exp M Invariants f x + \)" by simp } then have "eventually (\n. ereal(birkhoff_sum f n x / n) \ ereal(C/n + real_cond_exp M Invariants f x + \)) sequentially" by (simp add: eventually_at_top_dense) then have b: "limsup (\n. ereal(birkhoff_sum f n x / n)) \ limsup (\n. ereal(C/n + real_cond_exp M Invariants f x + \))" by (simp add: Limsup_mono) have "(\n. ereal(C*(1/real n) + real_cond_exp M Invariants f x + \)) \ ereal(C * 0 + real_cond_exp M Invariants f x + \)" by (intro tendsto_intros) then have "limsup (\n. ereal(C/real n + real_cond_exp M Invariants f x + \)) = real_cond_exp M Invariants f x + \" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by force then have "limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x + \" using b by simp } ultimately have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x + \" by auto then have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ ereal(real_cond_exp M Invariants f x) + \" by auto } then show ?thesis by (rule AE_upper_bound_inf_ereal) qed theorem birkhoff_theorem_AE_nonergodic: fixes f::"'a \ real" assumes "integrable M f" shows "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" proof - { fix x assume i: "limsup (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" and ii: "limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n)) \ real_cond_exp M Invariants (\x. -f x) x" and iii: "real_cond_exp M Invariants (\x. -f x) x = - real_cond_exp M Invariants f x" have "\n. birkhoff_sum (\x. -f x) n x = - birkhoff_sum f n x" using birkhoff_sum_cmult[where ?c = "-1" and ?f = f] by auto then have "\n. ereal(birkhoff_sum (\x. -f x) n x / n) = - ereal(birkhoff_sum f n x / n)" by auto moreover have "limsup (\n. - ereal(birkhoff_sum f n x / n)) = - liminf (\n. ereal(birkhoff_sum f n x /n))" by (rule ereal_Limsup_uminus) ultimately have "-liminf (\n. ereal(birkhoff_sum f n x /n)) = limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n))" by simp then have "-liminf (\n. ereal(birkhoff_sum f n x /n)) \ - real_cond_exp M Invariants f x" using ii iii by simp then have "liminf (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" by (simp add: ereal_uminus_le_reorder) then have "(\n. birkhoff_sum f n x /n) \ real_cond_exp M Invariants f x" using i by (simp add: limsup_le_liminf_real) } note * = this moreover have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" using birkhoff_aux2 assms by simp moreover have "AE x in M. limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n)) \ real_cond_exp M Invariants (\x. -f x) x" using birkhoff_aux2 assms by simp moreover have "AE x in M. real_cond_exp M Invariants (\x. -f x) x = - real_cond_exp M Invariants f x" using real_cond_exp_cmult[where ?c = "-1"] assms by force ultimately show ?thesis by auto qed text \If a function $f$ is integrable, then $E(f\circ T - f | I) = E(f\circ T | I) - E(f|I) = 0$. Hence, $S_n(f \circ T - f) / n$ converges almost everywhere to $0$, i.e., $f(T^n x)/n \to 0$. It is remarkable (and sometimes useful) that this holds under the weaker condition that $f\circ T - f$ is integrable (but not necessarily $f$), where this naive argument fails. The reason is that the Birkhoff sum of $f \circ T - f$ is $f\circ T^n - f$. If $n$ is such that $x$ and $T^n(x)$ belong to a set where $f$ is bounded, it follows that this Birkhoff sum is also bounded. Along such a sequence of times, $S_n(f\circ T - f)/n$ tends to $0$. By Poincare recurrence theorem, there are such times for almost every points. As it also converges to $E(f \circ T - f | I)$, it follows that this function is almost everywhere $0$. Then $f (T^n x)/n = S_n(f\circ T^n - f)/n - f/n$ tends almost surely to $E(f\circ T-f |I) = 0$. \ lemma limit_foTn_over_n: fixes f::"'a \ real" assumes [measurable]: "f \ borel_measurable M" and "integrable M (\x. f(T x) - f x)" shows "AE x in M. real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" "AE x in M. (\n. f((T^^n) x) / n) \ 0" proof - define E::"nat \ 'a set" where "E k = {x \ space M. \f x\ \ k}" for k have [measurable]: "E k \ sets M" for k unfolding E_def by auto have *: "(\k. E k) = space M" unfolding E_def by (auto simp add: real_arch_simple) define F::"nat \ 'a set" where "F k = recurrent_subset_infty (E k)" for k have [measurable]: "F k \ sets M" for k unfolding F_def by auto have **: "E k - F k \ null_sets M" for k unfolding F_def using Poincare_recurrence_thm by auto have "space M - (\k. F k) \ null_sets M" apply (rule null_sets_subset[of "(\k. E k - F k)"]) unfolding *[symmetric] using ** by auto with AE_not_in[OF this] have "AE x in M. x \ (\k. F k)" by auto moreover have "AE x in M. (\n. birkhoff_sum (\x. f(T x) - f x) n x / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" by (rule birkhoff_theorem_AE_nonergodic[OF assms(2)]) moreover have "real_cond_exp M Invariants (\x. f(T x) - f x) x = 0 \ (\n. f((T^^n) x) / n) \ 0" if H: "(\n. birkhoff_sum (\x. f(T x) - f x) n x / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" "x \ (\k. F k)" for x proof - have "f((T^^n) x) = birkhoff_sum (\x. f(T x) - f x) n x + f x" for n unfolding birkhoff_sum_def by (induction n, auto) then have "f((T^^n) x) / n = birkhoff_sum (\x. f(T x) - f x) n x / n + f x * (1/n)" for n by (auto simp add: divide_simps) moreover have "(\n. birkhoff_sum (\x. f(T x) - f x) n x / n + f x * (1/n)) \ real_cond_exp M Invariants (\x. f(T x) - f x) x + f x * 0" by (intro tendsto_intros H(1)) ultimately have lim: "(\n. f((T^^n) x) / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" by auto obtain k where "x \ F k" using H(2) by auto then have "infinite {n. (T^^n) x \ E k}" unfolding F_def recurrent_subset_infty_inf_returns by auto with infinite_enumerate[OF this] obtain r :: "nat \ nat" where r: "strict_mono r" "\n. r n \ {n. (T^^n) x \ E k}" by auto have A: "(\n. k * (1/r n)) \ real k * 0" apply (intro tendsto_intros) using LIMSEQ_subseq_LIMSEQ[OF lim_1_over_n \strict_mono r\] unfolding comp_def by auto have B: "\f((T^^(r n)) x) / r n\ \ k / (r n)" for n using r(2) unfolding E_def by (auto simp add: divide_simps) have "(\n. f((T^^(r n)) x) / r n) \ 0" apply (rule tendsto_rabs_zero_cancel, rule tendsto_sandwich[of "\n. 0" _ _ "\n. k * (1/r n)"]) using A B by auto moreover have "(\n. f((T^^(r n)) x) / r n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" using LIMSEQ_subseq_LIMSEQ[OF lim \strict_mono r\] unfolding comp_def by auto ultimately have *: "real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" using LIMSEQ_unique by auto then have "(\n. f((T^^n) x) / n) \ 0" using lim by auto then show ?thesis using * by auto qed ultimately show "AE x in M. real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" "AE x in M. (\n. f((T^^n) x) / n) \ 0" by auto qed text \We specialize the previous statement to the case where $f$ itself is integrable.\ lemma limit_foTn_over_n': fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. (\n. f((T^^n) x) / n) \ 0" by (rule limit_foTn_over_n, simp, rule Bochner_Integration.integrable_diff) (auto intro: assms T_integral_preserving(1)) text \It is often useful to show that a function is cohomologous to a nicer function, i.e., to prove that a given $f$ can be written as $f = g + u - u \circ T$ where $g$ is nicer than $f$. We show below that any integrable function is cohomologous to a function which is arbitrarily close to $E(f|I)$. This is an improved version of Lemma 2.1 in [Benoist-Quint, Annals of maths, 2011]. Note that the function $g$ to which $f$ is cohomologous is very nice (and, in particular, integrable), but the transfer function is only measurable in this argument. The fact that the control on conditional expectation is nevertheless preserved throughout the argument follows from Lemma~\verb+limit_foTn_over_n+ above.\ text \We start with the lemma (and the proof) of [BQ2011]. It shows that, if a function has a conditional expectation with respect to invariants which is positive, then it is cohomologous to a nonnegative function. The argument is the clever remark that $g = \max (0, \inf_n S_n f)$ and $u = \min (0, \inf_n S_n f)$ work (where these expressions are well defined as $S_n f$ tends to infinity thanks to our assumption).\ lemma cohomologous_approx_cond_exp_aux: fixes f::"'a \ real" assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x > 0" shows "\u g. u \ borel_measurable M \ (integrable M g) \ (AE x in M. g x \ 0 \ g x \ max 0 (f x)) \ (\x. f x = g x + u x - u (T x))" proof - define h::"'a \ real" where "h = (\x. (INF n\{1..}. birkhoff_sum f n x))" define u where "u = (\x. min (h x) 0)" define g where "g = (\x. f x - u x + u (T x))" have [measurable]: "h \ borel_measurable M" "u \ borel_measurable M" "g \ borel_measurable M" unfolding g_def h_def u_def by auto have "f x = g x + u x - u (T x)" for x unfolding g_def by auto { fix x assume H: "real_cond_exp M Invariants f x > 0" "(\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" have "eventually (\n. ereal(birkhoff_sum f n x / n) * ereal n = ereal(birkhoff_sum f n x)) sequentially" unfolding eventually_sequentially by (rule exI[of _ 1], auto) moreover have "(\n. ereal(birkhoff_sum f n x / n) * ereal n) \ ereal(real_cond_exp M Invariants f x) * \" apply (intro tendsto_intros) using H by auto ultimately have "(\n. ereal(birkhoff_sum f n x)) \ ereal(real_cond_exp M Invariants f x) * \" by (rule Lim_transform_eventually) then have "(\n. ereal(birkhoff_sum f n x)) \ \" using H by auto then have B: "\C. \n. C \ birkhoff_sum f n x" by (intro liminf_finite_then_bounded_below, simp add: liminf_PInfty) have "h x \ f x" unfolding h_def apply (rule cInf_lower) using B by force+ have "{birkhoff_sum f n (T x) |n. n \ {1..}} = {birkhoff_sum f (1+n) (x) - f x |n. n \ {1..}}" unfolding birkhoff_sum_cocycle by auto also have "... = {birkhoff_sum f n x - f x |n. n \ {2..}}" by (metis (no_types, hide_lams) Suc_1 Suc_eq_plus1_left Suc_le_D Suc_le_mono atLeast_iff) finally have *: "{birkhoff_sum f n (T x) |n. n \ {1..}} = (\t. t - (f x))`{birkhoff_sum f n x |n. n \ {2..}}" by auto have "h(T x) = Inf {birkhoff_sum f n (T x) |n. n \ {1..}}" unfolding h_def by (metis Setcompr_eq_image) also have "... = (\t\{birkhoff_sum f n x |n. n \ {2..}}. t - f x)" by (simp only: *) also have "... = (\t. t - (f x)) (Inf {birkhoff_sum f n x |n. n \ {2..}})" using B by (auto intro!: monoI bijI mono_bij_cInf [symmetric]) finally have I: "Inf {birkhoff_sum f n x |n. n \ {2..}} = f x + h (T x)" by auto have "max 0 (h x) + u x = h x" unfolding u_def by auto also have "... = Inf {birkhoff_sum f n x |n. n \ {1..}}" unfolding h_def by (metis Setcompr_eq_image) also have "... = Inf ({birkhoff_sum f n x |n. n \ {1}} \ {birkhoff_sum f n x |n. n \ {2..}})" by (auto intro!: arg_cong[of _ _ Inf], metis One_nat_def Suc_1 antisym birkhoff_sum_1(2) not_less_eq_eq, force) also have "Inf ({birkhoff_sum f n x |n. n \ {1}} \ {birkhoff_sum f n x |n. n \ {2..}}) = min (Inf {birkhoff_sum f n x |n. n \ {1}}) (Inf {birkhoff_sum f n x |n. n \ {2..}})" unfolding inf_min[symmetric] apply (intro cInf_union_distrib) using B by auto also have "... = min (f x) (f x + h (T x))" using I by auto also have "... = f x + u (T x)" unfolding u_def by auto finally have "max 0 (h x) = f x + u (T x) - u x" by auto then have "g x = max 0 (h x)" unfolding g_def by auto then have "g x \ 0 \ g x \ max 0 (f x)" using \h x \ f x\ by auto } then have *: "AE x in M. g x \ 0 \ g x \ max 0 (f x)" using assms(2) birkhoff_theorem_AE_nonergodic[OF assms(1)] by auto moreover have "integrable M g" apply (rule Bochner_Integration.integrable_bound[of _ f]) using * by (auto simp add: assms) ultimately have "u \ borel_measurable M \ integrable M g \ (AE x in M. 0 \ g x \ g x \ max 0 (f x)) \ (\x. f x = g x + u x - u (T x))" using \\x. f x = g x + u x - u (T x)\ \u \ borel_measurable M\ by auto then show ?thesis by blast qed text \To deduce the stronger version that $f$ is cohomologous to an arbitrarily good approximation of $E(f|I)$, we apply the previous lemma twice, to control successively the negative and the positive side. The sign control in the conclusion of the previous lemma implies that the second step does not spoil the first one.\ lemma cohomologous_approx_cond_exp: fixes f::"'a \ real" and B::"'a \ real" assumes [measurable]: "integrable M f" "B \ borel_measurable M" and "AE x in M. B x > 0" shows "\g u. u \ borel_measurable M \ integrable M g \ (\x. f x = g x + u x - u (T x)) \ (AE x in M. abs(g x - real_cond_exp M Invariants f x) \ B x)" proof - define C where "C = (\x. min (B x) 1)" have [measurable]: "integrable M C" apply (rule Bochner_Integration.integrable_bound[of _ "\_. (1::real)"], auto) unfolding C_def using assms(3) by auto have "C x \ B x" for x unfolding C_def by auto have "AE x in M. C x > 0" unfolding C_def using assms(3) by auto have AECI: "AE x in M. real_cond_exp M Invariants C x > 0" by (intro real_cond_exp_gr_c \integrable M C\ \AE x in M. C x > 0\) define f1 where "f1 = (\x. f x - real_cond_exp M Invariants f x)" have "integrable M f1" unfolding f1_def by (intro Bochner_Integration.integrable_diff \integrable M f\ real_cond_exp_int(1)) have "AE x in M. real_cond_exp M Invariants f1 x = real_cond_exp M Invariants f x - real_cond_exp M Invariants (real_cond_exp M Invariants f) x" unfolding f1_def apply (rule real_cond_exp_diff) by (intro Bochner_Integration.integrable_diff \integrable M f\ \integrable M C\ real_cond_exp_int(1))+ moreover have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (intro real_cond_exp_nested_subalg subalg \integrable M f\, auto) ultimately have AEf1: "AE x in M. real_cond_exp M Invariants f1 x = 0" by auto have A [measurable]: "integrable M (\x. f1 x + C x)" by (intro Bochner_Integration.integrable_add \integrable M f1\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. f1 x + C x) x = real_cond_exp M Invariants f1 x + real_cond_exp M Invariants C x" by (intro real_cond_exp_add \integrable M f1\ \integrable M C\) then have B: "AE x in M. real_cond_exp M Invariants (\x. f1 x + C x) x > 0" using AECI AEf1 by auto obtain u2 g2 where H2: "u2 \ borel_measurable M" "integrable M g2" "AE x in M. g2 x \ 0 \ g2 x \ max 0 (f1 x + C x)" "\x. f1 x + C x = g2 x + u2 x - u2 (T x)" using cohomologous_approx_cond_exp_aux[OF A B] by blast define f2 where "f2 = (\x. (g2 x - C x))" have *: "u2(T x) - u2 x = f2 x -f1 x" for x unfolding f2_def using H2(4)[of x] by auto have "AE x in M. f2 x \ - C x" using H2(3) unfolding f2_def by auto have "integrable M f2" unfolding f2_def by (intro Bochner_Integration.integrable_diff \integrable M g2\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. u2(T x) - u2 x) x = 0" proof (rule limit_foTn_over_n) show "integrable M (\x. u2(T x) - u2 x)" unfolding * by (intro Bochner_Integration.integrable_diff \integrable M f1\ \integrable M f2\) qed (simp add: \u2 \ borel_measurable M\) then have "AE x in M. real_cond_exp M Invariants (\x. f2 x - f1 x) x = 0" unfolding * by simp moreover have "AE x in M. real_cond_exp M Invariants (\x. f2 x - f1 x) x = real_cond_exp M Invariants f2 x - real_cond_exp M Invariants f1 x" by (intro real_cond_exp_diff \integrable M f2\ \integrable M f1\) ultimately have AEf2: "AE x in M. real_cond_exp M Invariants f2 x = 0" using AEf1 by auto have A [measurable]: "integrable M (\x. C x - f2 x)" by (intro Bochner_Integration.integrable_diff \integrable M f2\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. C x - f2 x) x = real_cond_exp M Invariants C x - real_cond_exp M Invariants f2 x" by (intro real_cond_exp_diff \integrable M f2\ \integrable M C\) then have B: "AE x in M. real_cond_exp M Invariants (\x. C x - f2 x) x > 0" using AECI AEf2 by auto obtain u3 g3 where H3: "u3 \ borel_measurable M" "integrable M g3" "AE x in M. g3 x \ 0 \ g3 x \ max 0 (C x - f2 x)" "\x. C x - f2 x = g3 x + u3 x - u3 (T x)" using cohomologous_approx_cond_exp_aux[OF A B] by blast define f3 where "f3 = (\x. C x - g3 x)" have "AE x in M. f3 x \ min (C x) (f2 x)" unfolding f3_def using H3(3) by auto then have "AE x in M. f3 x \ -C x" using \AE x in M. f2 x \ - C x\ \AE x in M. C x > 0\ by auto moreover have "AE x in M. f3 x \ C x" unfolding f3_def using H3(3) by auto ultimately have "AE x in M. abs(f3 x) \ C x" by auto then have *: "AE x in M. abs(f3 x) \ B x" using order_trans[OF _ \\x. C x \ B x\] by auto define g where "g = (\x. f3 x + real_cond_exp M Invariants f x)" define u where "u = (\x. u2 x - u3 x)" have "AE x in M. abs (g x - real_cond_exp M Invariants f x) \ B x" unfolding g_def using * by auto moreover have "f x = g x + u x - u(T x)" for x using H3(4)[of x] H2(4)[of x] unfolding u_def g_def f3_def f2_def f1_def by auto moreover have "u \ borel_measurable M" unfolding u_def using \u2 \ borel_measurable M\ \u3 \ borel_measurable M\ by auto moreover have "integrable M g" unfolding g_def f3_def by (intro Bochner_Integration.integrable_add Bochner_Integration.integrable_diff \integrable M C\ \integrable M g3\ \integrable M f\ real_cond_exp_int(1)) ultimately show ?thesis by auto qed subsubsection \$L^1$ version of Birkhoff theorem\ text \The $L^1$ convergence in Birkhoff theorem follows from the almost everywhere convergence and general considerations on $L^1$ convergence (Scheffe's lemma) explained in \verb+AE_and_int_bound_implies_L1_conv2+. This argument works neatly for nonnegative functions, the general case reduces to this one by taking the positive and negative parts of a given function. One could also prove it by truncation: for bounded functions, the $L^1$ convergence follows from the boundedness and almost sure convergence. The general case follows by density, but it is a little bit tedious to write as one need to make sure that the conditional expectation of the truncation converges to the conditional expectation of the original function. This is true in $L^1$ as the conditional expectation is a contraction in $L^1$, it follows almost everywhere after taking a subsequence. All in all, the argument based on Scheffe's lemma seems more economical.\ lemma birkhoff_lemma_L1: fixes f::"'a \ real" assumes "\x. f x \ 0" and [measurable]: "integrable M f" shows "(\n. \\<^sup>+x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0" proof (rule Scheffe_lemma2) show i: "integrable M (real_cond_exp M Invariants f)" using assms by (simp add: real_cond_exp_int(1)) show "AE x in M. (\n. birkhoff_sum f n x / real n) \ real_cond_exp M Invariants f x" using birkhoff_theorem_AE_nonergodic assms by simp fix n have [measurable]: "(\x. ennreal \birkhoff_sum f n x\) \ borel_measurable M" by measurable show [measurable]: "(\x. birkhoff_sum f n x / real n) \ borel_measurable M" by measurable have "AE x in M. real_cond_exp M Invariants f x \ 0" using assms(1) real_cond_exp_pos by simp then have *: "AE x in M. norm (real_cond_exp M Invariants f x) = real_cond_exp M Invariants f x" by auto have **: "(\ x. norm (real_cond_exp M Invariants f x) \M) = (\ x. real_cond_exp M Invariants f x \M)" apply (rule integral_cong_AE) using * by auto have "(\\<^sup>+x. ennreal (norm (real_cond_exp M Invariants f x)) \M) = (\ x. norm (real_cond_exp M Invariants f x) \M)" by (rule nn_integral_eq_integral) (auto simp add: i) also have "... = (\ x. real_cond_exp M Invariants f x \M)" using ** by simp also have "... = (\ x. f x \M)" using real_cond_exp_int(2) assms(2) by auto also have "... = (\x. norm(f x) \M)" using assms by auto also have "... = (\\<^sup>+x. norm(f x) \M)" by (rule nn_integral_eq_integral[symmetric], auto simp add: assms(2)) finally have eq: "(\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M) = (\\<^sup>+ x. norm(f x) \M)" by simp { fix x have "norm(birkhoff_sum f n x) \ birkhoff_sum (\x. norm(f x)) n x" using birkhoff_sum_abs by fastforce then have "norm(birkhoff_sum f n x) \ birkhoff_sum (\x. ennreal(norm(f x))) n x" unfolding birkhoff_sum_def by auto } then have "(\\<^sup>+x. norm(birkhoff_sum f n x) \M) \ (\\<^sup>+x. birkhoff_sum (\x. ennreal(norm(f x))) n x \M)" by (simp add: nn_integral_mono) also have "... = n * (\\<^sup>+x. norm(f x) \M)" by (rule birkhoff_sum_nn_integral, auto) also have "... = n * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" using eq by simp finally have *: "(\\<^sup>+x. norm(birkhoff_sum f n x) \M) \ n * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" by simp show "(\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x / real n)) \M) \ (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" proof (cases) assume "n = 0" then show ?thesis by auto next assume "\(n = 0)" then have "n > 0" by simp then have "1/ennreal(real n) \ 0" by simp have "(\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x / real n)) \M) = (\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x)) / ennreal(real n) \M)" using \n > 0\ by (auto simp: divide_ennreal) also have "... = (\\<^sup>+ x. (1/ennreal(real n)) * ennreal (norm (birkhoff_sum f n x)) \M)" by (simp add: \0 < n\ divide_ennreal_def mult.commute) also have "... = (1/ennreal(real n) * (\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x)) \M))" by (subst nn_integral_cmult) auto also have "... \ (1/ennreal(real n)) * (ennreal(real n) * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M))" using * by (intro mult_mono) (auto simp: ennreal_of_nat_eq_real_of_nat) also have "... = (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" using \n > 0\ by (auto simp del: ennreal_1 simp add: ennreal_1[symmetric] divide_ennreal ennreal_mult[symmetric] mult.assoc[symmetric]) simp finally show ?thesis by simp qed qed theorem birkhoff_theorem_L1_nonergodic: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "(\n. \\<^sup>+x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0" proof - define g where "g = (\x. max (f x) 0)" have g_int [measurable]: "integrable M g" unfolding g_def using assms by auto define h where "h = (\x. max (-f x) 0)" have h_int [measurable]: "integrable M h" unfolding h_def using assms by auto have "f = (\x. g x - h x)" unfolding g_def h_def by auto { fix n::nat assume "n > 0" have "\x. birkhoff_sum f n x = birkhoff_sum g n x - birkhoff_sum h n x" using birkhoff_sum_diff \f = (\x. g x - h x)\ by auto then have "\x. birkhoff_sum f n x / n = birkhoff_sum g n x / n - birkhoff_sum h n x / n" using \n > 0\ by (simp add: diff_divide_distrib) moreover have "AE x in M. real_cond_exp M Invariants g x - real_cond_exp M Invariants h x = real_cond_exp M Invariants f x" using AE_symmetric[OF real_cond_exp_diff] g_int h_int \f = (\x. g x - h x)\ by auto ultimately have "AE x in M. birkhoff_sum f n x / n - real_cond_exp M Invariants f x = (birkhoff_sum g n x / n - real_cond_exp M Invariants g x) - (birkhoff_sum h n x / n - real_cond_exp M Invariants h x)" by auto then have *: "AE x in M. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \ norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) + norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x)" by auto have "(\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. ennreal(norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x)) + norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" apply (rule nn_integral_mono_AE) using * by (simp add: ennreal_plus[symmetric] del: ennreal_plus) also have "... = (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" apply (rule nn_integral_add) apply auto using real_cond_exp_F_meas borel_measurable_cond_exp2 by measurable finally have "(\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" by simp } then have *: "eventually (\n. (\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) sequentially" using eventually_at_top_dense by auto have **: "eventually (\n. (\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0) sequentially" by simp have "(\n. (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M)) \ 0" apply (rule birkhoff_lemma_L1, auto simp add: g_int) unfolding g_def by auto moreover have "(\n. (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) \ 0" apply (rule birkhoff_lemma_L1, auto simp add: h_int) unfolding h_def by auto ultimately have "(\n. (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) \ 0" using tendsto_add[of _ 0 _ _ 0] by auto then show ?thesis using tendsto_sandwich[OF ** *] by auto qed subsubsection \Conservativity of skew products\ text \The behaviour of skew-products of the form $(x, y) \mapsto (Tx, y + f x)$ is directly related to Birkhoff theorem, as the iterates involve the Birkhoff sums in the fiber. Birkhoff theorem implies that such a skew product is conservative when the function $f$ has vanishing conditional expectation. To prove the theorem, assume by contradiction that a set $A$ with positive measure does not intersect its preimages. Replacing $A$ with a smaller set $C$, we can assume that $C$ is bounded in the $y$-direction, by a constant $N$, and also that all its nonempty vertical fibers, above the projection $Cx$, have a measure bounded from below. Then, by Birkhoff theorem, for any $r>0$, most of the first $n$ preimages of $C$ are contained in the set $\{|y| \leq r n+N\}$, of measure $O(r n)$. Hence, they can not be disjoint if $r < \mu(C)$. To make this argument rigorous, one should only consider the preimages whose $x$-component belongs to a set $Dx$ where the Birkhoff sums are small. This condition has a positive measure if $\mu(Cx) + \mu(Dx) > \mu(M)$, which one can ensure by taking $Dx$ large enough.\ theorem (in fmpt) skew_product_conservative: fixes f::"'a \ real" assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" shows "conservative_mpt (M \\<^sub>M lborel) (\(x,y). (T x, y + f x))" proof (rule conservative_mptI) let ?TS = "(\(x,y). (T x, y + f x))" let ?MS = "M \\<^sub>M (lborel::real measure)" have f_meas [measurable]: "f \ borel_measurable M" by auto have "mpt M T" by (simp add: mpt_axioms) with mpt_skew_product_real[OF this f_meas] show "mpt ?MS ?TS" by simp then interpret TS: mpt ?MS ?TS by auto fix A::"('a \ real) set" assume A1 [measurable]: "A \ sets ?MS" and A2: "emeasure ?MS A > 0" have "A = (\N::nat. A \ {(x,y). abs(y) \ N})" by (auto simp add: real_arch_simple) then have *: "emeasure ?MS (\N::nat. A \ {(x,y). abs(y) \ N}) > 0" using A2 by simp have "space ?MS = space M \ space (lborel::real measure)" using space_pair_measure by auto then have A_inc: "A \ space M \ space (lborel::real measure)" using sets.sets_into_space[OF A1] by auto { fix N::nat have "{(x, y). abs(y) \ real N \ x \ space M} = space M \ {-(real N)..(real N)}" by auto then have "{(x, y). \y\ \ real N \ x \ space M} \ sets ?MS" by auto then have "A \ {(x, y). \y\ \ real N \ x \ space M} \ sets ?MS" using A1 by auto moreover have "A \ {(x,y). abs(y) \ real N} = A \ {(x, y). \y\ \ real N \ x \ space M}" using A_inc by blast ultimately have "A \ {(x,y). abs(y) \ real N} \ sets ?MS" by auto } then have [measurable]: "\N::nat. A \ {(x, y). \y\ \ real N} \ sets (M \\<^sub>M borel)" by auto have "\N::nat. emeasure ?MS (A \ {(x,y). abs(y) \ N}) > 0" apply (rule emeasure_pos_unionE) using * by auto then obtain N::nat where N: "emeasure ?MS (A \ {(x,y). abs(y) \ N}) > 0" by auto define B where "B = A \ {(x,y). abs(y) \ N}" have B_meas [measurable]: "B \ sets (M \\<^sub>M lborel)" unfolding B_def by auto have "0 < emeasure (M \\<^sub>M lborel) B" unfolding B_def using N by auto also have "... = (\\<^sup>+x. emeasure lborel (Pair x -` B) \M)" apply (rule sigma_finite_measure.emeasure_pair_measure_alt) using B_meas by (auto simp add: lborel.sigma_finite_measure_axioms) finally have *: "(\\<^sup>+x. emeasure lborel (Pair x -` B) \M) > 0" by simp have "\Cx\sets M. \e::real>0. emeasure M Cx > 0 \ (\x \ Cx. emeasure lborel (Pair x -` B) \ e)" by (rule not_AE_zero_int_ennreal_E, auto simp add: *) then obtain Cx e where [measurable]: "Cx \ sets M" and Cxe: "e>(0::real)" "emeasure M Cx > 0" "\x. x \ Cx \ emeasure lborel (Pair x -` B) \ e" by blast define C where "C = B \ (Cx \ (UNIV::real set))" have C_meas [measurable]: "C \ sets (M \\<^sub>M lborel)" unfolding C_def using B_meas by auto have Cx_fibers: "\x. x \ Cx \ emeasure lborel (Pair x -` C) \ e" using Cxe(3) C_def by auto define c where "c = (measure M Cx)/2" have "c > 0" unfolding c_def using Cxe(2) by (simp add: emeasure_eq_measure) text \We will apply Birkhoff theorem to show that most preimages of $C$ at time $n$ are contained in a cylinder of height roughly $r n$, for some suitably small $r$. How small $r$ should be to get a contradiction can be determined at the end of the proof. It turns out that the good condition is the following one -- this is by no means obvious now.\ define r where "r = (if measure M (space M) = 0 then 1 else e * c / (4 * measure M (space M)))" have "r > 0" using \e > 0\ \c > 0\ unfolding r_def apply auto using measure_le_0_iff by fastforce have pos: "e*c-2*r*measure M (space M) > 0" using \e > 0\ \c > 0\ unfolding r_def by auto define Bgood where "Bgood = {x \ space M. (\n. birkhoff_sum f n x / n) \ 0}" have [measurable]: "Bgood \ sets M" unfolding Bgood_def by auto have *: "AE x in M. x \ Bgood" unfolding Bgood_def using birkhoff_theorem_AE_nonergodic[OF assms(1)] assms(2) by auto then have "emeasure M Bgood = emeasure M (space M)" by (intro emeasure_eq_AE) auto { fix x assume "x \ Bgood" then have "x \ space M" unfolding Bgood_def by auto have "(\n. birkhoff_sum f n x / n) \ 0" using \x \ Bgood\ unfolding Bgood_def by auto moreover have "0 \ {-r<..r>0\ by auto ultimately have "eventually (\n. birkhoff_sum f n x / n \ {-r<..0" "\n. n \ n0 \ birkhoff_sum f n x / n \ {-r<.. n0" then have "n>0" using \n0>0\ by auto with n0(2)[OF \n \ n0\] have "abs(birkhoff_sum f n x / n) \ r" by auto then have "abs(birkhoff_sum f n x) \ r * n" using \n>0\ by (simp add: divide_le_eq) } then have "x \ (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n})" using \x \ space M\ by blast } then have "AE x in M. x \ space M - (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n})" using * by auto then have eqM: "emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) = emeasure M (space M)" by (intro emeasure_eq_AE) auto have "(\n0. emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c) \ emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) + c" by (intro tendsto_intros Lim_emeasure_incseq) (auto simp add: incseq_def) moreover have "emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) + c > emeasure M (space M)" using eqM \c > 0\ emeasure_eq_measure by auto ultimately have "eventually (\n0. emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c > emeasure M (space M)) sequentially" unfolding order_tendsto_iff by auto then obtain n0 where n0: "emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c > emeasure M (space M)" using eventually_sequentially by auto define Dx where "Dx = {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}" have Dx_meas [measurable]: "Dx \ sets M" unfolding Dx_def by auto have "emeasure M Dx + c \ emeasure M (space M)" using n0 Dx_def by auto obtain n1::nat where n1: "n1 > max n0 ((measure M (space M) * 2 * N + e*c*n0 - e*c) / (e*c-2*r*measure M (space M)))" by (metis mult.commute mult.left_neutral numeral_One reals_Archimedean3 zero_less_numeral) then have "n1 > n0" by auto have n1_ineq: "n1 * (e*c-2*r*measure M (space M)) > (measure M (space M) * 2 * N + e*c*n0 - e*c)" using n1 pos by (simp add: pos_divide_less_eq) define D where "D = (\n. Dx \ {-r*n1-N..r*n1+N} \ (?TS^^n)-`C)" have Dn_meas [measurable]: "D n \ sets (M \\<^sub>M lborel)" for n unfolding D_def apply (rule TS.T_intersec_meas(2)) using C_meas by auto have "emeasure ?MS (D n) \ e * c" if "n \ {n0..n1}" for n proof - have "n \ n0" "n \ n1" using that by auto { fix x assume [simp]: "x \ space M" define F where "F = {y \ {-r*n1-N..r*n1+N}. y + birkhoff_sum f n x \ Pair ((T^^n)x) -`C}" have [measurable]: "F \ sets lborel" unfolding F_def by measurable { fix y::real have "(?TS^^n)(x,y) = ((T^^n)x, y + birkhoff_sum f n x)" using skew_product_real_iterates by simp then have "(indicator C ((?TS^^n) (x,y))::ennreal) = indicator Cx ((T^^n)x) * indicator (Pair ((T^^n)x) -`C) (y + birkhoff_sum f n x)" using C_def by (simp add: indicator_def) moreover have "(indicator (D n) (x, y)::ennreal) = indicator Dx x * indicator {-r*n1-N..r*n1+N} y * indicator C ((?TS^^n) (x,y))" unfolding D_def by (simp add: indicator_def) ultimately have "(indicator (D n) (x, y)::ennreal) = indicator Dx x * indicator {-r*n1-N..r*n1+N} y * indicator Cx ((T^^n)x) * indicator (Pair ((T^^n)x) -`C) (y + birkhoff_sum f n x)" by (simp add: mult.assoc) then have "(indicator (D n) (x, y)::ennreal) = indicator (Dx \ (T^^n)-`Cx) x * indicator F y" unfolding F_def by (simp add: indicator_def) } then have "(\\<^sup>+y. indicator (D n) (x, y) \lborel) = (\\<^sup>+y. indicator (Dx \ (T^^n)-`Cx) x * indicator F y \lborel)" by auto also have "... = indicator (Dx \ (T^^n)-`Cx) x * (\\<^sup>+y. indicator F y \lborel)" by (rule nn_integral_cmult, auto) also have "... = indicator (Dx \ (T^^n)-`Cx) x * emeasure lborel F" using \F \ sets lborel\ by auto finally have A: "(\\<^sup>+y. indicator (D n) (x, y) \lborel) = indicator (Dx \ (T^^n)-`Cx) x * emeasure lborel F" by simp have "(\\<^sup>+y. indicator (D n) (x, y) \lborel) \ ennreal e * indicator (Dx \ (T^^n)-`Cx) x" proof (cases) assume "indicator (Dx \ (T^^n)-`Cx) x = (0::ennreal)" then show ?thesis by auto next assume "\(indicator (Dx \ (T^^n)-`Cx) x = (0::ennreal))" then have "x \ Dx \ (T^^n)-`Cx" by (simp add: indicator_eq_0_iff) then have "x \ Dx" "(T^^n) x \ Cx" by auto then have "abs(birkhoff_sum f n x) \ r * n" using \n \ {n0..n1}\ Dx_def by auto then have *: "abs(birkhoff_sum f n x) \ r * n1" using \n \ n1\ \r>0\ by (meson of_nat_le_iff order_trans real_mult_le_cancel_iff2) have F_expr: "F = {-r*n1-N..r*n1+N} \ (+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C)" unfolding F_def by (auto simp add: add.commute) have "(Pair ((T^^n)x) -`C) \ {real_of_int (- int N)..real N}" unfolding C_def B_def by auto then have "((+)(birkhoff_sum f n x)) -` (Pair ((T^^n)x) -`C) \ {-N-birkhoff_sum f n x..N-birkhoff_sum f n x}" by auto also have "... \ {-r * n1 - N .. r * n1 + N}" using * by auto finally have "F = ((+)(birkhoff_sum f n x)) -` (Pair ((T^^n)x) -`C)" unfolding F_expr by auto then have "emeasure lborel F = emeasure lborel ((+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C))" by auto also have "... = emeasure lborel (((+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C)) \ space lborel)" by simp also have "... = emeasure (distr lborel borel ((+) (birkhoff_sum f n x))) (Pair ((T^^n)x) -`C)" apply (rule emeasure_distr[symmetric]) using C_meas by auto also have "... = emeasure lborel (Pair ((T^^n)x) -`C)" using lborel_distr_plus[of "birkhoff_sum f n x"] by simp also have "... \ e" using Cx_fibers \(T^^n) x \ Cx\ by auto finally have "emeasure lborel F \ e" by auto then show ?thesis using A by (simp add: indicator_def) qed } moreover have "emeasure ?MS (D n) = (\\<^sup>+x. (\\<^sup>+y. indicator (D n) (x, y) \lborel) \M)" using Dn_meas lborel.emeasure_pair_measure by blast ultimately have "emeasure ?MS (D n) \ (\\<^sup>+x. ennreal e * indicator (Dx \ (T ^^ n) -` Cx) x \M)" by (simp add: nn_integral_mono) also have "(\\<^sup>+x. ennreal e * indicator (Dx \ (T ^^ n) -` Cx) x \M) = e * (\\<^sup>+x. indicator (Dx \ (T ^^ n) -` Cx) x \M)" apply (rule nn_integral_cmult) using \e>0\ by auto also have "... = ennreal e * emeasure M (Dx \ (T ^^ n) -` Cx)" by simp finally have *: "emeasure ?MS (D n) \ ennreal e * emeasure M (Dx \ (T ^^ n) -` Cx)" by auto have "c + emeasure M (space M) \ emeasure M Dx + emeasure M Cx" using \emeasure M Dx + c \ emeasure M (space M)\ unfolding c_def by (auto simp: emeasure_eq_measure ennreal_plus[symmetric] simp del: ennreal_plus) also have "... = emeasure M Dx + emeasure M ((T^^n)--`Cx)" by (simp add: T_vrestr_same_emeasure(2)) also have "... = emeasure M (Dx \ ((T^^n)--`Cx)) + emeasure M (Dx \ ((T^^n)--`Cx))" by (rule emeasure_union_inter, auto) also have "... \ emeasure M (space M) + emeasure M (Dx \ ((T^^n)-`Cx))" proof - have "emeasure M (Dx \ ((T^^n)--`Cx)) \ emeasure M (space M)" by (rule emeasure_mono, auto simp add: sets.sets_into_space) moreover have "Dx \ ((T^^n)--`Cx) = Dx \ ((T^^n)-`Cx)" by (simp add: vrestr_intersec_in_space) ultimately show ?thesis by (metis add.commute add_left_mono) qed finally have "emeasure M (Dx \ ((T^^n)-`Cx)) \ c" by (simp add: emeasure_eq_measure) then have "ennreal e * emeasure M (Dx \ ((T^^n)-`Cx)) \ ennreal e * c" using \e > 0\ using mult_left_mono by fastforce with * show "emeasure ?MS (D n) \ e * c" using \0 \0 by (auto simp: ennreal_mult[symmetric]) qed have "\(disjoint_family_on D {n0..n1})" proof assume D: "disjoint_family_on D {n0..n1}" have "emeasure lborel {-r*n1-N..r*n1+N} = (r * real n1 + real N) - (- r * real n1 - real N)" apply (rule emeasure_lborel_Icc) using \r>0\ by auto then have *: "emeasure lborel {-r*n1-N..r*n1+N} = ennreal(2 * r * real n1 + 2 * real N)" by (auto simp: ac_simps) have "ennreal(e * c) * (real n1 - real n0 + 1) = ennreal(e*c) * card {n0..n1}" using \n1 > n0\ by (auto simp: ennreal_of_nat_eq_real_of_nat Suc_diff_le ac_simps of_nat_diff) also have "... = (\n\{n0..n1}. ennreal(e*c))" by (simp add: ac_simps) also have "... \ (\n\{n0..n1}. emeasure ?MS (D n))" using \\n. n \ {n0..n1} \ emeasure ?MS (D n) \ e * c\ by (meson sum_mono) also have "... = emeasure ?MS (\n\{n0..n1}. D n)" apply (rule sum_emeasure) using Dn_meas by (auto simp add: D) also have "... \ emeasure ?MS (space M \ {-r*n1-N..r*n1+N})" apply (rule emeasure_mono) unfolding D_def using sets.sets_into_space[OF Dx_meas] by auto also have "... = emeasure M (space M) * emeasure lborel {-r*n1-N..r*n1+N}" by (rule sigma_finite_measure.emeasure_pair_measure_Times, auto simp add: lborel.sigma_finite_measure_axioms) also have "... = emeasure M (space M) * ennreal(2 * r * real n1 + 2 * real N)" using * by auto finally have "ennreal(e * c) * (real n1- real n0+1) \ emeasure M (space M) * ennreal(2 * r * real n1 + 2 * real N)" by simp then have "e*c * (real n1- real n0 + 1) \ measure M (space M) * (2 * r * real n1 + 2 * real N)" using \0 \0 \0 \n0 < n1\ emeasure_eq_measure by (auto simp: ennreal_mult'[symmetric] simp del: ennreal_plus) then have "0 \ measure M (space M) * (2 * r * real n1 + 2 * real N) - e*c * (real n1- real n0 + 1)" by auto also have "... = (measure M (space M) * 2 * N + e*c*n0 - e*c) - n1 * (e*c-2*r*measure M (space M))" by algebra finally have "n1 * (e*c-2*r*measure M (space M)) \ measure M (space M) * 2 * N + e*c*n0 - e*c" by linarith then show False using n1_ineq by auto qed then obtain n m where nm: "n D n \ {}" unfolding disjoint_family_on_def by (metis inf_sup_aci(1) linorder_cases) define k where "k = m-n" then have "k>0" "D (n+k) \ D n \ {}" using nm by auto then have "((?TS^^(n+k))-`A) \ ((?TS^^n)-`A) \ {}" unfolding D_def C_def B_def by auto moreover have "((?TS^^(n+k))-`A) \ ((?TS^^n)-`A) = (?TS^^n)-`(((?TS^^k)-`A) \ A)" using funpow_add by (simp add: add.commute funpow_add set.compositionality) ultimately have "((?TS^^k)-`A) \ A \ {}" by auto then show "\k>0. ((?TS^^k)-`A) \ A \ {}" using \k>0\ by auto qed subsubsection \Oscillations around the limit in Birkhoff theorem\ text \In this paragraph, we prove that, in Birkhoff theorem with vanishing limit, the Birkhoff sums are infinitely many times arbitrarily close to $0$, both on the positive and the negative side. In the ergodic case, this statement implies for instance that if the Birkhoff sums of an integrable function tend to infinity almost everywhere, then the integral of the function can not vanish, it has to be strictly positive (while Birkhoff theorem per se does not exclude the convergence to infinity, at a rate slower than linear). This converts a qualitative information (convergence to infinity at an unknown rate) to a quantitative information (linear convergence to infinity). This result (sometimes known as Atkinson's Lemma) has been reinvented many times, for instance by Kesten and by Guivarch. It plays an important role in the study of random products of matrices. This is essentially a consequence of the conservativity of the corresponding skew-product, proved in \verb+skew_product_conservative+. Indeed, this implies that, starting from a small set $X \times [-e/2, e/2]$, the skew-product comes back infinitely often to itself, which implies that the Birkhoff sums at these times are bounded by $e$. To show that the Birkhoff sums come back to $[0,e]$ is a little bit more tricky. Argue by contradiction, and induce on $A \times [0,e/2]$ where $A$ is the set of points where the Birkhoff sums don't come back to $[0,e]$. Then the second coordinate decreases strictly when one iterates the skew product, which is not compatible with conservativity.\ lemma birkhoff_sum_small_asymp_lemma: assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" "e>(0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {0..e}}" proof - have [measurable]: "f \ borel_measurable M" by auto have [measurable]: "\N. {x \ space M. \N. \n\{N..}. birkhoff_sum f n x \ {0..e}} \ sets M" by auto { fix N assume "N>(0::nat)" define Ax where "Ax = {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}}" then have [measurable]: "Ax \ sets M" by auto define A where "A = Ax \ {0..e/2}" then have A_meas [measurable]: "A \ sets (M \\<^sub>M lborel)" by auto define TN where "TN = T^^N" interpret TN: fmpt M TN unfolding TN_def using fmpt_power by auto define fN where "fN = birkhoff_sum f N" have "TN.birkhoff_sum fN n x = birkhoff_sum f (n*N) x" for n x proof (induction n) case 0 then show ?case by auto next case (Suc n) have "TN.birkhoff_sum fN (Suc n) x = TN.birkhoff_sum fN n x + fN ((TN^^n) x)" using TN.birkhoff_sum_cocycle[of fN n 1] by auto also have "... = birkhoff_sum f (n*N) x + birkhoff_sum f N ((TN^^n) x)" using Suc.IH fN_def by auto also have "... = birkhoff_sum f (n*N+N) x" unfolding TN_def by (subst funpow_mult, subst mult.commute[of N n], rule birkhoff_sum_cocycle[of f "n*N" N x, symmetric]) finally show ?case by (simp add: add.commute) qed then have not0e: "\x n. x \ Ax \ n \ 0 \ TN.birkhoff_sum fN n x \ {0..e}" unfolding Ax_def by auto let ?TS = "(\(x,y). (T x, y + f x))" let ?MS = "M \\<^sub>M (lborel::real measure)" interpret TS: conservative_mpt ?MS ?TS by (rule skew_product_conservative, auto simp add: assms) let ?TSN = "(\(x,y). (TN x, y + fN x))" have *:"?TSN = ?TS^^N" unfolding TN_def fN_def using skew_product_real_iterates by auto interpret TSN: conservative_mpt ?MS ?TSN using * TS.conservative_mpt_power by auto define MA TA where "MA = restrict_space ?MS A" and "TA = TSN.induced_map A" interpret TA: conservative_mpt MA TA unfolding MA_def TA_def by (rule TSN.induced_map_conservative_mpt, measurable) have *: "\ x y. snd (TA (x,y)) = snd (x,y) + TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x" unfolding TA_def TSN.induced_map_def using TN.skew_product_real_iterates Pair_def by auto have [measurable]: "snd \ borel_measurable ?MS" by auto then have [measurable]: "snd \ borel_measurable MA" unfolding MA_def using measurable_restrict_space1 by blast have "AE z in MA. z \ TSN.recurrent_subset A" unfolding MA_def using TSN.induced_map_recurrent_typical(1)[OF A_meas]. moreover { fix z assume z: "z \ TSN.recurrent_subset A" define x y where "x = fst z" and "y = snd z" then have "z = (x,y)" by simp have "z \ A" using z "TSN.recurrent_subset_incl"(1) by auto then have "x \ Ax" "y \ {0..e/2}" unfolding A_def x_def y_def by auto define y2 where "y2 = y + TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x" have "y2 = snd (TA z)" unfolding y2_def using * \z = (x, y)\ by force moreover have "TA z \ A" unfolding TA_def using \z \ A\ TSN.induced_map_stabilizes_A by auto ultimately have "y2 \ {0..e/2}" unfolding A_def by auto have "TSN.return_time_function A (x,y) \ 0" using \z = (x,y)\ \z \ TSN.recurrent_subset A\ TSN.return_time0[of A] by fast then have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {0..e}" using not0e[OF \x \ Ax\] by auto moreover have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {-e..e}" using \y \ {0..e/2}\ \y2 \ {0..e/2}\ y2_def by auto ultimately have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {-e..<0}" by auto then have "y2 < y" using y2_def by auto then have "snd(TA z) < snd z" unfolding y_def using \y2 = snd (TA z)\ by auto } ultimately have a: "AE z in MA. snd(TA z) < snd z" by auto then have "AE z in MA. snd(TA z) \ snd z" by auto then have "AE z in MA. snd(TA z) = snd z" using TA.AE_decreasing_then_invariant[of snd] by auto then have "AE z in MA. False" using a by auto then have "space MA \ null_sets MA" by (simp add: AE_iff_null_sets) then have "emeasure MA A = 0" by (metis A_meas MA_def null_setsD1 space_restrict_space2) then have "emeasure ?MS A = 0" unfolding MA_def by (metis A_meas emeasure_restrict_space sets.sets_into_space sets.top space_restrict_space space_restrict_space2) moreover have "emeasure ?MS A = emeasure M Ax * emeasure lborel {0..e/2}" unfolding A_def by (intro lborel.emeasure_pair_measure_Times) auto ultimately have "emeasure M {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}} = 0" using \e>0\ Ax_def by simp then have "{x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}} \ null_sets M" by auto } then have "(\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}}) \ null_sets M" by (auto simp: greaterThan_0) moreover have "{x \ space M. \(infinite {n. birkhoff_sum f n x \ {0..e}})} \ (\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}})" proof fix x assume H: "x \ {x \ space M. \(infinite {n. birkhoff_sum f n x \ {0..e}})}" then have "x \ space M" by auto have *: "finite {n. birkhoff_sum f n x \ {0..e}}" using H by auto then obtain N where "\n. n \ N \ n \ {n. birkhoff_sum f n x \ {0..e}}" by (metis finite_nat_set_iff_bounded not_less) then have "x \ {x \ space M. \n\{N+1..}. birkhoff_sum f n x \ {0..e}}" using \x \ space M\ by auto moreover have "N+1>0" by auto ultimately show "x \ (\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}})" by auto qed ultimately show ?thesis unfolding eventually_ae_filter by auto qed theorem birkhoff_sum_small_asymp_pos_nonergodic: assumes [measurable]: "integrable M f" and "e > (0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" proof - define g where "g = (\x. f x - real_cond_exp M Invariants f x)" have g_meas [measurable]: "integrable M g" unfolding g_def using real_cond_exp_int(1)[OF assms(1)] assms(1) by auto have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (rule real_cond_exp_F_meas, auto simp add: real_cond_exp_int(1)[OF assms(1)]) then have *: "AE x in M. real_cond_exp M Invariants g x = 0" unfolding g_def using real_cond_exp_diff[OF assms(1) real_cond_exp_int(1)[OF assms(1)]] by auto have "AE x in M. infinite {n. birkhoff_sum g n x \ {0..e}}" by (rule birkhoff_sum_small_asymp_lemma, auto simp add: \e>0\ * g_meas) moreover { fix x assume "x \ space M" "infinite {n. birkhoff_sum g n x \ {0..e}}" { fix n assume H: "birkhoff_sum g n x \ {0..e}" have "birkhoff_sum g n x = birkhoff_sum f n x - birkhoff_sum (real_cond_exp M Invariants f) n x" unfolding g_def using birkhoff_sum_diff by auto also have "... = birkhoff_sum f n x - n * real_cond_exp M Invariants f x" using birkhoff_sum_of_invariants \x \ space M\ by auto finally have "birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}" using H by simp } then have "{n. birkhoff_sum g n x \ {0..e}} \ {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" by auto then have "infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" using \infinite {n. birkhoff_sum g n x \ {0..e}}\ finite_subset by blast } ultimately show ?thesis by auto qed theorem birkhoff_sum_small_asymp_neg_nonergodic: assumes [measurable]: "integrable M f" and "e > (0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" proof - define g where "g = (\x. real_cond_exp M Invariants f x - f x)" have g_meas [measurable]: "integrable M g" unfolding g_def using real_cond_exp_int(1)[OF assms(1)] assms(1) by auto have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (rule real_cond_exp_F_meas, auto simp add: real_cond_exp_int(1)[OF assms(1)]) then have *: "AE x in M. real_cond_exp M Invariants g x = 0" unfolding g_def using real_cond_exp_diff[OF real_cond_exp_int(1)[OF assms(1)] assms(1)] by auto have "AE x in M. infinite {n. birkhoff_sum g n x \ {0..e}}" by (rule birkhoff_sum_small_asymp_lemma, auto simp add: \e>0\ * g_meas) moreover { fix x assume "x \ space M" "infinite {n. birkhoff_sum g n x \ {0..e}}" { fix n assume H: "birkhoff_sum g n x \ {0..e}" have "birkhoff_sum g n x = birkhoff_sum (real_cond_exp M Invariants f) n x - birkhoff_sum f n x" unfolding g_def using birkhoff_sum_diff by auto also have "... = n * real_cond_exp M Invariants f x - birkhoff_sum f n x" using birkhoff_sum_of_invariants \x \ space M\ by auto finally have "birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}" using H by simp } then have "{n. birkhoff_sum g n x \ {0..e}} \ {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" by auto then have "infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" using \infinite {n. birkhoff_sum g n x \ {0..e}}\ finite_subset by blast } ultimately show ?thesis by auto qed subsubsection \Conditional expectation for the induced map\ text \Thanks to Birkhoff theorem, one can relate conditional expectations with respect to the invariant sigma algebra, for a map and for a corresponding induced map, as follows.\ proposition Invariants_cond_exp_induced_map: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" defines "MA \ restrict_space M A" and "TA \ induced_map A" and "fA \ induced_function A f" shows "AE x in MA. real_cond_exp MA (qmpt.Invariants MA TA) fA x = real_cond_exp M Invariants f x * real_cond_exp MA (qmpt.Invariants MA TA) (return_time_function A) x" proof - interpret A: fmpt MA TA unfolding MA_def TA_def by (rule induced_map_fmpt[OF assms(1)]) have "integrable M fA" unfolding fA_def using induced_function_integral_nonergodic(1) assms by auto then have "integrable MA fA" unfolding MA_def by (metis assms(1) integrable_mult_indicator integrable_restrict_space sets.Int_space_eq2) then have a: "AE x in MA. (\n. A.birkhoff_sum fA n x / n) \ real_cond_exp MA A.Invariants fA x" using A.birkhoff_theorem_AE_nonergodic by auto have "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" using birkhoff_theorem_AE_nonergodic assms(2) by auto then have b: "AE x in MA. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" unfolding MA_def by (metis (mono_tags, lifting) AE_restrict_space_iff assms(1) eventually_mono sets.Int_space_eq2) define phiA where "phiA = (\x. return_time_function A x)" have "integrable M phiA" unfolding phiA_def using return_time_integrable by auto then have "integrable MA phiA" unfolding MA_def by (metis assms(1) integrable_mult_indicator integrable_restrict_space sets.Int_space_eq2) then have c: "AE x in MA. (\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" using A.birkhoff_theorem_AE_nonergodic by auto have "A-recurrent_subset A \ null_sets M" using Poincare_recurrence_thm(1)[OF assms(1)] by auto then have "A - recurrent_subset A \ null_sets MA" unfolding MA_def by (metis Diff_subset assms(1) emeasure_restrict_space null_setsD1 null_setsD2 null_setsI sets.Int_space_eq2 sets_restrict_space_iff) then have "AE x in MA. x \ recurrent_subset A" by (simp add: AE_iff_null MA_def null_setsD2 set_diff_eq space_restrict_space2) moreover have "\x. x \ recurrent_subset A \ phiA x > 0" unfolding phiA_def using return_time0 by fastforce ultimately have *: "AE x in MA. phiA x > 0" by auto have d: "AE x in MA. real_cond_exp MA A.Invariants phiA x > 0" by (rule A.real_cond_exp_gr_c, auto simp add: * \integrable MA phiA\) { fix x assume A: "(\n. A.birkhoff_sum fA n x / n) \ real_cond_exp MA A.Invariants fA x" and B: "(\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" and C: "(\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" and D: "real_cond_exp MA A.Invariants phiA x > 0" define R where "R = (\n. A.birkhoff_sum phiA n x)" have D2: "ereal(real_cond_exp MA A.Invariants phiA x) > 0" using D by simp have "\n. real(R n)/ n = A.birkhoff_sum (\x. real(phiA x)) n x / n" unfolding R_def A.birkhoff_sum_def by auto moreover have "(\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" using C by auto ultimately have Rnn: "(\n. real(R n)/n) \ real_cond_exp MA A.Invariants phiA x" by presburger have "\n. ereal(real(R n))/ n = ereal(A.birkhoff_sum (\x. real(phiA x)) n x / n)" unfolding R_def A.birkhoff_sum_def by auto moreover have "(\n. ereal(A.birkhoff_sum (\x. real(phiA x)) n x / n)) \ real_cond_exp MA A.Invariants phiA x" using C by auto ultimately have i: "(\n. ereal(real(R n))/n) \ real_cond_exp MA A.Invariants phiA x" by auto have ii: "(\n. real n) \ \" by (rule id_nat_ereal_tendsto_PInf) have iii: "(\n. ereal(real(R n))/n * real n) \ \" using tendsto_mult_ereal_PInf[OF i D2 ii] by simp have "\n. n > 0 \ ereal(real(R n))/n * real n = R n" by auto then have "eventually (\n. ereal(real(R n))/n * real n = R n) sequentially" using eventually_at_top_dense by auto then have "(\n. real(R n)) \ \" using iii by (simp add: filterlim_cong) then have "(\n. birkhoff_sum f (R n) x / (R n)) \ real_cond_exp M Invariants f x" using limit_along_weak_subseq B by auto then have l: "(\n. (birkhoff_sum f (R n) x / (R n)) * ((R n)/n)) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" by (rule tendsto_mult, simp add: Rnn) obtain N where N: "\n. n > N \ R n > 0" using \(\n. real(R n)) \ \\ by (metis (full_types) eventually_at_top_dense filterlim_iff filterlim_weak_subseq) then have "\n. n > N \ (birkhoff_sum f (R n) x / (R n)) * ((R n)/n) = birkhoff_sum f (R n) x / n" by auto then have "eventually (\n. (birkhoff_sum f (R n) x / (R n)) * ((R n)/n) = birkhoff_sum f (R n) x / n) sequentially" by simp with tendsto_cong[OF this] have main_limit: "(\n. birkhoff_sum f (R n) x / n) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using l by auto have "\n. birkhoff_sum f (R n) x = A.birkhoff_sum fA n x" unfolding R_def fA_def phiA_def TA_def using induced_function_birkhoff_sum[OF assms(1)] by simp then have "\n. birkhoff_sum f (R n) x /n = A.birkhoff_sum fA n x / n" by simp then have "(\n. A.birkhoff_sum fA n x / n) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using main_limit by presburger then have "real_cond_exp MA A.Invariants fA x = real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using A LIMSEQ_unique by auto } then show ?thesis using a b c d unfolding phiA_def by auto qed corollary Invariants_cond_exp_induced_map_0: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" defines "MA \ restrict_space M A" and "TA \ induced_map A" and "fA \ induced_function A f" shows "AE x in MA. real_cond_exp MA (qmpt.Invariants MA TA) fA x = 0" proof - have "AE x in MA. real_cond_exp M Invariants f x = 0" unfolding MA_def apply (subst AE_restrict_space_iff) using assms(3) by auto then show ?thesis unfolding MA_def TA_def fA_def using Invariants_cond_exp_induced_map[OF assms(1) assms(2)] by auto qed end (*of locale fmpt*) end (*of Invariants.thy*) diff --git a/thys/Gromov_Hyperbolicity/Isometries.thy b/thys/Gromov_Hyperbolicity/Isometries.thy --- a/thys/Gromov_Hyperbolicity/Isometries.thy +++ b/thys/Gromov_Hyperbolicity/Isometries.thy @@ -1,3158 +1,3160 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Isometries\ theory Isometries imports Library_Complements Hausdorff_Distance begin text \Isometries, i.e., functions that preserve distances, show up very often in mathematics. We introduce a dedicated definition, and show its basic properties.\ definition isometry_on::"('a::metric_space) set \ ('a \ ('b::metric_space)) \ bool" where "isometry_on X f = (\x \ X. \y \ X. dist (f x) (f y) = dist x y)" definition isometry :: "('a::metric_space \ 'b::metric_space) \ bool" where "isometry f \ isometry_on UNIV f \ range f = UNIV" lemma isometry_on_subset: assumes "isometry_on X f" "Y \ X" shows "isometry_on Y f" using assms unfolding isometry_on_def by auto lemma isometry_onI [intro?]: assumes "\x y. x \ X \ y \ X \ dist (f x) (f y) = dist x y" shows "isometry_on X f" using assms unfolding isometry_on_def by auto lemma isometry_onD: assumes "isometry_on X f" "x \ X" "y \ X" shows "dist (f x) (f y) = dist x y" using assms unfolding isometry_on_def by auto lemma isometryI [intro?]: assumes "\x y. dist (f x) (f y) = dist x y" "range f = UNIV" shows "isometry f" unfolding isometry_def isometry_on_def using assms by auto lemma assumes "isometry_on X f" shows isometry_on_lipschitz: "1-lipschitz_on X f" and isometry_on_uniformly_continuous: "uniformly_continuous_on X f" and isometry_on_continuous: "continuous_on X f" proof - show "1-lipschitz_on X f" apply (rule lipschitz_onI) using isometry_onD[OF assms] by auto then show "uniformly_continuous_on X f" "continuous_on X f" using lipschitz_on_uniformly_continuous lipschitz_on_continuous_on by auto qed lemma isometryD: assumes "isometry f" shows "isometry_on UNIV f" "dist (f x) (f y) = dist x y" "range f = UNIV" "1-lipschitz_on UNIV f" "uniformly_continuous_on UNIV f" "continuous_on UNIV f" using assms unfolding isometry_def isometry_on_def apply auto using isometry_on_lipschitz isometry_on_uniformly_continuous isometry_on_continuous assms unfolding isometry_def by blast+ lemma isometry_on_injective: assumes "isometry_on X f" shows "inj_on f X" using assms inj_on_def isometry_on_def by force lemma isometry_on_compose: assumes "isometry_on X f" "isometry_on (f`X) g" shows "isometry_on X (\x. g(f x))" using assms unfolding isometry_on_def by auto lemma isometry_on_cong: assumes "isometry_on X f" "\x. x \ X \ g x = f x" shows "isometry_on X g" using assms unfolding isometry_on_def by auto lemma isometry_on_inverse: assumes "isometry_on X f" shows "isometry_on (f`X) (inv_into X f)" "\x. x \ X \ (inv_into X f) (f x) = x" "\y. y \ f`X \ f (inv_into X f y) = y" "bij_betw f X (f`X)" proof - show *: "bij_betw f X (f`X)" using assms unfolding bij_betw_def inj_on_def isometry_on_def by force show "isometry_on (f`X) (inv_into X f)" using assms unfolding isometry_on_def by (auto) (metis (mono_tags, lifting) dist_eq_0_iff inj_on_def inv_into_f_f) fix x assume "x \ X" then show "(inv_into X f) (f x) = x" using * by (simp add: bij_betw_def) next fix y assume "y \ f`X" then show "f (inv_into X f y) = y" by (simp add: f_inv_into_f) qed lemma isometry_inverse: assumes "isometry f" shows "isometry (inv f)" "bij f" using isometry_on_inverse[OF isometryD(1)[OF assms]] isometryD(3)[OF assms] unfolding isometry_def by (auto simp add: bij_imp_bij_inv bij_is_surj) lemma isometry_on_homeomorphism: assumes "isometry_on X f" shows "homeomorphism X (f`X) f (inv_into X f)" "homeomorphism_on X f" "X homeomorphic f`X" proof - show *: "homeomorphism X (f`X) f (inv_into X f)" apply (rule homeomorphismI) using uniformly_continuous_imp_continuous[OF isometry_on_uniformly_continuous] isometry_on_inverse[OF assms] assms by auto then show "X homeomorphic f`X" unfolding homeomorphic_def by auto show "homeomorphism_on X f" unfolding homeomorphism_on_def using * by auto qed lemma isometry_homeomorphism: fixes f::"('a::metric_space) \ ('b::metric_space)" assumes "isometry f" shows "homeomorphism UNIV UNIV f (inv f)" "(UNIV::'a set) homeomorphic (UNIV::'b set)" using isometry_on_homeomorphism[OF isometryD(1)[OF assms]] isometryD(3)[OF assms] by auto lemma isometry_on_closure: assumes "isometry_on X f" "continuous_on (closure X) f" shows "isometry_on (closure X) f" proof (rule isometry_onI) fix x y assume "x \ closure X" "y \ closure X" obtain u v::"nat \ 'a" where *: "\n. u n \ X" "u \ x" "\n. v n \ X" "v \ y" using \x \ closure X\ \y \ closure X\ unfolding closure_sequential by blast have "(\n. f (u n)) \ f x" using *(1) *(2) \x \ closure X\ \continuous_on (closure X) f\ unfolding comp_def continuous_on_closure_sequentially[of X f] by auto moreover have "(\n. f (v n)) \ f y" using *(3) *(4) \y \ closure X\ \continuous_on (closure X) f\ unfolding comp_def continuous_on_closure_sequentially[of X f] by auto ultimately have "(\n. dist (f (u n)) (f (v n))) \ dist (f x) (f y)" by (simp add: tendsto_dist) then have "(\n. dist (u n) (v n)) \ dist (f x) (f y)" using assms(1) *(1) *(3) unfolding isometry_on_def by auto moreover have "(\n. dist (u n) (v n)) \ dist x y" using *(2) *(4) by (simp add: tendsto_dist) ultimately show "dist (f x) (f y) = dist x y" using LIMSEQ_unique by auto qed lemma isometry_extend_closure: fixes f::"('a::metric_space) \ ('b::complete_space)" assumes "isometry_on X f" shows "\g. isometry_on (closure X) g \ (\x\X. g x = f x)" proof - obtain g where g: "\x. x \ X \ g x = f x" "uniformly_continuous_on (closure X) g" using uniformly_continuous_on_extension_on_closure[OF isometry_on_uniformly_continuous[OF assms]] by metis have "isometry_on (closure X) g" apply (rule isometry_on_closure, rule isometry_on_cong[OF assms]) using g uniformly_continuous_imp_continuous[OF g(2)] by auto then show ?thesis using g(1) by auto qed lemma isometry_on_complete_image: assumes "isometry_on X f" "complete X" shows "complete (f`X)" proof (rule completeI) fix u :: "nat \ 'b" assume u: "\n. u n \ f`X" "Cauchy u" define v where "v = (\n. inv_into X f (u n))" have "v n \ X" for n unfolding v_def by (simp add: inv_into_into u(1)) have "dist (v n) (v m) = dist (u n) (u m)" for m n using u(1) isometry_on_inverse[OF \isometry_on X f\] unfolding isometry_on_def v_def by (auto simp add: inv_into_into) then have "Cauchy v" using u(2) unfolding Cauchy_def by auto obtain l where "l \ X" "v \ l" apply (rule completeE[OF \complete X\ _ \Cauchy v\]) using \\n. v n \ X\ by auto have "(\n. f (v n)) \ f l" apply (rule continuous_on_tendsto_compose[OF isometry_on_continuous[OF \isometry_on X f\]]) using \\n. v n \ X\ \l \ X\ \v \ l\ by auto moreover have "f(v n) = u n" for n unfolding v_def by (simp add: f_inv_into_f u(1)) ultimately have "u \ f l" by auto then show "\m \ f`X. u \ m" using \l \ X\ by auto qed lemma isometry_on_id [simp]: "isometry_on A (\x. x)" "isometry_on A id" unfolding isometry_on_def by auto lemma isometry_on_add [simp]: "isometry_on A (\x. x + (t::'a::real_normed_vector))" unfolding isometry_on_def by auto lemma isometry_on_minus [simp]: "isometry_on A (\(x::'a::real_normed_vector). -x)" unfolding isometry_on_def by (auto simp add: dist_minus) lemma isometry_on_diff [simp]: "isometry_on A (\x. (t::'a::real_normed_vector) - x)" unfolding isometry_on_def by (auto, metis add_uminus_conv_diff dist_add_cancel dist_minus) lemma isometry_preserves_bounded: assumes "isometry_on X f" "A \ X" shows "bounded (f`A) \ bounded A" unfolding bounded_two_points using assms(2) isometry_onD[OF assms(1)] by auto (metis assms(2) rev_subsetD)+ lemma isometry_preserves_infdist: "infdist (f x) (f`A) = infdist x A" if "isometry_on X f" "A \ X" "x \ X" using that by (simp add: infdist_def image_comp isometry_on_def subset_iff) lemma isometry_preserves_hausdorff_distance: "hausdorff_distance (f`A) (f`B) = hausdorff_distance A B" if "isometry_on X f" "A \ X" "B \ X" using that isometry_preserves_infdist [OF that(1) that(2)] isometry_preserves_infdist [OF that(1) that(3)] isometry_preserves_bounded [OF that(1) that(2)] isometry_preserves_bounded [OF that(1) that(3)] by (simp add: hausdorff_distance_def image_comp subset_eq) lemma isometry_on_UNIV_iterates: fixes f::"('a::metric_space) \ 'a" assumes "isometry_on UNIV f" shows "isometry_on UNIV (f^^n)" by (induction n, auto, rule isometry_on_compose[of _ _ f], auto intro: isometry_on_subset[OF assms]) lemma isometry_iterates: fixes f::"('a::metric_space) \ 'a" assumes "isometry f" shows "isometry (f^^n)" using isometry_on_UNIV_iterates[OF isometryD(1)[OF assms], of n] bij_fn[OF isometry_inverse(2)[OF assms], of n] unfolding isometry_def by (simp add: bij_is_surj) section \Geodesic spaces\ text \A geodesic space is a metric space in which any pair of points can be joined by a geodesic segment, i.e., an isometrically embedded copy of a segment in the real line. Most spaces in geometry are geodesic. We introduce in this section the corresponding class of metric spaces. First, we study properties of general geodesic segments in metric spaces.\ subsection \Geodesic segments in general metric spaces\ definition geodesic_segment_between::"('a::metric_space) set \ 'a \ 'a \ bool" where "geodesic_segment_between G x y = (\g::(real \ 'a). g 0 = x \ g (dist x y) = y \ isometry_on {0..dist x y} g \ G = g`{0..dist x y})" definition geodesic_segment::"('a::metric_space) set \ bool" where "geodesic_segment G = (\x y. geodesic_segment_between G x y)" text \We also introduce the parametrization of a geodesic segment. It is convenient to use the following definition, which guarantees that the point is on $G$ even without checking that $G$ is a geodesic segment or that the parameter is in the reasonable range: this shortens some arguments below.\ definition geodesic_segment_param::"('a::metric_space) set \ 'a \ real \ 'a" where "geodesic_segment_param G x t = (if \w. w \ G \ dist x w = t then SOME w. w \ G \ dist x w = t else SOME w. w \ G)" lemma geodesic_segment_betweenI: assumes "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" shows "geodesic_segment_between G x y" unfolding geodesic_segment_between_def apply (rule exI[of _ g]) using assms by auto lemma geodesic_segmentI [intro, simp]: assumes "geodesic_segment_between G x y" shows "geodesic_segment G" unfolding geodesic_segment_def using assms by auto lemma geodesic_segmentI2 [intro]: assumes "isometry_on {a..b} g" "a \ (b::real)" shows "geodesic_segment_between (g`{a..b}) (g a) (g b)" "geodesic_segment (g`{a..b})" proof - define h where "h = (\t. g (t+a))" have *: "isometry_on {0..b-a} h" apply (rule isometry_onI) using \isometry_on {a..b} g\ \a \ b\ by (auto simp add: isometry_on_def h_def) have **: "dist (h 0) (h (b-a)) = b-a" using isometry_onD[OF \isometry_on {0..b-a} h\, of 0 "b-a"] \a \ b\ unfolding dist_real_def by auto have "geodesic_segment_between (h`{0..b-a}) (h 0) (h (b-a))" unfolding geodesic_segment_between_def apply (rule exI[of _ h]) unfolding ** using * by auto moreover have "g`{a..b} = h`{0..b-a}" unfolding h_def apply (auto simp add: image_iff) by (metis add.commute atLeastAtMost_iff diff_ge_0_iff_ge diff_right_mono le_add_diff_inverse) moreover have "h 0 = g a" "h (b-a) = g b" unfolding h_def by auto ultimately show "geodesic_segment_between (g`{a..b}) (g a) (g b)" by auto then show "geodesic_segment (g`{a..b})" unfolding geodesic_segment_def by auto qed lemma geodesic_segmentD: assumes "geodesic_segment_between G x y" shows "\g::(real \ _). (g t = x \ g (t + dist x y) = y \ isometry_on {t..t+dist x y} g \ G = g`{t..t+dist x y})" proof - obtain h where h: "h 0 = x" "h (dist x y) = y" "isometry_on {0..dist x y} h" "G = h`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) have * [simp]: "(\x. x - t) ` {t..t + dist x y} = {0..dist x y}" by auto define g where "g = (\s. h (s - t))" have "g t = x" "g (t + dist x y) = y" using h assms(1) unfolding g_def by auto moreover have "isometry_on {t..t+dist x y} g" unfolding g_def apply (rule isometry_on_compose[of _ _ h]) by (simp add: dist_real_def isometry_on_def, simp add: h(3)) moreover have "g` {t..t + dist x y} = G" unfolding g_def h(4) using * by (metis image_image) ultimately show ?thesis by auto qed lemma geodesic_segment_endpoints [simp]: assumes "geodesic_segment_between G x y" shows "x \ G" "y \ G" "G \ {}" using assms unfolding geodesic_segment_between_def by (auto, metis atLeastAtMost_iff image_eqI less_eq_real_def zero_le_dist) lemma geodesic_segment_commute: assumes "geodesic_segment_between G x y" shows "geodesic_segment_between G y x" proof - obtain g::"real\'a" where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) define h::"real\'a" where "h = (\t. g(dist x y-t))" have "(\t. dist x y -t)`{0..dist x y} = {0..dist x y}" by auto then have "h`{0..dist x y} = G" unfolding g(4) h_def by (metis image_image) moreover have "h 0 = y" "h (dist x y) = x" unfolding h_def using g by auto moreover have "isometry_on {0..dist x y} h" unfolding h_def apply (rule isometry_on_compose[of _ _ g]) using g(3) by auto ultimately show ?thesis unfolding geodesic_segment_between_def by (auto simp add: dist_commute) qed lemma geodesic_segment_dist: assumes "geodesic_segment_between G x y" "a \ G" shows "dist x a + dist a y = dist x y" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain t where t: "t \ {0..dist x y}" "a = g t" using g(4) assms by auto have "dist x a = t" using isometry_onD[OF g(3) _ t(1), of 0] unfolding g(1) dist_real_def t(2) using t(1) by auto moreover have "dist a y = dist x y - t" using isometry_onD[OF g(3) _ t(1), of "dist x y"] unfolding g(2) dist_real_def t(2) using t(1) by (auto simp add: dist_commute) ultimately show ?thesis by auto qed lemma geodesic_segment_dist_unique: assumes "geodesic_segment_between G x y" "a \ G" "b \ G" "dist x a = dist x b" shows "a = b" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain ta where ta: "ta \ {0..dist x y}" "a = g ta" using g(4) assms by auto have *: "dist x a = ta" unfolding g(1)[symmetric] ta(2) using isometry_onD[OF g(3), of 0 ta] unfolding dist_real_def using ta(1) by auto obtain tb where tb: "tb \ {0..dist x y}" "b = g tb" using g(4) assms by auto have "dist x b = tb" unfolding g(1)[symmetric] tb(2) using isometry_onD[OF g(3), of 0 tb] unfolding dist_real_def using tb(1) by auto then have "ta = tb" using * \dist x a = dist x b\ by auto then show "a = b" using ta(2) tb(2) by auto qed lemma geodesic_segment_union: assumes "dist x z = dist x y + dist y z" "geodesic_segment_between G x y" "geodesic_segment_between H y z" shows "geodesic_segment_between (G \ H) x z" "G \ H = {y}" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain h where h: "h (dist x y) = y" "h (dist x z) = z" "isometry_on {dist x y..dist x z} h" "H = h`{dist x y..dist x z}" unfolding \dist x z = dist x y + dist y z\ using geodesic_segmentD[OF \geodesic_segment_between H y z\, of "dist x y"] by auto define f where "f = (\t. if t \ dist x y then g t else h t)" have fg: "f t = g t" if "t \ dist x y" for t unfolding f_def using that by auto have fh: "f t = h t" if "t \ dist x y" for t unfolding f_def apply (cases "t > dist x y") using that g(2) h(1) by auto have "f 0 = x" "f (dist x z) = z" using fg fh g(1) h(2) assms(1) by auto have "f`{0..dist x z} = f`{0..dist x y} \ f`{dist x y..dist x z}" unfolding assms(1) image_Un[symmetric] by (simp add: ivl_disj_un_two_touch(4)) moreover have "f`{0..dist x y} = G" unfolding g(4) using fg image_cong by force moreover have "f`{dist x y..dist x z} = H" unfolding h(4) using fh image_cong by force ultimately have "f`{0..dist x z} = G \ H" by simp have Ifg: "dist (f s) (f t) = s-t" if "0 \ t" "t \ s" "s \ dist x y" for s t using that fg[of s] fg[of t] isometry_onD[OF g(3), of s t] unfolding dist_real_def by auto have Ifh: "dist (f s) (f t) = s-t" if "dist x y \ t" "t \ s" "s \ dist x z" for s t using that fh[of s] fh[of t] isometry_onD[OF h(3), of s t] unfolding dist_real_def by auto have I: "dist (f s) (f t) = s-t" if "0 \ t" "t \ s" "s \ dist x z" for s t proof - consider "t \ dist x y \ s \ dist x y" | "s \ dist x y" | "t \ dist x y" by fastforce then show ?thesis proof (cases) case 1 have "dist (f t) (f s) \ dist (f t) (f (dist x y)) + dist (f (dist x y)) (f s)" using dist_triangle by auto also have "... \ (dist x y - t) + (s - dist x y)" using that 1 Ifg[of t "dist x y"] Ifh[of "dist x y" s] by (auto simp add: dist_commute intro: mono_intros) finally have *: "dist (f t) (f s) \ s - t" by simp have "dist x z \ dist (f 0) (f t) + dist (f t) (f s) + dist (f s) (f (dist x z))" unfolding \f 0 = x\ \f (dist x z) = z\ using dist_triangle4 by auto also have "... \ t + dist (f t) (f s) + (dist x z - s)" using that 1 Ifg[of 0 t] Ifh[of s "dist x z"] by (auto simp add: dist_commute intro: mono_intros) finally have "s - t \ dist (f t) (f s)" by auto then show "dist (f s) (f t) = s-t" using * dist_commute by auto next case 2 then show ?thesis using Ifg that by auto next case 3 then show ?thesis using Ifh that by auto qed qed have "isometry_on {0..dist x z} f" unfolding isometry_on_def dist_real_def using I by (auto, metis abs_of_nonneg dist_commute dist_real_def le_cases zero_le_dist) then show "geodesic_segment_between (G \ H) x z" unfolding geodesic_segment_between_def using \f 0 = x\ \f (dist x z) = z\ \f`{0..dist x z} = G \ H\ by auto have "G \ H \ {y}" proof (auto) fix a assume a: "a \ G" "a \ H" obtain s where s: "s \ {0..dist x y}" "a = g s" using a g(4) by auto obtain t where t: "t \ {dist x y..dist x z}" "a = h t" using a h(4) by auto have "a = f s" using fg s by auto moreover have "a = f t" using fh t by auto ultimately have "s = t" using isometry_onD[OF \isometry_on {0..dist x z} f\, of s t] s(1) t(1) by auto then have "s = dist x y" using s t by auto then show "a = y" using s(2) g by auto qed then show "G \ H = {y}" using assms by auto qed lemma geodesic_segment_dist_le: assumes "geodesic_segment_between G x y" "a \ G" "b \ G" shows "dist a b \ dist x y" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain s t where st: "s \ {0..dist x y}" "t \ {0..dist x y}" "a = g s" "b = g t" using g(4) assms by auto have "dist a b = abs(s-t)" using isometry_onD[OF g(3) st(1) st(2)] unfolding st(3) st(4) dist_real_def by simp then show "dist a b \ dist x y" using st(1) st(2) unfolding dist_real_def by auto qed lemma geodesic_segment_param [simp]: assumes "geodesic_segment_between G x y" shows "geodesic_segment_param G x 0 = x" "geodesic_segment_param G x (dist x y) = y" "t \ {0..dist x y} \ geodesic_segment_param G x t \ G" "isometry_on {0..dist x y} (geodesic_segment_param G x)" "(geodesic_segment_param G x)`{0..dist x y} = G" "t \ {0..dist x y} \ dist x (geodesic_segment_param G x t) = t" "s \ {0..dist x y} \ t \ {0..dist x y} \ dist (geodesic_segment_param G x s) (geodesic_segment_param G x t) = abs(s-t)" "z \ G \ z = geodesic_segment_param G x (dist x z)" proof - obtain g::"real\'a" where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) have *: "g t \ G \ dist x (g t) = t" if "t \ {0..dist x y}" for t using isometry_onD[OF g(3), of 0 t] that g(1) g(4) unfolding dist_real_def by auto have G: "geodesic_segment_param G x t = g t" if "t \ {0..dist x y}" for t proof - have A: "geodesic_segment_param G x t \ G \ dist x (geodesic_segment_param G x t) = t" using *[OF that] unfolding geodesic_segment_param_def apply auto using *[OF that] by (metis (mono_tags, lifting) someI)+ obtain s where s: "geodesic_segment_param G x t = g s" "s \ {0..dist x y}" using A g(4) by auto have "s = t" using *[OF \s \ {0..dist x y}\] A unfolding s(1) by auto then show ?thesis using s by auto qed show "geodesic_segment_param G x 0 = x" "geodesic_segment_param G x (dist x y) = y" "t \ {0..dist x y} \ geodesic_segment_param G x t \ G" "isometry_on {0..dist x y} (geodesic_segment_param G x)" "(geodesic_segment_param G x)`{0..dist x y} = G" "t \ {0..dist x y} \ dist x (geodesic_segment_param G x t) = t" "s \ {0..dist x y} \ t \ {0..dist x y} \ dist (geodesic_segment_param G x s) (geodesic_segment_param G x t) = abs(s-t)" "z \ G \ z = geodesic_segment_param G x (dist x z)" using G g apply (auto simp add: rev_image_eqI) using G isometry_on_cong * atLeastAtMost_iff apply blast using G isometry_on_cong * atLeastAtMost_iff apply blast by (auto simp add: * dist_real_def isometry_onD) qed lemma geodesic_segment_param_in_segment: assumes "G \ {}" shows "geodesic_segment_param G x t \ G" unfolding geodesic_segment_param_def apply (auto, metis (mono_tags, lifting) someI) using assms some_in_eq by fastforce lemma geodesic_segment_reverse_param: assumes "geodesic_segment_between G x y" "t \ {0..dist x y}" shows "geodesic_segment_param G y (dist x y - t) = geodesic_segment_param G x t" proof - have * [simp]: "geodesic_segment_between G y x" using geodesic_segment_commute[OF assms(1)] by simp have "geodesic_segment_param G y (dist x y - t) \ G" apply (rule geodesic_segment_param(3)[of _ _ x]) using assms(2) by (auto simp add: dist_commute) moreover have "dist (geodesic_segment_param G y (dist x y - t)) x = t" using geodesic_segment_param(2)[OF *] geodesic_segment_param(7)[OF *, of "dist x y -t" "dist x y"] assms(2) by (auto simp add: dist_commute) moreover have "geodesic_segment_param G x t \ G" apply (rule geodesic_segment_param(3)[OF assms(1)]) using assms(2) by auto moreover have "dist (geodesic_segment_param G x t) x = t" using geodesic_segment_param(6)[OF assms] by (simp add: dist_commute) ultimately show ?thesis using geodesic_segment_dist_unique[OF assms(1)] by (simp add: dist_commute) qed lemma dist_along_geodesic_wrt_endpoint: assumes "geodesic_segment_between G x y" "u \ G" "v \ G" shows "dist u v = abs(dist u x - dist v x)" proof - have *: "u = geodesic_segment_param G x (dist x u)" "v = geodesic_segment_param G x (dist x v)" using assms by auto have "dist u v = dist (geodesic_segment_param G x (dist x u)) (geodesic_segment_param G x (dist x v))" using * by auto also have "... = abs(dist x u - dist x v)" apply (rule geodesic_segment_param(7)[OF assms(1)]) using assms apply auto using geodesic_segment_dist_le geodesic_segment_endpoints(1) by blast+ finally show ?thesis by (simp add: dist_commute) qed text \One often needs to restrict a geodesic segment to a subsegment. We introduce the tools to express this conveniently.\ definition geodesic_subsegment::"('a::metric_space) set \ 'a \ real \ real \ 'a set" where "geodesic_subsegment G x s t = G \ {z. dist x z \ s \ dist x z \ t}" text \A subsegment is always contained in the original segment.\ lemma geodesic_subsegment_subset: "geodesic_subsegment G x s t \ G" unfolding geodesic_subsegment_def by simp text \A subsegment is indeed a geodesic segment, and its endpoints and parametrization can be expressed in terms of the original segment.\ lemma geodesic_subsegment: assumes "geodesic_segment_between G x y" "0 \ s" "s \ t" "t \ dist x y" shows "geodesic_subsegment G x s t = (geodesic_segment_param G x)`{s..t}" "geodesic_segment_between (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (geodesic_segment_param G x t)" "\u. s \ u \ u \ t \ geodesic_segment_param (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (u - s) = geodesic_segment_param G x u" proof - show A: "geodesic_subsegment G x s t = (geodesic_segment_param G x)`{s..t}" proof (auto) fix y assume y: "y \ geodesic_subsegment G x s t" have "y = geodesic_segment_param G x (dist x y)" apply (rule geodesic_segment_param(8)[OF assms(1)]) using y geodesic_subsegment_subset by force moreover have "dist x y \ s \ dist x y \ t" using y unfolding geodesic_subsegment_def by auto ultimately show "y \ geodesic_segment_param G x ` {s..t}" by auto next fix u assume H: "s \ u" "u \ t" have *: "dist x (geodesic_segment_param G x u) = u" apply (rule geodesic_segment_param(6)[OF assms(1)]) using H assms by auto show "geodesic_segment_param G x u \ geodesic_subsegment G x s t" unfolding geodesic_subsegment_def using geodesic_segment_param_in_segment[OF geodesic_segment_endpoints(3)[OF assms(1)]] by (auto simp add: * H) qed have *: "isometry_on {s..t} (geodesic_segment_param G x)" by (rule isometry_on_subset[of "{0..dist x y}"]) (auto simp add: assms) show B: "geodesic_segment_between (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (geodesic_segment_param G x t)" unfolding A apply (rule geodesic_segmentI2) using * assms by auto fix u assume u: "s \ u" "u \ t" show "geodesic_segment_param (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (u - s) = geodesic_segment_param G x u" proof (rule geodesic_segment_dist_unique[OF B]) show "geodesic_segment_param (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (u - s) \ geodesic_subsegment G x s t" by (rule geodesic_segment_param_in_segment[OF geodesic_segment_endpoints(3)[OF B]]) show "geodesic_segment_param G x u \ geodesic_subsegment G x s t" unfolding A using u by auto have "dist (geodesic_segment_param G x s) (geodesic_segment_param (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (u - s)) = u - s" using B assms u by auto moreover have "dist (geodesic_segment_param G x s) (geodesic_segment_param G x u) = u -s" using assms u by auto ultimately show "dist (geodesic_segment_param G x s) (geodesic_segment_param (geodesic_subsegment G x s t) (geodesic_segment_param G x s) (u - s)) = dist (geodesic_segment_param G x s) (geodesic_segment_param G x u)" by simp qed qed text \The parameterizations of a segment and a subsegment sharing an endpoint coincide where defined.\ lemma geodesic_segment_subparam: assumes "geodesic_segment_between G x z" "geodesic_segment_between H x y" "H \ G" "t \ {0..dist x y}" shows "geodesic_segment_param G x t = geodesic_segment_param H x t" proof - have "geodesic_segment_param H x t \ G" using assms(3) geodesic_segment_param(3)[OF assms(2) assms(4)] by auto then have "geodesic_segment_param H x t = geodesic_segment_param G x (dist x (geodesic_segment_param H x t))" using geodesic_segment_param(8)[OF assms(1)] by auto then show ?thesis using geodesic_segment_param(6)[OF assms(2) assms(4)] by auto qed text \A segment contains a subsegment between any of its points\ lemma geodesic_subsegment_exists: assumes "geodesic_segment G" "x \ G" "y \ G" shows "\H. H \ G \ geodesic_segment_between H x y" proof - obtain a0 b0 where Ga0b0: "geodesic_segment_between G a0 b0" using assms(1) unfolding geodesic_segment_def by auto text \Permuting the endpoints if necessary, we can ensure that the first endpoint $a$ is closer to $x$ than $y$.\ have "\ a b. geodesic_segment_between G a b \ dist x a \ dist y a" proof (cases "dist x a0 \ dist y a0") case True show ?thesis apply (rule exI[of _ a0], rule exI[of _ b0]) using True Ga0b0 by auto next case False show ?thesis apply (rule exI[of _ b0], rule exI[of _ a0]) using Ga0b0 geodesic_segment_commute geodesic_segment_dist[OF Ga0b0 \x \ G\] geodesic_segment_dist[OF Ga0b0 \y \ G\] False by (auto simp add: dist_commute) qed then obtain a b where Gab: "geodesic_segment_between G a b" "dist x a \ dist y a" by auto have *: "0 \ dist x a" "dist x a \ dist y a" "dist y a \ dist a b" using Gab assms by (meson geodesic_segment_dist_le geodesic_segment_endpoints(1) zero_le_dist)+ have **: "x = geodesic_segment_param G a (dist x a)" "y = geodesic_segment_param G a (dist y a)" using Gab \x \ G\ \y \ G\ by (metis dist_commute geodesic_segment_param(8))+ define H where "H = geodesic_subsegment G a (dist x a) (dist y a)" have "H \ G" unfolding H_def by (rule geodesic_subsegment_subset) moreover have "geodesic_segment_between H x y" unfolding H_def using geodesic_subsegment(2)[OF Gab(1) *] ** by auto ultimately show ?thesis by auto qed text \A geodesic segment is homeomorphic to an interval.\ lemma geodesic_segment_homeo_interval: assumes "geodesic_segment_between G x y" shows "{0..dist x y} homeomorphic G" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) show ?thesis using isometry_on_homeomorphism(3)[OF g(3)] unfolding g(4) by simp qed text \Just like an interval, a geodesic segment is compact, connected, path connected, bounded, closed, nonempty, and proper.\ lemma geodesic_segment_topology: assumes "geodesic_segment G" shows "compact G" "connected G" "path_connected G" "bounded G" "closed G" "G \ {}" "proper G" proof - show "compact G" using assms geodesic_segment_homeo_interval homeomorphic_compactness unfolding geodesic_segment_def by force show "path_connected G" using assms is_interval_path_connected geodesic_segment_homeo_interval homeomorphic_path_connectedness unfolding geodesic_segment_def by (metis is_interval_cc) then show "connected G" using path_connected_imp_connected by auto show "bounded G" by (rule compact_imp_bounded, fact) show "closed G" by (rule compact_imp_closed, fact) show "G \ {}" using assms geodesic_segment_def geodesic_segment_endpoints(3) by auto show "proper G" using proper_of_compact \compact G\ by auto qed lemma geodesic_segment_between_x_x [simp]: "geodesic_segment_between {x} x x" "geodesic_segment {x}" "geodesic_segment_between G x x \ G = {x}" proof - show *: "geodesic_segment_between {x} x x" unfolding geodesic_segment_between_def apply (rule exI[of _ "\_. x"]) unfolding isometry_on_def by auto then show "geodesic_segment {x}" by auto show "geodesic_segment_between G x x \ G = {x}" using geodesic_segment_dist_le geodesic_segment_endpoints(2) * by fastforce qed lemma geodesic_segment_disconnection: assumes "geodesic_segment_between G x y" "z \ G" shows "(connected (G - {z})) = (z = x \ z = y)" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain t where t: "t \ {0..dist x y}" "z = g t" using \z \ G\ g(4) by auto have "({0..dist x y} - {t}) homeomorphic (G - {g t})" proof - have *: "isometry_on ({0..dist x y} - {t}) g" apply (rule isometry_on_subset[OF g(3)]) by auto have "({0..dist x y} - {t}) homeomorphic g`({0..dist x y} - {t})" by (rule isometry_on_homeomorphism(3)[OF *]) moreover have "g`({0..dist x y} - {t}) = G - {g t}" unfolding g(4) using isometry_on_injective[OF g(3)] t by (auto simp add: inj_onD) ultimately show ?thesis by auto qed moreover have "connected({0..dist x y} - {t}) = (t = 0 \ t = dist x y)" using t(1) by (auto simp add: connected_iff_interval, fastforce) ultimately have "connected (G - {z}) = (t = 0 \ t = dist x y)" unfolding \z = g t\[symmetric]using homeomorphic_connectedness by blast moreover have "(t = 0 \ t = dist x y) = (z = x \ z = y)" using t g apply auto by (metis atLeastAtMost_iff isometry_on_inverse(2) order_refl zero_le_dist)+ ultimately show ?thesis by auto qed lemma geodesic_segment_unique_endpoints: assumes "geodesic_segment_between G x y" "geodesic_segment_between G a b" shows "{x, y} = {a, b}" by (metis geodesic_segment_disconnection assms(1) assms(2) doubleton_eq_iff geodesic_segment_endpoints(1) geodesic_segment_endpoints(2)) lemma geodesic_segment_subsegment: assumes "geodesic_segment G" "H \ G" "compact H" "connected H" "H \ {}" shows "geodesic_segment H" proof - obtain x y where "geodesic_segment_between G x y" using assms unfolding geodesic_segment_def by auto obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) define L where "L = (inv_into {0..dist x y} g)`H" have "L \ {0..dist x y}" unfolding L_def using isometry_on_inverse[OF \isometry_on {0..dist x y} g\] assms(2) g(4) by auto have "isometry_on G (inv_into {0..dist x y} g)" using isometry_on_inverse[OF \isometry_on {0..dist x y} g\] g(4) by auto then have "isometry_on H (inv_into {0..dist x y} g)" using \H \ G\ isometry_on_subset by auto then have "H homeomorphic L" unfolding L_def using isometry_on_homeomorphism(3) by auto then have "compact L \ connected L" using assms homeomorphic_compactness homeomorphic_connectedness by blast then obtain a b where "L = {a..b}" using connected_compact_interval_1[of L] by auto have "a \ b" using \H \ {}\ \L = {a..b}\ unfolding L_def by auto then have "0 \ a" "b \ dist x y" using \L \ {0..dist x y}\ \L = {a..b}\ by auto have *: "H = g`{a..b}" by (metis L_def \L = {a..b}\ assms(2) g(4) image_inv_into_cancel) show "geodesic_segment H" unfolding * apply (rule geodesic_segmentI2[OF _ \a \ b\]) apply (rule isometry_on_subset[OF g(3)]) using \0 \ a\ \b \ dist x y\ by auto qed text \The image under an isometry of a geodesic segment is still obviously a geodesic segment.\ lemma isometry_preserves_geodesic_segment_between: assumes "isometry_on X f" "G \ X" "geodesic_segment_between G x y" shows "geodesic_segment_between (f`G) (f x) (f y)" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) then have *: "f`G = (f o g) `{0..dist x y}" "f x = (f o g) 0" "f y = (f o g) (dist x y)" by auto show ?thesis unfolding * apply (intro geodesic_segmentI2(1)) unfolding comp_def apply (rule isometry_on_compose[of _ g]) using g(3) g(4) assms by (auto intro: isometry_on_subset) qed text \The sum of distances $d(w, x) + d(w, y)$ can be controlled using the distance from $w$ to a geodesic segment between $x$ and $y$.\ lemma geodesic_segment_distance: assumes "geodesic_segment_between G x y" shows "dist w x + dist w y \ dist x y + 2 * infdist w G" proof - have "\z \ G. infdist w G = dist w z" apply (rule infdist_proper_attained) using assms by (auto simp add: geodesic_segment_topology) then obtain z where z: "z \ G" "infdist w G = dist w z" by auto have "dist w x + dist w y \ (dist w z + dist z x) + (dist w z + dist z y)" by (intro mono_intros) also have "... = dist x z + dist z y + 2 * dist w z" by (auto simp add: dist_commute) also have "... = dist x y + 2 * infdist w G" using z(1) assms geodesic_segment_dist unfolding z(2) by auto finally show ?thesis by auto qed text \If a point $y$ is on a geodesic segment between $x$ and its closest projection $p$ on a set $A$, then $p$ is also a closest projection of $y$, and the closest projection set of $y$ is contained in that of $x$.\ lemma proj_set_geodesic_same_basepoint: assumes "p \ proj_set x A" "geodesic_segment_between G p x" "y \ G" shows "p \ proj_set y A" proof (rule proj_setI) show "p \ A" using assms proj_setD by auto have *: "dist y p \ dist y q" if "q \ A" for q proof - have "dist p y + dist y x = dist p x" using assms geodesic_segment_dist by blast also have "... \ dist q x" using proj_set_dist_le[OF \q \ A\ assms(1)] by (simp add: dist_commute) also have "... \ dist q y + dist y x" by (intro mono_intros) finally show ?thesis by (simp add: dist_commute) qed have "dist y p \ Inf (dist y ` A)" apply (rule cINF_greatest) using \p \ A\ * by auto then show "dist y p \ infdist y A" unfolding infdist_def using \p \ A\ by auto qed lemma proj_set_subset: assumes "p \ proj_set x A" "geodesic_segment_between G p x" "y \ G" shows "proj_set y A \ proj_set x A" proof - have "z \ proj_set x A" if "z \ proj_set y A" for z proof (rule proj_setI) show "z \ A" using that proj_setD by auto have "dist x z \ dist x y + dist y z" by (intro mono_intros) also have "... \ dist x y + dist y p" using proj_set_dist_le[OF proj_setD(1)[OF \p \ proj_set x A\] that] by auto also have "... = dist x p" using assms geodesic_segment_commute geodesic_segment_dist by blast also have "... = infdist x A" using proj_setD(2)[OF assms(1)] by simp finally show "dist x z \ infdist x A" by simp qed then show ?thesis by auto qed lemma proj_set_thickening: assumes "p \ proj_set x Z" "0 \ D" "D \ dist p x" "geodesic_segment_between G p x" shows "geodesic_segment_param G p D \ proj_set x (\z\Z. cball z D)" proof (rule proj_setI') have "dist p (geodesic_segment_param G p D) = D" using geodesic_segment_param(7)[OF assms(4), of 0 D] unfolding geodesic_segment_param(1)[OF assms(4)] using assms by simp then show "geodesic_segment_param G p D \ (\z\Z. cball z D)" using proj_setD(1)[OF \p \ proj_set x Z\] by force show "dist x (geodesic_segment_param G p D) \ dist x y" if "y \ (\z\Z. cball z D)" for y proof - obtain z where y: "y \ cball z D" "z \ Z" using \y \ (\z\Z. cball z D)\ by auto have "dist (geodesic_segment_param G p D) x + D = dist p x" using geodesic_segment_param(7)[OF assms(4), of D "dist p x"] unfolding geodesic_segment_param(2)[OF assms(4)] using assms by simp also have "... \ dist z x" using proj_setD(2)[OF \p \ proj_set x Z\] infdist_le[OF \z \ Z\, of x] by (simp add: dist_commute) also have "... \ dist z y + dist y x" by (intro mono_intros) also have "... \ D + dist y x" using y by simp finally show ?thesis by (simp add: dist_commute) qed qed lemma proj_set_thickening': assumes "p \ proj_set x Z" "0 \ D" "D \ E" "E \ dist p x" "geodesic_segment_between G p x" shows "geodesic_segment_param G p D \ proj_set (geodesic_segment_param G p E) (\z\Z. cball z D)" proof - define H where "H = geodesic_subsegment G p D (dist p x)" have H1: "geodesic_segment_between H (geodesic_segment_param G p D) x" apply (subst geodesic_segment_param(2)[OF \geodesic_segment_between G p x\, symmetric]) unfolding H_def apply (rule geodesic_subsegment(2)) using assms by auto have H2: "geodesic_segment_param G p E \ H" unfolding H_def using assms geodesic_subsegment(1) by force have "geodesic_segment_param G p D \ proj_set x (\z\Z. cball z D)" apply (rule proj_set_thickening) using assms by auto then show ?thesis by (rule proj_set_geodesic_same_basepoint[OF _ H1 H2]) qed text \It is often convenient to use \emph{one} geodesic between $x$ and $y$, even if it is not unique. We introduce a notation for such a choice of a geodesic, denoted \verb+{x--S--y}+ for such a geodesic that moreover remains in the set $S$. We also enforce the condition \verb+{x--S--y} = {y--S--x}+. When there is no such geodesic, we simply take \verb+{x--S--y} = {x, y}+ for definiteness. It would be even better to enforce that, if $a$ is on \verb+{x--S--y}+, then \verb+{x--S--y}+ is the union of \verb+{x--S--a}+ and \verb+{a--S--y}+, but I do not know if such a choice is always possible -- such a choice of geodesics is called a geodesic bicombing. We also write \verb+{x--y}+ for \verb+{x--UNIV--y}+.\ definition some_geodesic_segment_between::"'a::metric_space \ 'a set \ 'a \ 'a set" ("(1{_--_--_})") where "some_geodesic_segment_between = (SOME f. \ x y S. f x S y = f y S x \ (if (\G. geodesic_segment_between G x y \ G \ S) then (geodesic_segment_between (f x S y) x y \ (f x S y \ S)) else f x S y = {x, y}))" abbreviation some_geodesic_segment_between_UNIV::"'a::metric_space \ 'a \ 'a set" ("(1{_--_})") where "some_geodesic_segment_between_UNIV x y \ {x--UNIV--y}" text \We prove that there is such a choice of geodesics, compatible with direction reversal. What we do is choose arbitrarily a geodesic between $x$ and $y$ if it exists, and then use the geodesic between $\min(x, y)$ and $\max(x,y)$, for any total order on the space, to ensure that we get the same result from $x$ to $y$ or from $y$ to $x$.\ lemma some_geodesic_segment_between_exists: "\f. \ x y S. f x S y = f y S x \ (if (\G. geodesic_segment_between G x y \ G \ S) then (geodesic_segment_between (f x S y) x y \ (f x S y \ S)) else f x S y = {x, y})" proof - define g::"'a \ 'a set \ 'a \ 'a set" where "g = (\x S y. if (\G. geodesic_segment_between G x y \ G \ S) then (SOME G. geodesic_segment_between G x y \ G \ S) else {x, y})" have g1: "geodesic_segment_between (g x S y) x y \ (g x S y \ S)" if "\G. geodesic_segment_between G x y \ G \ S" for x y S unfolding g_def using someI_ex[OF that] by auto have g2: "g x S y = {x, y}" if "\(\G. geodesic_segment_between G x y \ G \ S)" for x y S unfolding g_def using that by auto obtain r::"'a rel" where r: "well_order_on UNIV r" using well_order_on by auto have A: "x = y" if "(x, y) \ r" "(y, x) \ r" for x y using r that unfolding well_order_on_def linear_order_on_def partial_order_on_def antisym_def by auto have B: "(x, y) \ r \ (y, x) \ r" for x y using r unfolding well_order_on_def linear_order_on_def total_on_def partial_order_on_def preorder_on_def refl_on_def by force define f where "f = (\x S y. if (x, y) \ r then g x S y else g y S x)" have "f x S y = f y S x" for x y S unfolding f_def using r A B by auto moreover have "geodesic_segment_between (f x S y) x y \ (f x S y \ S)" if "\G. geodesic_segment_between G x y \ G \ S" for x y S unfolding f_def using g1 geodesic_segment_commute that by smt moreover have "f x S y = {x, y}" if "\(\G. geodesic_segment_between G x y \ G \ S)" for x y S unfolding f_def using g2 that geodesic_segment_commute doubleton_eq_iff by metis ultimately show ?thesis by metis qed lemma some_geodesic_commute: "{x--S--y} = {y--S--x}" unfolding some_geodesic_segment_between_def by (auto simp add: someI_ex[OF some_geodesic_segment_between_exists]) lemma some_geodesic_segment_description: "(\G. geodesic_segment_between G x y \ G \ S) \ geodesic_segment_between {x--S--y} x y" "(\(\G. geodesic_segment_between G x y \ G \ S)) \ {x--S--y} = {x, y}" unfolding some_geodesic_segment_between_def by (simp add: someI_ex[OF some_geodesic_segment_between_exists])+ text \Basic topological properties of our chosen set of geodesics.\ lemma some_geodesic_compact [simp]: "compact {x--S--y}" apply (cases "\G. geodesic_segment_between G x y \ G \ S") using some_geodesic_segment_description[of x y] geodesic_segment_topology[of "{x--S--y}"] geodesic_segment_def apply auto by blast lemma some_geodesic_closed [simp]: "closed {x--S--y}" by (rule compact_imp_closed[OF some_geodesic_compact[of x S y]]) lemma some_geodesic_bounded [simp]: "bounded {x--S--y}" by (rule compact_imp_bounded[OF some_geodesic_compact[of x S y]]) lemma some_geodesic_endpoints [simp]: "x \ {x--S--y}" "y \ {x--S--y}" "{x--S--y} \ {}" apply (cases "\G. geodesic_segment_between G x y \ G \ S") using some_geodesic_segment_description[of x y S] apply auto apply (cases "\G. geodesic_segment_between G x y \ G \ S") using some_geodesic_segment_description[of x y S] apply auto apply (cases "\G. geodesic_segment_between G x y \ G \ S") using geodesic_segment_endpoints(3) by (auto, blast) lemma some_geodesic_subsegment: assumes "H \ {x--S--y}" "compact H" "connected H" "H \ {}" shows "geodesic_segment H" apply (cases "\G. geodesic_segment_between G x y \ G \ S") using some_geodesic_segment_description[of x y] geodesic_segment_subsegment[OF _ assms] geodesic_segment_def apply auto[1] using some_geodesic_segment_description[of x y] assms by (metis connected_finite_iff_sing finite.emptyI finite.insertI finite_subset geodesic_segment_between_x_x(2)) lemma some_geodesic_in_subset: assumes "x \ S" "y \ S" shows "{x--S--y} \ S" apply (cases "\G. geodesic_segment_between G x y \ G \ S") unfolding some_geodesic_segment_between_def by (simp add: assms someI_ex[OF some_geodesic_segment_between_exists])+ lemma some_geodesic_same_endpoints [simp]: "{x--S--x} = {x}" apply (cases "\G. geodesic_segment_between G x x \ G \ S") apply (meson geodesic_segment_between_x_x(3) some_geodesic_segment_description(1)) by (simp add: some_geodesic_segment_description(2)) subsection \Geodesic subsets\ text \A subset is \emph{geodesic} if any two of its points can be joined by a geodesic segment. We prove basic properties of such a subset in this paragraph -- notably connectedness. A basic example is given by convex subsets of vector spaces, as closed segments are geodesic.\ definition geodesic_subset::"('a::metric_space) set \ bool" where "geodesic_subset S = (\x\S. \y\S. \G. geodesic_segment_between G x y \ G \ S)" lemma geodesic_subsetD: assumes "geodesic_subset S" "x \ S" "y \ S" shows "geodesic_segment_between {x--S--y} x y" using assms some_geodesic_segment_description(1) unfolding geodesic_subset_def by blast lemma geodesic_subsetI: assumes "\x y. x \ S \ y \ S \ \G. geodesic_segment_between G x y \ G \ S" shows "geodesic_subset S" using assms unfolding geodesic_subset_def by auto lemma geodesic_subset_empty: "geodesic_subset {}" using geodesic_subsetI by auto lemma geodesic_subset_singleton: "geodesic_subset {x}" by (auto intro!: geodesic_subsetI geodesic_segment_between_x_x(1)) lemma geodesic_subset_path_connected: assumes "geodesic_subset S" shows "path_connected S" proof - have "\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" if "x \ S" "y \ S" for x y proof - define G where "G = {x--S--y}" have *: "geodesic_segment_between G x y" "G \ S" "x \ G" "y \ G" using assms that by (auto simp add: G_def geodesic_subsetD some_geodesic_in_subset that(1) that(2)) then have "path_connected G" using geodesic_segment_topology(3) unfolding geodesic_segment_def by auto then have "\g. path g \ path_image g \ G \ pathstart g = x \ pathfinish g = y" using * unfolding path_connected_def by auto then show ?thesis using \G \ S\ by auto qed then show ?thesis unfolding path_connected_def by auto qed text \To show that a segment in a normed vector space is geodesic, we will need to use its length parametrization, which is given in the next lemma.\ lemma closed_segment_as_isometric_image: "((\t. x + (t/dist x y) *\<^sub>R (y - x))`{0..dist x y}) = closed_segment x y" proof (auto simp add: closed_segment_def image_iff) fix t assume H: "0 \ t" "t \ dist x y" show "\u. x + (t / dist x y) *\<^sub>R (y - x) = (1 - u) *\<^sub>R x + u *\<^sub>R y \ 0 \ u \ u \ 1" - apply (rule exI[of _ "t/dist x y"]) using H apply (auto simp add: algebra_simps divide_simps) - by (metis add_diff_cancel_left' add_divide_distrib dist_not_less_zero dist_pos_lt divide_eq_1_iff linordered_field_class.sign_simps(29) scaleR_left.add scaleR_one) + apply (rule exI[of _ "t/dist x y"]) + using H apply (auto simp add: algebra_simps divide_simps) + apply (metis add_diff_cancel_left' add_diff_eq add_divide_distrib dist_eq_0_iff scaleR_add_left vector_fraction_eq_iff) + done next fix u::real assume H: "0 \ u" "u \ 1" show "\t\{0..dist x y}. (1 - u) *\<^sub>R x + u *\<^sub>R y = x + (t / dist x y) *\<^sub>R (y - x)" apply (rule bexI[of _ "u * dist x y"]) using H by (auto simp add: algebra_simps mult_left_le_one_le) qed proposition closed_segment_is_geodesic: fixes x y::"'a::real_normed_vector" shows "isometry_on {0..dist x y} (\t. x + (t/dist x y) *\<^sub>R (y - x))" "geodesic_segment_between (closed_segment x y) x y" "geodesic_segment (closed_segment x y)" proof - show *: "isometry_on {0..dist x y} (\t. x + (t/dist x y) *\<^sub>R (y - x))" unfolding isometry_on_def dist_norm apply (cases "x = y") by (auto simp add: scaleR_diff_left[symmetric] diff_divide_distrib[symmetric] norm_minus_commute) show "geodesic_segment_between (closed_segment x y) x y" unfolding closed_segment_as_isometric_image[symmetric] apply (rule geodesic_segment_betweenI[OF _ _ *]) by auto then show "geodesic_segment (closed_segment x y)" by auto qed text \We deduce that a convex set is geodesic.\ proposition convex_is_geodesic: assumes "convex (S::'a::real_normed_vector set)" shows "geodesic_subset S" proof (rule geodesic_subsetI) fix x y assume H: "x \ S" "y \ S" show "\G. geodesic_segment_between G x y \ G \ S" apply (rule exI[of _ "closed_segment x y"]) apply (auto simp add: closed_segment_is_geodesic) using H assms convex_contains_segment by blast qed subsection \Geodesic spaces\ text \In this subsection, we define geodesic spaces (metric spaces in which there is a geodesic segment joining any pair of points). We specialize the previous statements on geodesic segments to these situations.\ class geodesic_space = metric_space + assumes geodesic: "geodesic_subset (UNIV::('a::metric_space) set)" text \The simplest example of a geodesic space is a real normed vector space. Significant examples also include graphs (with the graph distance), Riemannian manifolds, and $CAT(\kappa)$ spaces.\ instance real_normed_vector \ geodesic_space by (standard, simp add: convex_is_geodesic) lemma (in geodesic_space) some_geodesic_is_geodesic_segment [simp]: "geodesic_segment_between {x--y} x (y::'a)" "geodesic_segment {x--y}" using some_geodesic_segment_description(1)[of x y] geodesic_subsetD[OF geodesic] by (auto, blast) lemma (in geodesic_space) some_geodesic_connected [simp]: "connected {x--y}" "path_connected {x--y}" by (auto intro!: geodesic_segment_topology) text \In geodesic spaces, we restate as simp rules all properties of the geodesic segment parametrizations.\ lemma (in geodesic_space) geodesic_segment_param_in_geodesic_spaces [simp]: "geodesic_segment_param {x--y} x 0 = x" "geodesic_segment_param {x--y} x (dist x y) = y" "t \ {0..dist x y} \ geodesic_segment_param {x--y} x t \ {x--y}" "isometry_on {0..dist x y} (geodesic_segment_param {x--y} x)" "(geodesic_segment_param {x--y} x)`{0..dist x y} = {x--y}" "t \ {0..dist x y} \ dist x (geodesic_segment_param {x--y} x t) = t" "s \ {0..dist x y} \ t \ {0..dist x y} \ dist (geodesic_segment_param {x--y} x s) (geodesic_segment_param {x--y} x t) = abs(s-t)" "z \ {x--y} \ z = geodesic_segment_param {x--y} x (dist x z)" using geodesic_segment_param[OF some_geodesic_is_geodesic_segment(1)[of x y]] by auto subsection \Uniquely geodesic spaces\ text \In this subsection, we define uniquely geodesic spaces, i.e., geodesic spaces in which, additionally, there is a unique geodesic between any pair of points.\ class uniquely_geodesic_space = geodesic_space + assumes uniquely_geodesic: "\x y G H. geodesic_segment_between G x y \ geodesic_segment_between H x y \ G = H" text \To prove that a geodesic space is uniquely geodesic, it suffices to show that there is no loop, i.e., if two geodesic segments intersect only at their endpoints, then they coincide. Indeed, assume this holds, and consider two geodesics with the same endpoints. If they differ at some time $t$, then consider the last time $a$ before $t$ where they coincide, and the first time $b$ after $t$ where they coincide. Then the restrictions of the two geodesics to $[a,b]$ give a loop, and a contradiction.\ lemma (in geodesic_space) uniquely_geodesic_spaceI: assumes "\G H x (y::'a). geodesic_segment_between G x y \ geodesic_segment_between H x y \ G \ H = {x, y} \ x = y" "geodesic_segment_between G x y" "geodesic_segment_between H x (y::'a)" shows "G = H" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) obtain h where h: "h 0 = x" "h (dist x y) = y" "isometry_on {0..dist x y} h" "H = h`{0..dist x y}" by (meson \geodesic_segment_between H x y\ geodesic_segment_between_def) have "g t = h t" if "t \ {0..dist x y}" for t proof (rule ccontr) assume "g t \ h t" define Z where "Z = {s \ {0..dist x y}. g s = h s}" have "0 \ Z" "dist x y \ Z" unfolding Z_def using g h by auto have "t \ Z" unfolding Z_def using \g t \ h t\ by auto have [simp]: "closed Z" proof - have *: "Z = (\s. dist (g s) (h s))-`{0} \ {0..dist x y}" unfolding Z_def by auto show ?thesis unfolding * apply (rule closed_vimage_Int) using isometry_on_continuous[OF g(3)] isometry_on_continuous[OF h(3)] continuous_on_dist by auto qed define a where "a = Sup (Z \ {0..t})" have a: "a \ Z \ {0..t}" unfolding a_def apply (rule closed_contains_Sup, auto) using \0 \ Z\ that by auto then have "h a = g a" unfolding Z_def by auto define b where "b = Inf (Z \ {t..dist x y})" have b: "b \ Z \ {t..dist x y}" unfolding b_def apply (rule closed_contains_Inf, auto) using \dist x y \ Z\ that by auto then have "h b = g b" unfolding Z_def by auto have notZ: "s \ Z" if "s \ {a<.. t") case True assume "s \ Z" then have *: "s \ Z \ {0..t}" using that a True by auto have "s \ a" unfolding a_def apply (rule cSup_upper) using * by auto then show False using that by auto next case False assume "s \ Z" then have *: "s \ Z \ {t..dist x y}" using that b False by auto have "s \ b" unfolding b_def apply (rule cInf_lower) using * by auto then show False using that by auto qed have "t \ {a<..t \ Z\ less_eq_real_def by auto then have "a \ b" by auto then have "dist (h a) (h b) = b-a" using isometry_onD[OF h(3), of a b] a b that unfolding dist_real_def by auto then have "dist (h a) (h b) > 0" using \t \ {a<.. by auto then have "h a \ h b" by auto define G2 where "G2 = g`{a..b}" define H2 where "H2 = h`{a..b}" have "G2 \ H2 \ {h a, h b}" proof fix z assume z: "z \ G2 \ H2" obtain sg where sg: "z = g sg" "sg \ {a..b}" using z unfolding G2_def by auto obtain sh where sh: "z = h sh" "sh \ {a..b}" using z unfolding H2_def by auto have "sg = dist x z" using isometry_onD[OF g(3), of 0 sg] a b sg(2) unfolding sg(1) g(1)[symmetric] dist_real_def by auto moreover have "sh = dist x z" using isometry_onD[OF h(3), of 0 sh] a b sh(2) unfolding sh(1) h(1)[symmetric] dist_real_def by auto ultimately have "sg = sh" by auto then have "sh \ Z" using sg(1) sh(1) a b sh(2) unfolding Z_def by auto then have "sh \ {a, b}" using notZ sh(2) by (metis IntD2 atLeastAtMost_iff atLeastAtMost_singleton greaterThanLessThan_iff inf_bot_left insertI2 insert_inter_insert not_le) then show "z \ {h a, h b}" using sh(1) by auto qed then have "G2 \ H2 = {h a, h b}" using \h a = g a\ \h b = g b\ \a \ b\ unfolding H2_def G2_def apply auto unfolding \h a = g a\[symmetric] \h b = g b\[symmetric] by auto moreover have "geodesic_segment_between G2 (h a) (h b)" unfolding G2_def \h a = g a\ \h b = g b\ apply (rule geodesic_segmentI2) apply (rule isometry_on_subset[OF g(3)]) using a b that by auto moreover have "geodesic_segment_between H2 (h a) (h b)" unfolding H2_def apply (rule geodesic_segmentI2) apply (rule isometry_on_subset[OF h(3)]) using a b that by auto ultimately have "h a = h b" using assms(1) by auto then show False using \h a \ h b\ by simp qed then show "G = H" using g(4) h(4) by (simp add: image_def) qed context uniquely_geodesic_space begin lemma geodesic_segment_unique: "geodesic_segment_between G x y = (G = {x--(y::'a)})" using uniquely_geodesic[of _ x y] by (meson some_geodesic_is_geodesic_segment) lemma geodesic_segment_dist': assumes "dist x z = dist x y + dist y z" shows "y \ {x--z}" "{x--z} = {x--y} \ {y--z}" proof - have "geodesic_segment_between ({x--y} \ {y--z}) x z" using geodesic_segment_union[OF assms] by auto then show "{x--z} = {x--y} \ {y--z}" using geodesic_segment_unique by auto then show "y \ {x--z}" by auto qed lemma geodesic_segment_expression: "{x--z} = {y. dist x z = dist x y + dist y z}" using geodesic_segment_dist'(1) geodesic_segment_dist[OF some_geodesic_is_geodesic_segment(1)] by auto lemma geodesic_segment_split: assumes "(y::'a) \ {x--z}" shows "{x--z} = {x--y} \ {y--z}" "{x--y} \ {y--z} = {y}" apply (metis assms geodesic_segment_dist geodesic_segment_dist'(2) some_geodesic_is_geodesic_segment(1)) apply (rule geodesic_segment_union(2)[of x z], auto simp add: assms) using assms geodesic_segment_expression by blast lemma geodesic_segment_subparam': assumes "y \ {x--z}" "t \ {0..dist x y}" shows "geodesic_segment_param {x--z} x t = geodesic_segment_param {x--y} x t" apply (rule geodesic_segment_subparam[of _ _ z _ y]) using assms apply auto using geodesic_segment_split(1)[OF assms(1)] by auto end (*of context uniquely_geodesic_space*) subsection \A complete metric space with middles is geodesic.\ text \A complete space in which every pair of points has a middle (i.e., a point $m$ which is half distance of $x$ and $y$) is geodesic: to construct a geodesic between $x_0$ and $y_0$, first choose a middle $m$, then middles of the pairs $(x_0,m)$ and $(m, y_0)$, and so on. This will define the geodesic on dyadic points (and this is indeed an isometry on these dyadic points. Then, extend it by uniform continuity to the whole segment $[0, dist x0 y0]$. The formal proof will be done in a locale where $x_0$ and $y_0$ are fixed, for notational simplicity. We define inductively the sequence of middles, in a function \verb+geod+ of two natural variables: $geod n m$ corresponds to the image of the dyadic point $m/2^n$. It is defined inductively, by $geod (n+1) (2m) = geod n m$, and $geod (n+1) (2m+1)$ is a middle of $geod n m$ and $geod n (m+1)$. This is not a completely classical inductive definition, so one has to use \verb+function+ to define it. Then, one checks inductively that it has all the properties we want, and use it to define the geodesic segment on dyadic points. We will not use a canonical representative for a dyadic point, but any representative (i.e., numerator and denominator will not have to be coprime) -- this will not create problems as $geod$ does not depend on the choice of the representative, by construction.\ locale complete_space_with_middle = fixes x0 y0::"'a::complete_space" assumes middles: "\x y::'a. \z. dist x z = (dist x y)/2 \ dist z y = (dist x y)/2" begin definition middle::"'a \ 'a \ 'a" where "middle x y = (SOME z. dist x z = (dist x y)/2 \ dist z y = (dist x y)/2)" lemma middle: "dist x (middle x y) = (dist x y)/2" "dist (middle x y) y = (dist x y)/2" unfolding middle_def using middles[of x y] by (metis (mono_tags, lifting) someI_ex)+ function geod::"nat \ nat \ 'a" where "geod 0 0 = x0" |"geod 0 (Suc m) = y0" |"geod (Suc n) (2 * m) = geod n m" |"geod (Suc n) (Suc (2*m)) = middle (geod n m) (geod n (Suc m))" apply (auto simp add: double_not_eq_Suc_double) by (metis One_nat_def dvd_mult_div_cancel list_decode.cases odd_Suc_minus_one odd_two_times_div_two_nat) termination by lexicographic_order text \By induction, the distance between successive points is $D/2^n$.\ lemma geod_distance_successor: "\a < 2^n. dist (geod n a) (geod n (Suc a)) = dist x0 y0 / 2^n" proof (induction n) case 0 show ?case by auto next case (Suc n) show ?case proof (auto) fix a::nat assume a: "a < 2 * 2^n" obtain m where m: "a = 2 * m \ a = Suc (2 * m)" by (metis geod.elims) then have "m < 2^n" using a by auto consider "a = 2 * m" | "a = Suc(2*m)" using m by auto then show "dist (geod (Suc n) a) (geod (Suc n) (Suc a)) = dist x0 y0 / (2 * 2 ^ n)" proof (cases) case 1 show ?thesis unfolding 1 apply auto unfolding middle using Suc.IH \m < 2^n\ by auto next case 2 have *: "Suc (Suc (2 * m)) = 2 * (Suc m)" by auto show ?thesis unfolding 2 apply auto unfolding * geod.simps(3) middle using Suc.IH \m < 2^n\ by auto qed qed qed lemma geod_mult: "geod n a = geod (n + k) (a * 2^k)" apply (induction k, auto) using geod.simps(3) by (metis mult.left_commute) lemma geod_0: "geod n 0 = x0" by (induction n, auto, metis geod.simps(3) semiring_normalization_rules(10)) lemma geod_end: "geod n (2^n) = y0" by (induction n, auto) text \By the triangular inequality, the distance between points separated by $(b-a)/2^n$ is at most $D * (b-a)/2^n$.\ lemma geod_upper: assumes "a \ b" "b \ 2^n" shows "dist (geod n a) (geod n b) \ (b-a) * dist x0 y0 / 2^n" proof - have *: "a+k > 2^n \ dist (geod n a) (geod n (a+k)) \ k * dist x0 y0 / 2^n" for k proof (induction k) case 0 then show ?case by auto next case (Suc k) show ?case proof (cases "2 ^ n < a + Suc k") case True then show ?thesis by auto next case False then have *: "a + k < 2 ^ n" by auto have "dist (geod n a) (geod n (a + Suc k)) \ dist (geod n a) (geod n (a+k)) + dist (geod n (a+k)) (geod n (a+Suc k))" using dist_triangle by auto also have "... \ k * dist x0 y0 / 2^n + dist x0 y0 / 2^n" using Suc.IH * geod_distance_successor by auto finally show ?thesis by (simp add: add_divide_distrib distrib_left mult.commute) qed qed show ?thesis using *[of "b-a"] assms by (simp add: of_nat_diff) qed text \In fact, the distance is exactly $D * (b-a)/2^n$, otherwise the extremities of the interval would be closer than $D$, a contradiction.\ lemma geod_dist: assumes "a \ b" "b \ 2^n" shows "dist (geod n a) (geod n b) = (b-a) * dist x0 y0 / 2^n" proof - have "dist (geod n a) (geod n b) \ (real b-a) * dist x0 y0 / 2^n" using geod_upper[of a b n] assms by auto moreover have "\ (dist (geod n a) (geod n b) < (real b-a) * dist x0 y0 / 2^n)" proof (rule ccontr, simp) assume *: "dist (geod n a) (geod n b) < (real b-a) * dist x0 y0 / 2^n" have "dist x0 y0 = dist (geod n 0) (geod n (2^n))" using geod_0 geod_end by auto also have "... \ dist (geod n 0) (geod n a) + dist (geod n a) (geod n b) + dist (geod n b) (geod n (2^n))" using dist_triangle4 by auto also have "... < a * dist x0 y0 / 2^n + (real b-a) * dist x0 y0 / 2^n + (2^n - real b) * dist x0 y0 / 2^n" using * assms geod_upper[of 0 a n] geod_upper[of b "2^n" n] by (auto intro: mono_intros) also have "... = dist x0 y0" using assms by (auto simp add: algebra_simps divide_simps) finally show "False" by auto qed ultimately show ?thesis by auto qed text \We deduce the same statement but for points that are not on the same level, by putting them on a common multiple level.\ lemma geod_dist2: assumes "a \ 2^n" "b \ 2^p" "a/2^n \ b / 2^p" shows "dist (geod n a) (geod p b) = (b/2^p - a/2^n) * dist x0 y0" proof - define r where "r = max n p" define ar where "ar = a * 2^(r - n)" have a: "ar / 2^r = a / 2^n" unfolding ar_def r_def by (auto simp add: divide_simps semiring_normalization_rules(26)) have A: "geod r ar = geod n a" unfolding ar_def r_def using geod_mult[of n a "max n p - n"] by auto define br where "br = b * 2^(r - p)" have b: "br / 2^r = b / 2^p" unfolding br_def r_def by (auto simp add: divide_simps semiring_normalization_rules(26)) have B: "geod r br = geod p b" unfolding br_def r_def using geod_mult[of p b "max n p - p"] by auto have "dist (geod n a) (geod p b) = dist (geod r ar) (geod r br)" using A B by auto also have "... = (real br - ar) * dist x0 y0 / 2 ^r" apply (rule geod_dist) using \a/2^n \ b / 2^p\ unfolding a[symmetric] b[symmetric] apply (auto simp add: divide_simps) using \b \ 2^p\ b apply (auto simp add: divide_simps) by (metis br_def le_add_diff_inverse2 max.cobounded2 mult.commute mult_le_mono2 r_def semiring_normalization_rules(26)) also have "... = (real br / 2^r - real ar / 2^r) * dist x0 y0" by (auto simp add: algebra_simps divide_simps) finally show ?thesis using a b by auto qed text \Same thing but without a priori ordering of the points.\ lemma geod_dist3: assumes "a \ 2^n" "b \ 2^p" shows "dist (geod n a) (geod p b) = abs(b/2^p - a/2^n) * dist x0 y0" apply (cases "a /2^n \ b/2^p", auto) apply (rule geod_dist2[OF assms], auto) apply (subst dist_commute, rule geod_dist2[OF assms(2) assms(1)], auto) done text \Finally, we define a geodesic by extending what we have already defined on dyadic points, thanks to the result of isometric extension of isometries taking their values in complete spaces.\ lemma geod: shows "\g. isometry_on {0..dist x0 y0} g \ g 0 = x0 \ g (dist x0 y0) = y0" proof (cases "x0 = y0") case True show ?thesis apply (rule exI[of _ "\_. x0"]) unfolding isometry_on_def using True by auto next case False define A where "A = {(real k/2^n) * dist x0 y0 |k n. k \ 2^n}" have "{0..dist x0 y0} \ closure A" proof (auto simp add: closure_approachable dist_real_def) fix t::real assume t: "0 \ t" "t \ dist x0 y0" fix e:: real assume "e > 0" then obtain n::nat where n: "dist x0 y0/e < 2^n" using one_less_numeral_iff real_arch_pow semiring_norm(76) by blast define k where "k = floor (2^n * t/ dist x0 y0)" have "k \ 2^n * t/ dist x0 y0" unfolding k_def by auto also have "... \ 2^n" using t False by (auto simp add: algebra_simps divide_simps) finally have "k \ 2^n" by auto have "k \ 0" using t False unfolding k_def by auto define l where "l = nat k" have "k = int l" "l \ 2^n" using \k \ 0\ \k \ 2^n\ nat_le_iff unfolding l_def by auto have "abs (2^n * t/dist x0 y0 - k) \ 1" unfolding k_def by linarith then have "abs(t - k/2^n * dist x0 y0) \ dist x0 y0 / 2^n" by (auto simp add: algebra_simps divide_simps False) also have "... < e" using n \e > 0\ by (auto simp add: algebra_simps divide_simps) finally have "abs(t - k/2^n * dist x0 y0) < e" by auto then have "abs(t - l/2^n * dist x0 y0) < e" using \k = int l\ by auto moreover have "l/2^n * dist x0 y0 \ A" unfolding A_def using \l \ 2^n\ by auto ultimately show "\u\A. abs(u - t) < e" by force qed text \For each dyadic point, we choose one representation of the form $K/2^N$, it is not important for us that it is the minimal one.\ define index where "index = (\t. SOME i. t = real (fst i)/2^(snd i) * dist x0 y0 \ (fst i) \ 2^(snd i))" define K where "K = (\t. fst (index t))" define N where "N = (\t. snd (index t))" have t: "t = K t/ 2^(N t) * dist x0 y0 \ K t \ 2^(N t)" if "t \ A" for t proof - obtain n k::nat where "t = k/2^n * dist x0 y0" "k \ 2^n" using \t\ A\ unfolding A_def by auto then have *: "\i. t = real (fst i)/2^(snd i) * dist x0 y0 \ (fst i) \ 2^(snd i)" by auto show ?thesis unfolding K_def N_def index_def using someI_ex[OF *] by auto qed text \We can now define our function on dyadic points.\ define f where "f = (\t. geod (N t) (K t))" have "0 \ A" unfolding A_def by auto have "f 0 = x0" proof - have "0 = K 0 /2^(N 0) * dist x0 y0" using t \0 \ A\ by auto then have "K 0 = 0" using False by auto then show ?thesis unfolding f_def using geod_0 by auto qed have "dist x0 y0 = (real 1/2^0) * dist x0 y0" by auto then have "dist x0 y0 \ A" unfolding A_def by force have "f (dist x0 y0) = y0" proof - have "dist x0 y0 = K (dist x0 y0) / 2^(N (dist x0 y0)) * dist x0 y0" using t \dist x0 y0 \ A\ by auto then have "K (dist x0 y0) = 2^(N(dist x0 y0))" using False by (auto simp add: divide_simps) then show ?thesis unfolding f_def using geod_end by auto qed text \By construction, it is an isometry on dyadic points.\ have "isometry_on A f" proof (rule isometry_onI) fix s t assume inA: "s \ A" "t \ A" have "dist (f s) (f t) = abs (K t/2^(N t) - K s/2^(N s)) * dist x0 y0" unfolding f_def apply (rule geod_dist3) using t inA by auto also have "... = abs(K t/2^(N t) * dist x0 y0 - K s/2^(N s) * dist x0 y0)" by (auto simp add: abs_mult_pos left_diff_distrib) also have "... = abs(t - s)" using t inA by auto finally show "dist (f s) (f t) = dist s t" unfolding dist_real_def by auto qed text \We can thus extend it to an isometry on the closure of dyadic points. It is the desired geodesic.\ then obtain g where g: "isometry_on (closure A) g" "\t. t \ A \ g t = f t" using isometry_extend_closure by metis have "isometry_on {0..dist x0 y0} g" by (rule isometry_on_subset[OF \isometry_on (closure A) g\ \{0..dist x0 y0} \ closure A\]) moreover have "g 0 = x0" using g(2)[OF \0 \ A\] \f 0 = x0\ by simp moreover have "g (dist x0 y0) = y0" using g(2)[OF \dist x0 y0 \ A\] \f (dist x0 y0) = y0\ by simp ultimately show ?thesis by auto qed end text \We can now complete the proof that a complete space with middles is in fact geodesic: all the work has been done in the locale \verb+complete_space_with_middle+, in Lemma~\verb+geod+.\ theorem complete_with_middles_imp_geodesic: assumes "\x y::('a::complete_space). \m. dist x m = dist x y /2 \ dist m y = dist x y /2" shows "OFCLASS('a, geodesic_space_class)" proof (standard, rule geodesic_subsetI) fix x0 y0::'a interpret complete_space_with_middle x0 y0 apply standard using assms by auto have "\g. g 0 = x0 \ g (dist x0 y0) = y0 \ isometry_on {0..dist x0 y0} g" using geod by auto then show "\G. geodesic_segment_between G x0 y0 \ G \ UNIV" unfolding geodesic_segment_between_def by auto qed section \Quasi-isometries\ text \A $(\lambda, C)$ quasi-isometry is a function which behaves like an isometry, up to an additive error $C$ and a multiplicative error $\lambda$. It can be very different from an isometry on small scales (for instance, the function integer part is a quasi-isometry between $\mathbb{R}$ and $\mathbb{Z}$), but on large scales it captures many important features of isometries. When the space is unbounded, one checks easily that $C \geq 0$ and $\lambda \geq 1$. As this is the only case of interest (any two bounded sets are quasi-isometric), we incorporate this requirement in the definition.\ definition quasi_isometry_on::"real \ real \ ('a::metric_space) set \ ('a \ ('b::metric_space)) \ bool" ("_ _ -quasi'_isometry'_on" [1000, 999]) where "lambda C-quasi_isometry_on X f = ((lambda \ 1) \ (C \ 0) \ (\x \ X. \y \ X. (dist (f x) (f y) \ lambda * dist x y + C \ dist (f x) (f y) \ (1/lambda) * dist x y - C)))" abbreviation quasi_isometry :: "real \ real \ ('a::metric_space \ 'b::metric_space) \ bool" ("_ _ -quasi'_isometry" [1000, 999]) where "quasi_isometry lambda C f \ lambda C-quasi_isometry_on UNIV f" subsection \Basic properties of quasi-isometries\ lemma quasi_isometry_onD: assumes "lambda C-quasi_isometry_on X f" shows "\x y. x \ X \ y \ X \ dist (f x) (f y) \ lambda * dist x y + C" "\x y. x \ X \ y \ X \ dist (f x) (f y) \ (1/lambda) * dist x y - C" "lambda \ 1" "C \ 0" using assms unfolding quasi_isometry_on_def by auto lemma quasi_isometry_onI [intro]: assumes "\x y. x \ X \ y \ X \ dist (f x) (f y) \ lambda * dist x y + C" "\x y. x \ X \ y \ X \ dist (f x) (f y) \ (1/lambda) * dist x y - C" "lambda \ 1" "C \ 0" shows "lambda C-quasi_isometry_on X f" using assms unfolding quasi_isometry_on_def by auto lemma isometry_quasi_isometry_on: assumes "isometry_on X f" shows "1 0-quasi_isometry_on X f" using assms unfolding isometry_on_def quasi_isometry_on_def by auto lemma quasi_isometry_on_change_params: assumes "lambda C-quasi_isometry_on X f" "mu \ lambda" "D \ C" shows "mu D-quasi_isometry_on X f" proof (rule quasi_isometry_onI) have P1: "lambda \ 1" "C \ 0" using quasi_isometry_onD[OF assms(1)] by auto then show P2: "mu \ 1" "D \ 0" using assms by auto fix x y assume inX: "x \ X" "y \ X" have "dist (f x) (f y) \ lambda * dist x y + C" using quasi_isometry_onD[OF assms(1)] inX by auto also have "... \ mu * dist x y + D" using assms by (auto intro!: mono_intros) finally show "dist (f x) (f y) \ mu * dist x y + D" by simp have "dist (f x) (f y) \ (1/lambda) * dist x y - C" using quasi_isometry_onD[OF assms(1)] inX by auto moreover have "(1/lambda) * dist x y + (- C) \ (1/mu) * dist x y + (- D)" apply (intro mono_intros) using P1 P2 assms by (auto simp add: divide_simps) ultimately show "dist (f x) (f y) \ (1/mu) * dist x y - D" by simp qed lemma quasi_isometry_on_subset: assumes "lambda C-quasi_isometry_on X f" "Y \ X" shows "lambda C-quasi_isometry_on Y f" using assms unfolding quasi_isometry_on_def by auto lemma quasi_isometry_on_perturb: assumes "lambda C-quasi_isometry_on X f" "D \ 0" "\x. x \ X \ dist (f x) (g x) \ D" shows "lambda (C + 2 * D)-quasi_isometry_on X g" proof (rule quasi_isometry_onI) show "lambda \ 1" "C + 2 * D \ 0" using \D \ 0\ quasi_isometry_onD[OF assms(1)] by auto fix x y assume *: "x \ X" "y \ X" have "dist (g x) (g y) \ dist (f x) (f y) + 2 * D" using assms(3)[OF *(1)] assms(3)[OF *(2)] dist_triangle4[of "g x" "g y" "f x" "f y"] by (simp add: dist_commute) then show "dist (g x) (g y) \ lambda * dist x y + (C + 2 * D)" using quasi_isometry_onD(1)[OF assms(1) *] by auto have "dist (g x) (g y) \ dist (f x) (f y) - 2 * D" using assms(3)[OF *(1)] assms(3)[OF *(2)] dist_triangle4[of "f x" "f y" "g x" "g y"] by (simp add: dist_commute) then show "dist (g x) (g y) \ (1/lambda) * dist x y - (C + 2 * D)" using quasi_isometry_onD(2)[OF assms(1) *] by auto qed lemma quasi_isometry_on_compose: assumes "lambda C-quasi_isometry_on X f" "mu D-quasi_isometry_on Y g" "f`X \ Y" shows "(lambda * mu) (C * mu + D)-quasi_isometry_on X (g o f)" proof (rule quasi_isometry_onI) have I: "lambda \ 1" "C \ 0" "mu \ 1" "D \ 0" using quasi_isometry_onD[OF assms(1)] quasi_isometry_onD[OF assms(2)] by auto then show "lambda * mu \ 1" "C * mu + D \ 0" by (auto, metis dual_order.order_iff_strict le_numeral_extra(2) mult_le_cancel_right1 order.strict_trans1) fix x y assume inX: "x \ X" "y \ X" then have inY: "f x \ Y" "f y \ Y" using \f`X \ Y\ by auto have "dist ((g o f) x) ((g o f) y) \ mu * dist (f x) (f y) + D" using quasi_isometry_onD(1)[OF assms(2) inY] by simp also have "... \ mu * (lambda * dist x y + C) + D" using \mu \ 1\ quasi_isometry_onD(1)[OF assms(1) inX] by auto finally show "dist ((g o f) x) ((g o f) y) \ (lambda * mu) * dist x y + (C * mu + D)" by (auto simp add: algebra_simps) have "(1/(lambda * mu)) * dist x y - (C * mu + D) \ (1/(lambda * mu)) * dist x y - (C/mu + D)" using \mu \ 1\ \C \ 0\ apply (auto, auto simp add: divide_simps) by (metis eq_iff less_eq_real_def mult.commute mult_eq_0_iff mult_le_cancel_right1 order.trans) also have "... = (1/mu) * ((1/lambda) * dist x y - C) - D" by (auto simp add: algebra_simps) also have "... \ (1/mu) * dist (f x) (f y) - D" using \mu \ 1\ quasi_isometry_onD(2)[OF assms(1) inX] by (auto simp add: divide_simps) also have "... \ dist ((g o f) x) ((g o f) y)" using quasi_isometry_onD(2)[OF assms(2) inY] by auto finally show "1 / (lambda * mu) * dist x y - (C * mu + D) \ dist ((g \ f) x) ((g \ f) y)" by auto qed lemma quasi_isometry_on_bounded: assumes "lambda C-quasi_isometry_on X f" "bounded X" shows "bounded (f`X)" proof (cases "X = {}") case True then show ?thesis by auto next case False obtain x where "x \ X" using False by auto obtain e where e: "\z. z \ X \ dist x z \ e" using bounded_any_center assms(2) by metis have "dist (f x) y \ C + lambda * e" if "y \ f`X" for y proof - obtain z where *: "z \ X" "y = f z" using \y \ f`X\ by auto have "dist (f x) y \ lambda * dist x z + C" unfolding \y = f z\ using * quasi_isometry_onD(1)[OF assms(1) \x \ X\ \z \ X\] by (auto simp add: add_mono) also have "... \ C + lambda * e" using e[OF \z \ X\] quasi_isometry_onD(3)[OF assms(1)] by auto finally show ?thesis by simp qed then show ?thesis unfolding bounded_def by auto qed lemma quasi_isometry_on_empty: assumes "C \ 0" "lambda \ 1" shows "lambda C-quasi_isometry_on {} f" using assms unfolding quasi_isometry_on_def by auto text \Quasi-isometries change the distance to a set by at most $\lambda \cdot + C$, this follows readily from the fact that this inequality holds pointwise.\ lemma quasi_isometry_on_infdist: assumes "lambda C-quasi_isometry_on X f" "w \ X" "S \ X" shows "infdist (f w) (f`S) \ lambda * infdist w S + C" "infdist (f w) (f`S) \ (1/lambda) * infdist w S - C" proof - have "lambda \ 1" "C \ 0" using quasi_isometry_onD[OF assms(1)] by auto show "infdist (f w) (f`S) \ lambda * infdist w S + C" proof (cases "S = {}") case True then show ?thesis using \C \ 0\ unfolding infdist_def by auto next case False then have "(INF x\S. dist (f w) (f x)) \ (INF x\S. lambda * dist w x + C)" apply (rule cINF_superset_mono) apply (meson bdd_belowI2 zero_le_dist) using assms by (auto intro!: quasi_isometry_onD(1)[OF assms(1)]) also have "... = (INF t\(dist w)`S. lambda * t + C)" by (auto simp add: image_comp) also have "... = lambda * Inf ((dist w)`S) + C" apply (rule continuous_at_Inf_mono[symmetric]) unfolding mono_def using \lambda \ 1\ False by (auto intro!: continuous_intros) finally show ?thesis unfolding infdist_def using False by (auto simp add: image_comp) qed show "1 / lambda * infdist w S - C \ infdist (f w) (f ` S)" proof (cases "S = {}") case True then show ?thesis using \C \ 0\ unfolding infdist_def by auto next case False then have "(1/lambda) * infdist w S - C = (1/lambda) * Inf ((dist w)`S) - C" unfolding infdist_def by auto also have "... = (INF t\(dist w)`S. (1/lambda) * t - C)" apply (rule continuous_at_Inf_mono) unfolding mono_def using \lambda \ 1\ False by (auto simp add: divide_simps intro!: continuous_intros) also have "... = (INF x\S. (1/lambda) * dist w x - C)" by (auto simp add: image_comp) also have "... \ (INF x\S. dist (f w) (f x))" apply (rule cINF_superset_mono[OF False]) apply (rule bdd_belowI2[of _ "-C"]) using assms \lambda \ 1\ apply simp apply simp apply (rule quasi_isometry_onD(2)[OF assms(1)]) using assms by auto finally show ?thesis unfolding infdist_def using False by (auto simp add: image_comp) qed qed subsection \Quasi-isometric isomorphisms\ text \The notion of isomorphism for quasi-isometries is not that it should be a bijection, as it is a coarse notion, but that it is a bijection up to a bounded displacement. For instance, the inclusion of $\mathbb{Z}$ in $\mathbb{R}$ is a quasi-isometric isomorphism between these spaces, whose (quasi)-inverse (which is non-unique) is given by the function integer part. This is formalized in the next definition.\ definition quasi_isometry_between::"real \ real \ ('a::metric_space) set \ ('b::metric_space) set \ ('a \ 'b) \ bool" ("_ _ -quasi'_isometry'_between" [1000, 999]) where "lambda C-quasi_isometry_between X Y f = ((lambda C-quasi_isometry_on X f) \ (f`X \ Y) \ (\y\Y. \x\X. dist (f x) y \ C))" definition quasi_isometric::"('a::metric_space) set \ ('b::metric_space) set \ bool" where "quasi_isometric X Y = (\lambda C f. lambda C-quasi_isometry_between X Y f)" lemma quasi_isometry_betweenD: assumes "lambda C-quasi_isometry_between X Y f" shows "lambda C-quasi_isometry_on X f" "f`X \ Y" "\y. y \ Y \ \x\X. dist (f x) y \ C" "\x y. x \ X \ y \ X \ dist (f x) (f y) \ lambda * dist x y + C" "\x y. x \ X \ y \ X \ dist (f x) (f y) \ (1/lambda) * dist x y - C" "lambda \ 1" "C \ 0" using assms unfolding quasi_isometry_between_def quasi_isometry_on_def by auto lemma quasi_isometry_betweenI: assumes "lambda C-quasi_isometry_on X f" "f`X \ Y" "\y. y \ Y \ \x\X. dist (f x) y \ C" shows "lambda C-quasi_isometry_between X Y f" using assms unfolding quasi_isometry_between_def by auto lemma quasi_isometry_on_between: assumes "lambda C-quasi_isometry_on X f" shows "lambda C-quasi_isometry_between X (f`X) f" using assms unfolding quasi_isometry_between_def quasi_isometry_on_def by force lemma quasi_isometry_between_change_params: assumes "lambda C-quasi_isometry_between X Y f" "mu \ lambda" "D \ C" shows "mu D-quasi_isometry_between X Y f" proof (rule quasi_isometry_betweenI) show "mu D-quasi_isometry_on X f" by (rule quasi_isometry_on_change_params[OF quasi_isometry_betweenD(1)[OF assms(1)] assms(2) assms(3)]) show "f`X \ Y" using quasi_isometry_betweenD[OF assms(1)] by auto fix y assume "y \ Y" show "\x\X. dist (f x) y \ D" using quasi_isometry_betweenD(3)[OF assms(1) \y \ Y\] \D \ C\ by force qed lemma quasi_isometry_subset: assumes "X \ Y" "\y. y \ Y \ \x\X. dist x y \ C" "C \ 0" shows "1 C-quasi_isometry_between X Y (\x. x)" unfolding quasi_isometry_between_def using assms by auto lemma isometry_quasi_isometry_between: assumes "isometry f" shows "1 0-quasi_isometry_between UNIV UNIV f" using assms unfolding quasi_isometry_between_def quasi_isometry_on_def isometry_def isometry_on_def surj_def by (auto) metis proposition quasi_isometry_inverse: assumes "lambda C-quasi_isometry_between X Y f" shows "\g. lambda (3 * C * lambda)-quasi_isometry_between Y X g \ (\x\X. dist x (g (f x)) \ 3 * C * lambda) \ (\y\Y. dist y (f (g y)) \ 3 * C * lambda)" proof - define g where "g = (\y. SOME x. x \ X \ dist (f x) y \ C)" have *: "g y \ X \ dist (f (g y)) y \ C" if "y \ Y" for y unfolding g_def using quasi_isometry_betweenD(3)[OF assms that] by (metis (no_types, lifting) someI_ex) have "lambda \ 1" "C \ 0" using quasi_isometry_betweenD[OF assms] by auto have "C \ 3 * C * lambda" using \lambda \ 1\ \C \ 0\ - by (metis linordered_field_class.sign_simps(24) mult_le_cancel_left1 not_less numeral_le_iff one_eq_numeral_iff order.trans semiring_norm(68)) + by (simp add: algebra_simps mult_ge1_mono) then have A: "dist y (f (g y)) \ 3 * C * lambda" if "y \ Y" for y using *[OF that] by (simp add: dist_commute) have B: "dist x (g (f x)) \ 3 * C * lambda" if "x \ X" for x proof - have "f x \ Y" using that quasi_isometry_betweenD(2)[OF assms] by auto have "(1/lambda) * dist x (g (f x)) - C \ dist (f x) (f (g (f x)))" apply (rule quasi_isometry_betweenD(5)[OF assms]) using that *[OF \f x \ Y\] by auto also have "... \ C" using *[OF \f x \ Y\] by (simp add: dist_commute) finally have "dist x (g (f x)) \ 2 * C * lambda" using \lambda \ 1\ \C \ 0\ by (simp add: divide_simps) also have "... \ 3 * C * lambda" using \lambda \ 1\ \C \ 0\ by (simp add: divide_simps) finally show ?thesis by auto qed have "lambda (3 * C * lambda)-quasi_isometry_on Y g" proof (rule quasi_isometry_onI) show "lambda \ 1" "3 * C * lambda \ 0" using \lambda \ 1\ \C \ 0\ by auto fix y1 y2 assume inY: "y1 \ Y" "y2 \ Y" then have inX: "g y1 \ X" "g y2 \ X" using * by auto have "dist y1 y2 \ dist y1 (f (g y1)) + dist (f (g y1)) (f (g y2)) + dist (f (g y2)) y2" using dist_triangle4 by auto also have "... \ C + dist (f (g y1)) (f (g y2)) + C" using *[OF inY(1)] *[OF inY(2)] by (auto simp add: dist_commute intro: add_mono) also have "... \ C + (lambda * dist (g y1) (g y2) + C) + C" using quasi_isometry_betweenD(4)[OF assms inX] by (auto intro: add_mono) finally have "dist y1 y2 - 3 * C \ lambda * dist (g y1) (g y2)" by auto then have "dist (g y1) (g y2) \ (1/lambda) * dist y1 y2 - 3 * C / lambda" using \lambda \ 1\ by (auto simp add: divide_simps mult.commute) moreover have "3 * C / lambda \ 3 * C * lambda" using \lambda \ 1\ \C \ 0\ apply (auto simp add: divide_simps mult_le_cancel_left1) by (metis dual_order.order_iff_strict less_1_mult mult.left_neutral) ultimately show "dist (g y1) (g y2) \ (1/lambda) * dist y1 y2 - 3 * C * lambda" by auto have "(1/lambda) * dist (g y1) (g y2) - C \ dist (f (g y1)) (f (g y2))" using quasi_isometry_betweenD(5)[OF assms inX] by auto also have "... \ dist (f (g y1)) y1 + dist y1 y2 + dist y2 (f (g y2))" using dist_triangle4 by auto also have "... \ C + dist y1 y2 + C" using *[OF inY(1)] *[OF inY(2)] by (auto simp add: dist_commute intro: add_mono) finally show "dist (g y1) (g y2) \ lambda * dist y1 y2 + 3 * C * lambda" using \lambda \ 1\ by (auto simp add: divide_simps algebra_simps) qed then have "lambda (3 * C * lambda)-quasi_isometry_between Y X g" proof (rule quasi_isometry_betweenI) show "g ` Y \ X" using * by auto fix x assume "x \ X" have "f x \ Y" "dist (g (f x)) x \ 3 * C * lambda" using B[OF \x \ X\] quasi_isometry_betweenD(2)[OF assms] \x \ X\ by (auto simp add: dist_commute) then show "\y\Y. dist (g y) x \ 3 * C * lambda" by blast qed then show ?thesis using A B by blast qed proposition quasi_isometry_compose: assumes "lambda C-quasi_isometry_between X Y f" "mu D-quasi_isometry_between Y Z g" shows "(lambda * mu) (C * mu + 2 * D)-quasi_isometry_between X Z (g o f)" proof (rule quasi_isometry_betweenI) have "(lambda * mu) (C * mu + D)-quasi_isometry_on X (g \ f)" by (rule quasi_isometry_on_compose[OF quasi_isometry_betweenD(1)[OF assms(1)] quasi_isometry_betweenD(1)[OF assms(2)] quasi_isometry_betweenD(2)[OF assms(1)]]) then show "(lambda * mu) (C * mu + 2 * D)-quasi_isometry_on X (g \ f)" apply (rule quasi_isometry_on_change_params) using quasi_isometry_betweenD(7)[OF assms(2)] by auto show "(g \ f) ` X \ Z" using quasi_isometry_betweenD(2)[OF assms(1)] quasi_isometry_betweenD(2)[OF assms(2)] by auto fix z assume "z \ Z" obtain y where y: "y \ Y" "dist (g y) z \ D" using quasi_isometry_betweenD(3)[OF assms(2) \z \ Z\] by auto obtain x where x: "x \ X" "dist (f x) y \ C" using quasi_isometry_betweenD(3)[OF assms(1) \y \ Y\] by auto have "dist ((g o f) x) z \ dist (g (f x)) (g y) + dist (g y) z" using dist_triangle by auto also have "... \ (mu * dist (f x) y + D) + D" apply (rule add_mono, rule quasi_isometry_betweenD(4)[OF assms(2)]) using x y quasi_isometry_betweenD(2)[OF assms(1)] by auto also have "... \ C * mu + 2 * D" using x(2) quasi_isometry_betweenD(6)[OF assms(2)] by auto finally show "\x\X. dist ((g \ f) x) z \ C * mu + 2 * D" using x(1) by auto qed theorem quasi_isometric_equiv_rel: "quasi_isometric X X" "quasi_isometric X Y \ quasi_isometric Y Z \ quasi_isometric X Z" "quasi_isometric X Y \ quasi_isometric Y X" proof - show "quasi_isometric X X" unfolding quasi_isometric_def using quasi_isometry_subset[of X X 0] by auto assume H: "quasi_isometric X Y" then show "quasi_isometric Y X" unfolding quasi_isometric_def using quasi_isometry_inverse by blast assume "quasi_isometric Y Z" then show "quasi_isometric X Z" using H unfolding quasi_isometric_def using quasi_isometry_compose by blast qed text \Many interesting properties in geometric group theory are invariant under quasi-isometry. We prove the most basic ones here.\ lemma quasi_isometric_empty: assumes "X = {}" "quasi_isometric X Y" shows "Y = {}" using assms unfolding quasi_isometric_def quasi_isometry_between_def quasi_isometry_on_def by blast lemma quasi_isometric_bounded: assumes "bounded X" "quasi_isometric X Y" shows "bounded Y" proof (cases "X = {}") case True show ?thesis using quasi_isometric_empty[OF True assms(2)] by auto next case False obtain lambda C f where QI: "lambda C-quasi_isometry_between X Y f" using assms(2) unfolding quasi_isometric_def by auto obtain x where "x \ X" using False by auto obtain e where e: "\z. z \ X \ dist x z \ e" using bounded_any_center assms(1) by metis have "dist (f x) y \ 2 * C + lambda * e" if "y \ Y" for y proof - obtain z where *: "z \ X" "dist (f z) y \ C" using quasi_isometry_betweenD(3)[OF QI \y \ Y\] by auto have "dist (f x) y \ dist (f x) (f z) + dist (f z) y" using dist_triangle by auto also have "... \ (lambda * dist x z + C) + C" using * quasi_isometry_betweenD(4)[OF QI \x \ X\ \z \ X\] by (auto simp add: add_mono) also have "... \ 2 * C + lambda * e" using quasi_isometry_betweenD(6)[OF QI] e[OF \z \ X\] by (auto simp add: algebra_simps) finally show ?thesis by simp qed then show ?thesis unfolding bounded_def by auto qed lemma quasi_isometric_bounded_iff: assumes "bounded X" "X \ {}" "bounded Y" "Y \ {}" shows "quasi_isometric X Y" proof - obtain x y where "x \ X" "y \ Y" using assms by auto obtain C where C: "\z. z \ Y \ dist y z \ C" using \bounded Y\ bounded_any_center by metis have "C \ 0" using C[OF \y \ Y\] by auto obtain D where D: "\z. z \ X \ dist x z \ D" using \bounded X\ bounded_any_center by metis have "D \ 0" using D[OF \x \ X\] by auto define f::"'a \ 'b" where "f = (\_. y)" have "1 (C + 2 * D)-quasi_isometry_between X Y f" proof (rule quasi_isometry_betweenI) show "f`X \ Y" unfolding f_def using \y \ Y\ by auto show "1 (C + 2 * D)-quasi_isometry_on X f" proof (rule quasi_isometry_onI, auto simp add: \C \ 0\ \D \ 0\ f_def) fix a b assume "a \ X" "b \ X" have "dist a b \ dist a x + dist x b" using dist_triangle by auto also have "... \ D + D" using D[OF \a \ X\] D[OF \b \ X\] by (auto simp add: dist_commute) finally show "dist a b \ C + 2 * D" using \C \ 0\ by auto qed show "\a\X. dist (f a) z \ C + 2 * D" if "z \ Y" for z unfolding f_def using \x \ X\ C[OF \z \ Y\] \D \ 0\ by auto qed then show ?thesis unfolding quasi_isometric_def by auto qed subsection \Quasi-isometries of Euclidean spaces.\ text \A less trivial fact is that the dimension of euclidean spaces is invariant under quasi-isometries. It is proved below using growth argument, as quasi-isometries preserve the growth rate. The growth of the space is asymptotic behavior of the number of well-separated points that fit in a ball of radius $R$, when $R$ tends to infinity. Up to a suitable equivalence, it is clearly a quasi-isometry invariance. We show below that, in a Euclidean space of dimension $d$, the growth is like $R^d$: the upper bound is obtained by using the fact that we have disjoint balls inside a big ball, hence volume controls conclude the argument, while the lower bound is obtained by considering integer points.\ text \First, we show that the growth rate of a Euclidean space of dimension $d$ is bounded from above by $R^d$, using the control on measure of disjoint balls and a volume argument.\ proposition growth_rate_euclidean_above: fixes D::real assumes "D > (0::real)" and H: "F \ cball (0::'a::euclidean_space) R" "R \ 0" "\x y. x \ F \ y \ F \ x \ y \ dist x y \ D" shows "finite F \ card F \ 1 + ((6/D)^(DIM('a))) * R^(DIM('a))" proof - define C::real where "C = ((6/D)^(DIM('a)))" have "C \ 0" unfolding C_def using \D > 0\ by auto have "D/3 \ 0" using assms by auto have "finite F \ card F \ 1 + C * R^(DIM('a))" proof (cases "R < D/2") case True have "x = y" if "x \ F" "y \ F" for x y proof (rule ccontr) assume "\(x = y)" then have "D \ dist x y" using H \x \ F\ \y \ F\ by auto also have "... \ dist x 0 + dist 0 y" by (rule dist_triangle) also have "... \ R + R" using H(1) \x \ F\ \y \ F\ by (intro add_mono, auto) also have "... < D" using \R < D/2\ by auto finally show False by simp qed then have "finite F \ card F \ 1" using finite_at_most_singleton by auto moreover have "1 + 0 * R^(DIM('a)) \ 1 + C * R^(DIM('a))" using \C \ 0\ \R \ 0\ by (auto intro: mono_intros) ultimately show ?thesis by auto next case False have "card G \ 1 + C * R^(DIM('a))" if "G \ F" "finite G" for G proof - have "norm y \ 2*R" if "y \ cball x (D/3)" "x \ G" for x y proof - have "norm y = dist 0 y" by auto also have "... \ dist 0 x + dist x y" by (rule dist_triangle) also have "... \ R + D/3" using \x \ G\ \G \ F\ \y \ cball x (D/3)\ \F \ cball 0 R\ by (auto intro: add_mono) finally show ?thesis using False \D > 0\ by auto qed then have I: "(\x\G. cball x (D/3)) \ cball 0 (2*R)" by auto have "disjoint_family_on (\x. cball x (D/3)) G" unfolding disjoint_family_on_def proof (auto) fix a b x assume *: "a \ G" "b \ G" "a \ b" "dist a x * 3 \ D" "dist b x * 3 \ D" then have "D \ dist a b" using H \G \ F\ by auto also have "... \ dist a x + dist x b" by (rule dist_triangle) also have "... \ D/3 + D/3" using * by (auto simp add: dist_commute intro: mono_intros) also have "... < D" using \D > 0\ by auto finally show False by simp qed have "2 * R \ 0" using \R \ 0\ by auto define A where "A = measure lborel (cball (0::'a) 1)" have "A > 0" unfolding A_def using lebesgue_measure_ball_pos by auto have "card G * ((D/3)^(DIM('a)) * A) = (\x\G. ((D/3)^(DIM('a)) * A))" by auto also have "... = (\x\G. measure lborel (cball x (D/3)))" unfolding lebesgue_measure_ball[OF \D/3 \ 0\] A_def by auto also have "... = measure lborel (\x\G. cball x (D/3))" apply (rule measure_finite_Union[symmetric, OF \finite G\ _ \disjoint_family_on (\x. cball x (D/3)) G\]) apply auto using emeasure_bounded_finite less_imp_neq by auto also have "... \ measure lborel (cball (0::'a) (2*R))" apply (rule measure_mono_fmeasurable) using I \finite G\ emeasure_bounded_finite unfolding fmeasurable_def by auto also have "... = (2*R)^(DIM('a)) * A" unfolding A_def using lebesgue_measure_ball[OF \2*R \ 0\] by auto finally have "card G * (D/3)^(DIM('a)) \ (2*R)^(DIM('a))" using \A > 0\ by (auto simp add: divide_simps) then have "card G \ C * R^(DIM('a))" unfolding C_def using \D > 0\ apply (auto simp add: algebra_simps divide_simps) by (metis numeral_times_numeral power_mult_distrib semiring_norm(12) semiring_norm(14)) then show ?thesis by auto qed then show "finite F \ card F \ 1 + C * R^(DIM('a))" by (rule finite_finite_subset_caract') qed then show ?thesis unfolding C_def by blast qed text \Then, we show that the growth rate of a Euclidean space of dimension $d$ is bounded from below by $R^d$, using integer points.\ proposition growth_rate_euclidean_below: fixes D::real assumes "R \ 0" shows "\F. (F \ cball (0::'a::euclidean_space) R \ (\x\F. \y\F. x = y \ dist x y \ D) \ finite F \ card F \ (1/((max D 1) * DIM('a)))^(DIM('a)) * R^(DIM('a)))" proof - define E where "E = max D 1" have "E > 0" unfolding E_def by auto define c where "c = (1/(E * DIM('a)))^(DIM('a))" have "c > 0" unfolding c_def using \E > 0\ by auto define n where "n = nat (floor (R/(E * DIM('a)))) + 1" then have "n > 0" using \R \ 0\ by auto have "R/(E * DIM('a)) \ n" unfolding n_def by linarith then have "c * R^(DIM('a)) \ n^(DIM('a))" unfolding c_def power_mult_distrib[symmetric] by (auto simp add: \0 < E\ \0 \ R\ less_imp_le power_mono) have "n-1 \ R/(E * DIM('a))" unfolding n_def using \R \ 0\ \E > 0\ by auto then have "E * DIM('a) * (n-1) \ R" using \R \ 0\ \E > 0\ by (simp add: mult.commute pos_le_divide_eq) text \We want to consider the set of linear combinations of basis elements with integer coefficients bounded by $n$ (multiplied by $E$ to guarantee the $D$ separation). The formal way to write these elements is to consider all the functions from the basis to $\{0,\dotsc, n-1\}$, and associate to such a function $f$ the point $\sum E f(i) \cdot i$ where the sum is over all basis elements $i$. This is what the next definition does.\ define F::"'a set" where "F = (\f. (\i\Basis. (E * real (f i)) *\<^sub>R i))`((Basis::('a set)) \\<^sub>E {0.. (Basis::('a set)) \\<^sub>E {0.. Basis \\<^sub>E {0..i\Basis. (E * real (f i)) *\<^sub>R i) = (\i\Basis. (E * real (g i)) *\<^sub>R i)" for f g proof (rule ext) fix i show "f i = g i" proof (cases "i \ Basis") case True then have "E * real(f i) = E * real(g i)" using inner_sum_left_Basis[OF True, of "\i. E * real(f i)"] inner_sum_left_Basis[OF True, of "\i. E * real(g i)"] that(3) by auto then show "f i = g i" using \E > 0\ by auto next case False then have "f i = undefined" "g i = undefined" using that by auto then show "f i = g i" by auto qed qed then have "inj_on (\f. (\i\Basis. (E * real (f i)) *\<^sub>R i)) ((Basis::('a set)) \\<^sub>E {0..\<^sub>E {0..n > 0\ using card_infinite by force have "card F \ c * R^(DIM('a))" using \c * R^(DIM('a)) \ n^(DIM('a))\ \card F = n^(DIM('a))\ by auto have separation: "dist x y \ D" if "x \ F" "y \ F" "x \ y" for x y proof - obtain f where x: "f \ (Basis::('a set)) \\<^sub>E {0..i\Basis. (E * real (f i)) *\<^sub>R i)" using \x \ F\ unfolding F_def by auto obtain g where y: "g \ (Basis::('a set)) \\<^sub>E {0..i\Basis. (E * real (g i)) *\<^sub>R i)" using \y \ F\ unfolding F_def by auto obtain i where "f i \ g i" using x y \x \y\ by force moreover have "f j = g j" if "j \ Basis" for j using x(1) y(1) that by fastforce ultimately have "i \ Basis" by auto have "D \ E" unfolding E_def by auto also have "... \ abs(E * (real (f i) - real (g i)))" using \E > 0\ using \f i \ g i\ by (auto simp add: divide_simps abs_mult) also have "... = abs(inner x i - inner y i)" unfolding x(2) y(2) inner_sum_left_Basis[OF \i \ Basis\] by (auto simp add: algebra_simps) also have "... = abs(inner (x-y) i)" by (simp add: inner_diff_left) also have "... \ norm (x-y)" using Basis_le_norm[OF \i \ Basis\] by blast finally show "dist x y \ D" by (simp add: dist_norm) qed have "norm x \ R" if "x \ F" for x proof - obtain f where x: "f \ (Basis::('a set)) \\<^sub>E {0..i\Basis. (E * real (f i)) *\<^sub>R i)" using \x \ F\ unfolding F_def by auto then have "norm x = norm (\i\Basis. (E * real (f i)) *\<^sub>R i)" by simp also have "... \ (\i\Basis. norm((E * real (f i)) *\<^sub>R i))" by (rule norm_sum) also have "... = (\i\Basis. abs(E * real (f i)))" by auto also have "... = (\i\Basis. E * real (f i))" using \E > 0\ by auto also have "... \ (\i\(Basis::'a set). E * (n-1))" apply (rule sum_mono) using PiE_mem[OF x(1)] \E > 0\ apply (auto simp add: divide_simps) using \n > 0\ by fastforce also have "... = DIM('a) * E * (n-1)" by auto finally show "norm x \ R" using \E * DIM('a) * (n-1) \ R\ by (auto simp add: algebra_simps) qed then have "F \ cball 0 R" by auto then show ?thesis using \card F \ c * R^(DIM('a))\ \finite F\ separation c_def E_def by blast qed text \As the growth is invariant under quasi-isometries, we deduce that it is impossible to map quasi-isometrically a Euclidean space in a space of strictly smaller dimension.\ proposition quasi_isometry_on_euclidean: fixes f::"'a::euclidean_space\'b::euclidean_space" assumes "lambda C-quasi_isometry_on UNIV f" shows "DIM('a) \ DIM('b)" proof - have C: "lambda \ 1" "C \ 0" using quasi_isometry_onD[OF assms] by auto define D where "D = lambda * (C+1)" define Ca where "Ca = (1/((max D 1) * DIM('a)))^(DIM('a))" have "Ca > 0" unfolding Ca_def by auto have A: "\R::real. R \ 0 \ (\F. (F \ cball (0::'a::euclidean_space) R \ (\x\F. \y\F. x = y \ dist x y \ D) \ finite F \ card F \ Ca * R^(DIM('a))))" using growth_rate_euclidean_below[of _ D] unfolding Ca_def by blast define Cb::real where "Cb = ((6/1)^(DIM('b)))" have B: "\F (R::real). (F \ cball (0::'b::euclidean_space) R \ R \ 0 \ (\x\F. \y\F. x = y \ dist x y \ 1) \ (finite F \ card F \ 1 + Cb * R^(DIM('b))))" using growth_rate_euclidean_above[of 1] unfolding Cb_def by fastforce have M: "Ca * R^(DIM('a)) \ 1 + Cb * (lambda * R + C + norm(f 0))^(DIM('b))" if "R \ 0" for R::real proof - obtain F::"'a set" where F: "F \ cball 0 R" "\x\F. \y\F. x = y \ dist x y \ D" "finite F" "card F \ Ca * R^(DIM('a))" using A[OF \R \ 0\] by auto define G where "G = f`F" have *: "dist (f x) (f y) \ 1" if "x \ y" "x \ F" "y \ F" for x y proof - have "dist x y \ D" using that F(2) by auto have "1 = (1/lambda) * D - C" using \lambda \ 1\ unfolding D_def by auto also have "... \ (1/lambda) * dist x y - C" using \dist x y \ D\ \lambda \ 1\ by (auto simp add: divide_simps) also have "... \ dist (f x) (f y)" using quasi_isometry_onD[OF assms] by auto finally show ?thesis by simp qed then have "inj_on f F" unfolding inj_on_def by force then have "card G = card F" unfolding G_def by (simp add: card_image) then have "card G \ Ca * R^(DIM('a))" using F by auto moreover have "finite G \ card G \ 1 + Cb * (lambda * R + C + norm(f 0))^(DIM('b))" proof (rule B) show "0 \ lambda * R + C + norm (f 0)" using \R \ 0\ \C \ 0\ \lambda \ 1\ by auto show "\x\G. \y\G. x = y \ 1 \ dist x y" using * unfolding G_def by (auto, metis) show "G \ cball 0 (lambda * R + C + norm (f 0))" unfolding G_def proof (auto) fix x assume "x \ F" have "norm (f x) \ norm (f 0) + dist (f x) (f 0)" by (metis dist_0_norm dist_triangle2) also have "... \ norm (f 0) + (lambda * dist x 0 + C)" by (intro mono_intros quasi_isometry_onD(1)[OF assms]) auto also have "... \ norm (f 0) + lambda * R + C" using \x \ F\ \F \ cball 0 R\ \lambda \ 1\ by auto finally show "norm (f x) \ lambda * R + C + norm (f 0)" by auto qed qed ultimately show "Ca * R^(DIM('a)) \ 1 + Cb * (lambda * R + C + norm(f 0))^(DIM('b))" by auto qed define CB where "CB = max Cb 0" have "CB \ 0" "CB \ Cb" unfolding CB_def by auto define D::real where "D = (1 + CB * (lambda + C + norm(f 0))^(DIM('b)))/Ca" have Rineq: "R^(DIM('a)) \ D * R^(DIM('b))" if "R \ 1" for R::real proof - have "Ca * R^(DIM('a)) \ 1 + Cb * (lambda * R + C + norm(f 0))^(DIM('b))" using M \R \ 1\ by auto also have "... \ 1 + CB * (lambda * R + C + norm(f 0))^(DIM('b))" using \CB \ Cb\ \lambda \ 1\ \R \ 1\ \C \ 0\ by (auto intro!: mult_right_mono) also have "... \ R^(DIM('b)) + CB * (lambda * R + C * R + norm(f 0) * R)^(DIM('b))" using \lambda \ 1\ \R \ 1\ \C \ 0\ \CB \ 0\ by (auto intro!: mono_intros) also have "... = (1 + CB * (lambda + C + norm(f 0))^(DIM('b))) * R^(DIM('b))" by (auto simp add: algebra_simps power_mult_distrib[symmetric]) finally show ?thesis using \Ca > 0\ unfolding D_def by (auto simp add: divide_simps algebra_simps) qed show "DIM('a) \ DIM('b)" proof (rule ccontr) assume "\(DIM('a) \ DIM('b))" then obtain n where "DIM('a) = DIM('b) + n" "n > 0" by (metis less_imp_add_positive not_le) have "D \ 1" using Rineq[of 1] by auto define R where "R = 2 * D" then have "R \ 1" using \D \ 1\ by auto have "R^n * R^(DIM('b)) = R^(DIM('a))" unfolding \DIM('a) = DIM('b) + n\ by (auto simp add: power_add) also have "... \ D * R^(DIM('b))" using Rineq[OF \R \ 1\] by auto finally have "R^n \ D" using \R \ 1\ by auto moreover have "2 * D \ R^n" unfolding R_def using \D \ 1\ \n > 0\ by (metis One_nat_def Suc_leI \1 \ R\ \R \ 2 * D\ less_eq_real_def power_increasing_iff power_one power_one_right) ultimately show False using \D \ 1\ by auto qed qed text \As a particular case, we deduce that two quasi-isometric Euclidean spaces have the same dimension.\ theorem quasi_isometric_euclidean: assumes "quasi_isometric (UNIV::'a::euclidean_space set) (UNIV::'b::euclidean_space set)" shows "DIM('a) = DIM('b)" proof - obtain lambda C and f::"'a \'b" where "lambda C-quasi_isometry_on UNIV f" using assms unfolding quasi_isometric_def quasi_isometry_between_def by auto then have *: "DIM('a) \ DIM('b)" using quasi_isometry_on_euclidean by auto have "quasi_isometric (UNIV::'b::euclidean_space set) (UNIV::'a::euclidean_space set)" using quasi_isometric_equiv_rel(3)[OF assms] by auto then obtain lambda C and f::"'b \'a" where "lambda C-quasi_isometry_on UNIV f" unfolding quasi_isometric_def quasi_isometry_between_def by auto then have "DIM('b) \ DIM('a)" using quasi_isometry_on_euclidean by auto then show ?thesis using * by auto qed text \A different (and important) way to prove the above statement would be to use asymptotic cones. Here, it can be done in an elementary way: start with a quasi-isometric map $f$, and consider a limit (defined with a ultrafilter) of $x\mapsto f(n x)/n$. This is a map which contracts and expands the distances by at most $\lambda$. In particular, it is a homeomorphism on its image. No such map exists if the dimension of the target is smaller than the dimension of the source (invariance of domain theorem, already available in the library). The above argument using growth is more elementary to write, though.\ subsection \Quasi-geodesics\ text \A quasi-geodesic is a quasi-isometric embedding of a real segment into a metric space. As the embedding need not be continuous, a quasi-geodesic does not have to be compact, nor connected, which can be a problem. However, in a geodesic space, it is always possible to deform a quasi-geodesic into a continuous one (at the price of worsening the quasi-isometry constants). This is the content of the proposition \verb+quasi_geodesic_made_lipschitz+ below, which is a variation around Lemma III.H.1.11 in~\cite{bridson_haefliger}. The strategy of the proof is simple: assume that the quasi-geodesic $c$ is defined on $[a,b]$. Then, on the points $a$, $a+C/\lambda$, $\cdots$, $a+ N \cdot C/\lambda$, $b$, take $d$ equal to $c$, where $N$ is chosen so that the distance between the last point and $b$ is in $[C/\lambda, 2C/\lambda)$. In the intervals, take $d$ to be geodesic.\ proposition (in geodesic_space) quasi_geodesic_made_lipschitz: fixes c::"real \ 'a" assumes "lambda C-quasi_isometry_on {a..b} c" "dist (c a) (c b) \ 2 * C" shows "\d. continuous_on {a..b} d \ d a = c a \ d b = c b \ (\x\{a..b}. dist (c x) (d x) \ 4 * C) \ lambda (4 * C)-quasi_isometry_on {a..b} d \ (2 * lambda)-lipschitz_on {a..b} d \ hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C" proof - consider "C = 0" | "C > 0 \ b \ a" | "C > 0 \ a < b \ b \ a + 2 * C/lambda" | "C > 0 \ a +2 * C/lambda < b" using quasi_isometry_onD(4)[OF assms(1)] by fastforce then show ?thesis proof (cases) text \If the original function is Lipschitz, we can use it directly.\ case 1 have "lambda-lipschitz_on {a..b} c" apply (rule lipschitz_onI) using 1 quasi_isometry_onD[OF assms(1)] by auto then have a: "(2 * lambda)-lipschitz_on {a..b} c" apply (rule lipschitz_on_mono) using quasi_isometry_onD[OF assms(1)] assms by (auto simp add: divide_simps) then have b: "continuous_on {a..b} c" using lipschitz_on_continuous_on by blast have "continuous_on {a..b} c \ c a = c a \ c b = c b \ (\x\{a..b}. dist (c x) (c x) \ 4 * C) \ lambda (4 * C)-quasi_isometry_on {a..b} c \ (2 * lambda)-lipschitz_on {a..b} c \ hausdorff_distance (c`{a..b}) (c`{a..b}) \ 2 * C" using 1 a b assms(1) by auto then show ?thesis by blast next text \If the original interval is empty, anything will do.\ case 2 then have "b < a" using assms(2) less_eq_real_def by auto then have *: "{a..b} = {}" by auto have a: "(2 * lambda)-lipschitz_on {a..b} c" unfolding * apply (rule lipschitz_intros) using quasi_isometry_onD[OF assms(1)] assms by (auto simp add: divide_simps) then have b: "continuous_on {a..b} c" using lipschitz_on_continuous_on by blast have "continuous_on {a..b} c \ c a = c a \ c b = c b \ (\x\{a..b}. dist (c x) (c x) \ 4 * C) \ lambda (4 * C)-quasi_isometry_on {a..b} c \ (2 * lambda)-lipschitz_on {a..b} c \ hausdorff_distance (c`{a..b}) (c`{a..b}) \ 2 * C" using a b quasi_isometry_on_empty assms(1) quasi_isometry_onD[OF assms(1)] * assms by auto then show ?thesis by blast next text \If the original interval is short, we can use a direct geodesic interpolation between its endpoints\ case 3 then have C: "C > 0" "lambda \ 1" using quasi_isometry_onD[OF assms(1)] by auto have [mono_intros]: "1/lambda \ lambda" using C by (simp add: divide_simps mult_ge1_powers(1)) have "a < b" using 3 by simp have "2 * C \ dist (c a) (c b)" using assms by auto also have "... \ lambda * dist a b + C" using quasi_isometry_onD[OF assms(1)] \a < b\ by auto also have "... = lambda * (b-a) + C" using \a < b\ dist_real_def by auto finally have *: "C \ (b-a) * lambda" by (auto simp add: algebra_simps) define d where "d = (\x. geodesic_segment_param {(c a)--(c b)} (c a) ((dist (c a) (c b) /(b-a)) * (x-a)))" have dend: "d a = c a" "d b = c b" unfolding d_def using \a < b\ by auto have Lip: "(2 * lambda)-lipschitz_on {a..b} d" proof - have "(1 * (((2 * lambda)) * (1+0)))-lipschitz_on {a..b} (\x. geodesic_segment_param {(c a)--(c b)} (c a) ((dist (c a) (c b) /(b-a)) * (x-a)))" proof (rule lipschitz_on_compose2[of _ _ "\x. ((dist (c a) (c b) /(b-a)) * (x-a))"], intro lipschitz_intros) have "(\x. dist (c a) (c b) / (b-a) * (x - a)) ` {a..b} \ {0..dist (c a) (c b)}" apply auto using \a < b\ by (auto simp add: algebra_simps divide_simps intro: mult_right_mono) moreover have "1-lipschitz_on {0..dist (c a) (c b)} (geodesic_segment_param {c a--c b} (c a))" by (rule isometry_on_lipschitz, simp) ultimately show "1-lipschitz_on ((\x. dist (c a) (c b) / (b-a) * (x - a)) ` {a..b}) (geodesic_segment_param {c a--c b} (c a))" using lipschitz_on_subset by auto have "dist (c a) (c b) \ lambda * dist a b + C" apply (rule quasi_isometry_onD(1)[OF assms(1)]) using \a < b\ by auto also have "... = lambda * (b - a) + C" unfolding dist_real_def using \a < b\ by auto also have "... \ 2 * lambda * (b-a)" using * by (auto simp add: algebra_simps) finally show "\dist (c a) (c b) / (b - a)\ \ 2 * lambda" using \a < b\ by (auto simp add: divide_simps) qed then show ?thesis unfolding d_def by auto qed have dist_c_d: "dist (c x) (d x) \ 4 * C" if H: "x \ {a..b}" for x proof - have "(x-a) + (b - x) \ 2 * C/lambda" using that 3 by auto then consider "x-a \ C/lambda" | "b - x \ C/lambda" by linarith then have "\v\{a,b}. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ a]) using 1 H by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ b]) using 2 H by (auto simp add: dist_real_def) qed then obtain v where v: "v \ {a,b}" "dist x v \ C/lambda" by auto have "dist (c x) (d x) \ dist (c x) (c v) + dist (c v) (d v) + dist (d v) (d x)" by (intro mono_intros) also have "... \ (lambda * dist x v + C) + 0 + ((2 * lambda) * dist v x)" apply (intro mono_intros quasi_isometry_onD(1)[OF assms(1)] that lipschitz_onD[OF Lip]) using v \a < b\ dend by auto also have "... \ (lambda * (C/lambda) + C) + 0 + ((2 * lambda) * (C/lambda))" apply (intro mono_intros) using C v by (auto simp add: metric_space_class.dist_commute) finally show ?thesis using C by (auto simp add: algebra_simps divide_simps) qed text \A similar argument shows that the Hausdorff distance between the images is bounded by $2C$.\ have "hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C" proof (rule hausdorff_distanceI2) show "0 \ 2 * C" using C by auto fix z assume "z \ c`{a..b}" then obtain x where x: "x \ {a..b}" "z = c x" by auto have "(x-a) + (b - x) \ 2 * C/lambda" using x 3 by auto then consider "x-a \ C/lambda" | "b - x \ C/lambda" by linarith then have "\v\{a,b}. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ a]) using 1 x by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ b]) using 2 x by (auto simp add: dist_real_def) qed then obtain v where v: "v \ {a,b}" "dist x v \ C/lambda" by auto have "dist z (d v) = dist (c x) (c v)" unfolding x(2) using v dend by auto also have "... \ lambda * dist x v + C" apply (rule quasi_isometry_onD(1)[OF assms(1)]) using v(1) x(1) by auto also have "... \ lambda * (C/lambda) + C" apply (intro mono_intros) using C v(2) by auto also have "... = 2 * C" using C by (simp add: divide_simps) finally have *: "dist z (d v) \ 2 * C" by simp show "\y\d ` {a..b}. dist z y \ 2 * C" apply (rule bexI[of _ "d v"]) using * v(1) \a < b\ by auto next fix z assume "z \ d`{a..b}" then obtain x where x: "x \ {a..b}" "z = d x" by auto have "(x-a) + (b - x) \ 2 * C/lambda" using x 3 by auto then consider "x-a \ C/lambda" | "b - x \ C/lambda" by linarith then have "\v\{a,b}. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ a]) using 1 x by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ b]) using 2 x by (auto simp add: dist_real_def) qed then obtain v where v: "v \ {a,b}" "dist x v \ C/lambda" by auto have "dist z (c v) = dist (d x) (d v)" unfolding x(2) using v dend by auto also have "... \ 2 * lambda * dist x v" apply (rule lipschitz_onD(1)[OF Lip]) using v(1) x(1) by auto also have "... \ 2 * lambda * (C/lambda)" apply (intro mono_intros) using C v(2) by auto also have "... = 2 * C" using C by (simp add: divide_simps) finally have *: "dist z (c v) \ 2 * C" by simp show "\y\c`{a..b}. dist z y \ 2 * C" apply (rule bexI[of _ "c v"]) using * v(1) \a < b\ by auto qed have "lambda (4 * C)-quasi_isometry_on {a..b} d" proof show "1 \ lambda" using C by auto show "0 \ 4 * C" using C by auto show "dist (d x) (d y) \ lambda * dist x y + 4 * C" if "x \ {a..b}" "y \ {a..b}" for x y proof - have "dist (d x) (d y) \ 2 * lambda * dist x y" apply (rule lipschitz_onD[OF Lip]) using that by auto also have "... = lambda * dist x y + lambda * dist x y" by auto also have "... \ lambda * dist x y + lambda * (2 * C/lambda)" apply (intro mono_intros) using 3 that C unfolding dist_real_def by auto also have "... = lambda * dist x y + 2 * C" using C by (simp add: algebra_simps divide_simps) finally show ?thesis using C by auto qed show "1 / lambda * dist x y - 4 * C \ dist (d x) (d y)" if "x \ {a..b}" "y \ {a..b}" for x y proof - have "1/lambda * dist x y - 4 * C \ lambda * dist x y - 2 * C" apply (intro mono_intros) using C by auto also have "... \ lambda * (2 * C/lambda) - 2 * C" apply (intro mono_intros) using that 3 C unfolding dist_real_def by auto also have "... = 0" using C by (auto simp add: algebra_simps divide_simps) also have "... \ dist (d x) (d y)" by auto finally show ?thesis by simp qed qed then have "continuous_on {a..b} d \ d a = c a \ d b = c b \ lambda (4 * C)-quasi_isometry_on {a..b} d \ (\x\{a..b}. dist (c x) (d x) \ 4 *C) \ (2*lambda)-lipschitz_on {a..b} d \ hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C" using dist_c_d \d a = c a\ \d b = c b\ \(2*lambda)-lipschitz_on {a..b} d\ \hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C\ lipschitz_on_continuous_on by auto then show ?thesis by auto next text \Now, for the only nontrivial case, we use geodesic interpolation between the points $a$, $a + C/\lambda$, $\cdots$, $a+N\cdot C/\lambda$, $b'$, $b$ where $N$ is chosen so that the distance between $a+N C/\lambda$ and $b$ belongs to $[2C/\lambda, 3C/\lambda)$, and $b'$ is the middle of this interval. This gives a decomposition into intervals of length at most $3/2\cdot C/\lambda$.\ case 4 then have C: "C > 0" "lambda \ 1" using quasi_isometry_onD[OF assms(1)] by auto have "a < b" using 4 C by (smt divide_pos_pos) have [mono_intros]: "1/lambda \ lambda" using C by (simp add: divide_simps mult_ge1_powers(1)) define N where "N = floor((b-a)/(C/lambda)) - 2" have N: "N \ (b-a)/(C/lambda)-2" "(b-a)/(C/lambda) \ N + (3::real)" unfolding N_def by linarith+ have "2 < (b-a)/(C/lambda)" using C 4 by (auto simp add: divide_simps algebra_simps) then have N0 : "0 \ N" unfolding N_def by auto define p where "p = (\t::int. a + (C/lambda) * t)" have pmono: "p i \ p j" if "i \ j" for i j unfolding p_def using that C by (auto simp add: algebra_simps divide_simps) have pmono': "p i < p j" if "i < j" for i j unfolding p_def using that C by (auto simp add: algebra_simps divide_simps) have "p (N+1) \ b" unfolding p_def using C N by (auto simp add: algebra_simps divide_simps) then have pb: "p i \ b" if "i \ {0..N}" for i using that pmono by (meson atLeastAtMost_iff linear not_le order_trans zle_add1_eq_le) have bpN: "b - p N \ {2 * C/lambda .. 3 * C/lambda}" unfolding p_def using C N apply (auto simp add: divide_simps) by (auto simp add: algebra_simps) have "p N < b" using pmono'[of N "N+1"] \p (N+1) \ b\ by auto define b' where "b' = (b + p N)/2" have b': "p N < b'" "b' < b" using \p N < b\ unfolding b'_def by auto have pb': "p i \ b'" if "i \ {0..N}" for i using pmono[of i N] b' that by auto text \Introduce the set $A$ along which one will discretize.\ define A where "A = p`{0..N} \ {b', b}" have "finite A" unfolding A_def by auto have "b \ A" unfolding A_def by auto have "p 0 \ A" unfolding A_def using \0 \ N\ by auto moreover have pa: "p 0 = a" unfolding p_def by auto ultimately have "a \ A" by auto have "A \ {a..b}" unfolding A_def using \a < b\ b' pa pb pmono N0 by fastforce then have "b' \ {a..b' < b\ by auto have A : "finite A" "A \ {a..b}" "a \ A" "b \ A" "a < b" by fact+ have nx: "next_in A x = x + C/lambda" if "x \ A" "x \ b" "x \ b'" "x \ p N" for x proof (rule next_inI[OF A]) show "x \ {a..x \ A\ \A \ {a..b}\ \x \ b\ by auto obtain i where i: "x = p i" "i \ {0..N}" using \x \ A\ \x \ b\ \x \ b'\ unfolding A_def by auto have *: "p (i+1) = x + C/lambda" unfolding i(1) p_def by (auto simp add: algebra_simps) have "i \ N" using that i by auto then have "i + 1 \ {0..N}" using \i \ {0..N}\ by auto then have "p (i+1) \ A" unfolding A_def by fastforce then show "x + C/lambda \ A" unfolding * by auto show "x < x + C / lambda" using C by auto show "{x<.. A = {}" proof (auto) fix y assume y: "y \ A" "x < y" "y < x + C/lambda" consider "y = b" | "y = b'" | "\j\i. y = p j" | "\j>i. y = p j" using \y \ A\ not_less unfolding A_def by auto then show False proof (cases) case 1 have "x + C/lambda \ b" unfolding *[symmetric] using \i + 1 \ {0..N}\ pb by auto then show False using y(3) unfolding 1 i(1) by auto next case 2 have "x + C/lambda \ b'" unfolding *[symmetric] using \i + 1 \ {0..N}\ pb' by auto then show False using y(3) unfolding 2 i(1) by auto next case 3 then obtain j where j: "j \ i" "y = p j" by auto have "y \ x" unfolding j(2) i(1) using pmono[OF \j \ i\] by simp then show False using \x < y\ by auto next case 4 then obtain j where j: "j > i" "y = p j" by auto then have "i+1 \ j" by auto have "x + C/lambda \ y" unfolding j(2) *[symmetric] using pmono[OF \i+1 \ j\] by auto then show False using \y < x + C/lambda\ by auto qed qed qed have npN: "next_in A (p N) = b'" proof (rule next_inI[OF A]) show "p N \ {a..0 \ N\ \p N < b\ by auto show "p N < b'" by fact show "b' \ A" unfolding A_def by auto show "{p N<.. A = {}" unfolding A_def using pmono b' by force qed have nb': "next_in A (b') = b" proof (rule next_inI[OF A]) show "b' \ {a..b' < b\ by auto show "b' < b" by fact show "b \ A" by fact show "{b'<.. A = {}" unfolding A_def using pmono b' by force qed have gap: "next_in A x - x \ {C/lambda.. 3/2 * C/lambda}" if "x \ A - {b}" for x proof (cases "x = p N \ x = b'") case True then show ?thesis using npN nb' bpN b'_def by force next case False have *: "next_in A x = x + C/lambda" apply (rule nx) using that False by auto show ?thesis unfolding * using C by (auto simp add: algebra_simps divide_simps) qed text \We can now define the function $d$, by geodesic interpolation between points in $A$.\ define d where "d x = (if x \ A then c x else geodesic_segment_param {c (prev_in A x) -- c (next_in A x)} (c (prev_in A x)) ((x - prev_in A x)/(next_in A x - prev_in A x) * dist (c(prev_in A x)) (c(next_in A x))))" for x have "d a = c a" "d b = c b" unfolding d_def using \a \ A\ \b \ A\ by auto text \To prove the Lipschitz continuity, we argue that $d$ is Lipschitz on finitely many intervals, that cover the interval $[a,b]$, the intervals between points in $A$. There is a formula for $d$ on them (the nontrivial point is that the above formulas for $d$ match at the boundaries).\ have *: "d x = geodesic_segment_param {(c u)--(c v)} (c u) ((dist (c u) (c v) /(v-u)) * (x-u))" if "u \ A - {b}" "v = next_in A u" "x \ {u..v}" for x u v proof - have "u \ {a..A \ {a..b}\ by fastforce have H: "u \ A" "v \ A" "u < v" "A \ {u<..u \ {a..] by auto consider "x = u" | "x = v" | "x \ {u<..x \ {u..v}\ by fastforce then show ?thesis proof (cases) case 1 then have "d x = c u" unfolding d_def using \u \ A- {b}\ \A \ {a..b}\ by auto then show ?thesis unfolding 1 by auto next case 2 then have "d x = c v" unfolding d_def using \v \ A\ \A \ {a..b}\ by auto then show ?thesis unfolding 2 using \u < v\ by auto next case 3 have *: "prev_in A x = u" apply (rule prev_inI[OF A]) using 3 H \A \ {a..b}\ by auto have **: "next_in A x = v" apply (rule next_inI[OF A]) using 3 H \A \ {a..b}\ by auto show ?thesis unfolding d_def * ** using 3 H \A \ {u<.. \A \ {a..b}\ by (auto simp add: algebra_simps) qed qed text \From the above formula, we deduce that $d$ is Lipschitz on those intervals.\ have lip0: "(lambda + C / (next_in A u - u))-lipschitz_on {u..next_in A u} d" if "u \ A - {b}" for u proof - define v where "v = next_in A u" have "u \ {a..A \ {a..b}\ by fastforce have "u \ A" "v \ A" "u < v" "A \ {u<..u \ {a..] by auto have "(1 * (((lambda + C / (next_in A u - u))) * (1+0)))-lipschitz_on {u..v} (\x. geodesic_segment_param {(c u)--(c v)} (c u) ((dist (c u) (c v) /(v-u)) * (x-u)))" proof (rule lipschitz_on_compose2[of _ _ "\x. ((dist (c u) (c v) /(v-u)) * (x-u))"], intro lipschitz_intros) have "(\x. dist (c u) (c v) / (v - u) * (x - u)) ` {u..v} \ {0..dist (c u) (c v)}" apply auto using \u < v\ by (auto simp add: algebra_simps divide_simps intro: mult_right_mono) moreover have "1-lipschitz_on {0..dist (c u) (c v)} (geodesic_segment_param {c u--c v} (c u))" by (rule isometry_on_lipschitz, simp) ultimately show "1-lipschitz_on ((\x. dist (c u) (c v) / (v - u) * (x - u)) ` {u..v}) (geodesic_segment_param {c u--c v} (c u))" using lipschitz_on_subset by auto have "dist (c u) (c v) \ lambda * dist u v + C" apply (rule quasi_isometry_onD(1)[OF assms(1)]) using \u \ A\ \v \ A\ \A \ {a..b}\ by auto also have "... = lambda * (v - u) + C" unfolding dist_real_def using \u < v\ by auto finally show "\dist (c u) (c v) / (v - u)\ \ lambda + C / (next_in A u - u)" using \u < v\ unfolding v_def by (auto simp add: divide_simps) qed then show ?thesis using *[OF \u \ A -{b}\ \v = next_in A u\] unfolding v_def by (auto intro: lipschitz_on_transform) qed have lip: "(2 * lambda)-lipschitz_on {u..next_in A u} d" if "u \ A - {b}" for u proof (rule lipschitz_on_mono[OF lip0[OF that]], auto) define v where "v = next_in A u" have "u \ {a..A \ {a..b}\ by fastforce have "u \ A" "v \ A" "u < v" "A \ {u<..u \ {a..] by auto have Duv: "v - u \ {C/lambda .. 2 * C/lambda}" unfolding v_def using gap[OF \u \ A - {b}\] by simp then show " C / (next_in A u - u) \ lambda" using \u < v\ C unfolding v_def by (auto simp add: algebra_simps divide_simps) qed text \The Lipschitz continuity of $d$ now follows from its Lipschitz continuity on each subinterval in $I$.\ have Lip: "(2 * lambda)-lipschitz_on {a..b} d" apply (rule lipschitz_on_closed_Union[of "{{u..next_in A u} |u. u \ A - {b}}" _ "\x. x"]) using lip \finite A\ C intervals_decomposition[OF A] using assms by auto then have "continuous_on {a..b} d" using lipschitz_on_continuous_on by auto text \$d$ has good upper controls on each basic interval.\ have QI0: "dist (d x) (d y) \ lambda * dist x y + C" if H: "u \ A - {b}" "x \ {u..next_in A u}" "y \ {u..next_in A u}" for u x y proof - have "u < next_in A u" using H(1) A next_in_basics(2)[OF A] by auto moreover have "dist x y \ next_in A u - u" unfolding dist_real_def using H by auto ultimately have *: "dist x y / (next_in A u - u) \ 1" by (simp add: divide_simps) have "dist (d x) (d y) \ (lambda + C / (next_in A u - u)) * dist x y" by (rule lipschitz_onD[OF lip0[OF H(1)] H(2) H(3)]) also have "... = lambda * dist x y + C * (dist x y / (next_in A u - u))" by (simp add: algebra_simps) also have "... \ lambda * dist x y + C * 1" apply (intro mono_intros) using C * by auto finally show ?thesis by simp qed text \We can now show that $c$ and $d$ are pointwise close. This follows from the fact that they coincide on $A$ and are well controlled in between (for $c$, this is a consequence of the choice of $A$. For $d$, it follows from the fact that it is geodesic in the intervals).\ have dist_c_d: "dist (c x) (d x) \ 4 * C" if "x \ {a..b}" for x proof - obtain u where u: "u \ A - {b}" "x \ {u..next_in A u}" using \x \ {a..b}\ intervals_decomposition[OF A] by blast have "(x-u) + (next_in A u - x) \ 2 * C/lambda" using gap[OF u(1)] by auto then consider "x-u \ C/lambda" | "next_in A u - x \ C/lambda" by linarith then have "\v\A. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ u]) using 1 u by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ "next_in A u"]) using 2 u A(2) by (auto simp add: dist_real_def intro!:next_in_basics[OF A]) qed then obtain v where v: "v \ A" "dist x v \ C/lambda" by auto have "dist (c x) (d x) \ dist (c x) (c v) + dist (c v) (d v) + dist (d v) (d x)" by (intro mono_intros) also have "... \ (lambda * dist x v + C) + 0 + ((2 * lambda) * dist v x)" apply (intro mono_intros quasi_isometry_onD(1)[OF assms(1)] that lipschitz_onD[OF Lip]) using A(2) \v \ A\ apply blast using \v \ A\ d_def apply auto[1] using A(2) \v \ A\ by blast also have "... \ (lambda * (C/lambda) + C) + 0 + ((2 * lambda) * (C/lambda))" apply (intro mono_intros) using v(2) C by (auto simp add: metric_space_class.dist_commute) finally show ?thesis using C by (auto simp add: algebra_simps divide_simps) qed text \A similar argument shows that the Hausdorff distance between the images is bounded by $2C$.\ have "hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C" proof (rule hausdorff_distanceI2) show "0 \ 2 * C" using C by auto fix z assume "z \ c`{a..b}" then obtain x where x: "x \ {a..b}" "z = c x" by auto then obtain u where u: "u \ A - {b}" "x \ {u..next_in A u}" using intervals_decomposition[OF A] by blast have "(x-u) + (next_in A u - x) \ 2 * C/lambda" using gap[OF u(1)] by auto then consider "x-u \ C/lambda" | "next_in A u - x \ C/lambda" by linarith then have "\v\A. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ u]) using 1 u by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ "next_in A u"]) using 2 u A(2) by (auto simp add: dist_real_def intro!:next_in_basics[OF A]) qed then obtain v where v: "v \ A" "dist x v \ C/lambda" by auto have "dist z (d v) = dist (c x) (c v)" unfolding x(2) d_def using \v \ A\ by auto also have "... \ lambda * dist x v + C" apply (rule quasi_isometry_onD(1)[OF assms(1)]) using v(1) A(2) x(1) by auto also have "... \ lambda * (C/lambda) + C" apply (intro mono_intros) using C v(2) by auto also have "... = 2 * C" using C by (simp add: divide_simps) finally have *: "dist z (d v) \ 2 * C" by simp show "\y\d ` {a..b}. dist z y \ 2 * C" apply (rule bexI[of _ "d v"]) using * v(1) A(2) by auto next fix z assume "z \ d`{a..b}" then obtain x where x: "x \ {a..b}" "z = d x" by auto then obtain u where u: "u \ A - {b}" "x \ {u..next_in A u}" using intervals_decomposition[OF A] by blast have "(x-u) + (next_in A u - x) \ 2 * C/lambda" using gap[OF u(1)] by auto then consider "x-u \ C/lambda" | "next_in A u - x \ C/lambda" by linarith then have "\v\A. dist x v \ C/lambda" proof (cases) case 1 show ?thesis apply (rule bexI[of _ u]) using 1 u by (auto simp add: dist_real_def) next case 2 show ?thesis apply (rule bexI[of _ "next_in A u"]) using 2 u A(2) by (auto simp add: dist_real_def intro!:next_in_basics[OF A]) qed then obtain v where v: "v \ A" "dist x v \ C/lambda" by auto have "dist z (c v) = dist (d x) (d v)" unfolding x(2) d_def using \v \ A\ by auto also have "... \ 2 * lambda * dist x v" apply (rule lipschitz_onD(1)[OF Lip]) using v(1) A(2) x(1) by auto also have "... \ 2 * lambda * (C/lambda)" apply (intro mono_intros) using C v(2) by auto also have "... = 2 * C" using C by (simp add: divide_simps) finally have *: "dist z (c v) \ 2 * C" by simp show "\y\c`{a..b}. dist z y \ 2 * C" apply (rule bexI[of _ "c v"]) using * v(1) A(2) by auto qed text \From the above controls, we check that $d$ is a quasi-isometry, with explicit constants.\ have "lambda (4 * C)-quasi_isometry_on {a..b} d" proof show "1 \ lambda" using C by auto show "0 \ 4 * C" using C by auto have I : "dist (d x) (d y) \ lambda * dist x y + 4 * C" if H: "x \ {a..b}" "y \ {a..b}" "x < y" for x y proof - obtain u where u: "u \ A - {b}" "x \ {u..next_in A u}" using intervals_decomposition[OF A] H(1) by force have "u \ {a.. A" using next_in_basics(1)[OF A \u \ {a..] by auto obtain v where v: "v \ A - {b}" "y \ {v..next_in A v}" using intervals_decomposition[OF A] H(2) by force have "v \ {a.. v" using u(1) next_in_basics(3)[OF A, OF \v \ {a..] by auto show ?thesis proof (cases "u = v") case True have "dist (d x) (d y) \ lambda * dist x y + C" apply (rule QI0[OF u]) using v(2) True by auto also have "... \ lambda * dist x y + 4 * C" using C by auto finally show ?thesis by simp next case False then have "u < v" using \u \ v\ by auto then have "next_in A u \ v" using v(1) next_in_basics(3)[OF A, OF \u \ {a..] by auto have d1: "d (next_in A u) = c (next_in A u)" using \next_in A u \ A\ unfolding d_def by auto have d2: "d v = c v" using v(1) unfolding d_def by auto have "dist (d x) (d y) \ dist (d x) (d (next_in A u)) + dist (d (next_in A u)) (d v) + dist (d v) (d y)" by (intro mono_intros) also have "... \ (lambda * dist x (next_in A u) + C) + (lambda * dist (next_in A u) v + C) + (lambda * dist v y + C)" apply (intro mono_intros) apply (rule QI0[OF u]) using u(2) apply simp apply (simp add: d1 d2) apply (rule quasi_isometry_onD(1)[OF assms(1)]) using \next_in A u \ A\ \A \ {a..b}\ apply auto[1] using \v \ A - {b}\ \A \ {a..b}\ apply auto[1] apply (rule QI0[OF v(1)]) using v(2) by auto also have "... = lambda * dist x y + 3 * C" unfolding dist_real_def using \x \ {u..next_in A u}\ \y \ {v..next_in A v}\ \x < y\ \next_in A u \ v\ by (auto simp add: algebra_simps) finally show ?thesis using C by simp qed qed show "dist (d x) (d y) \ lambda * dist x y + 4 * C" if H: "x \ {a..b}" "y \ {a..b}" for x y proof - consider "x < y" | "x = y" | "x > y" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using I[OF H(1) H(2) 1] by simp next case 2 show ?thesis unfolding 2 using C by auto next case 3 show ?thesis using I [OF H(2) H(1) 3] by (simp add: metric_space_class.dist_commute) qed qed text \The lower bound is more tricky. We separate the case where $x$ and $y$ are in the same interval, when they are in different nearby intervals, and when they are in different separated intervals. The latter case is more difficult. In this case, one of the intervals has length $C/\lambda$ and the other one has length at most $3/2\cdot C/\lambda$. There, we approximate $dist (d x) (d y)$ by $dist (d u') (d v')$ where $u'$ and $v'$ are suitable endpoints of the intervals containing respectively $x$ and $y$. We use the inner endpoint (between $x$ and $y$) if the distance between $x$ or $y$ and this point is less than $2/5$ of the length of the interval, and the outer endpoint otherwise. The reason is that, with the outer endpoints, we get right away an upper bound for the distance between $x$ and $y$, while this is not the case with the inner endpoints where there is an additional error. The equilibrium is reached at proportion $2/5$. \ have J : "dist (d x) (d y) \ (1/lambda) * dist x y - 4 * C" if H: "x \ {a..b}" "y \ {a..b}" "x < y" for x y proof - obtain u where u: "u \ A - {b}" "x \ {u..next_in A u}" using intervals_decomposition[OF A] H(1) by force have "u \ {a.. A" using next_in_basics(1)[OF A \u \ {a..] by auto obtain v where v: "v \ A - {b}" "y \ {v..next_in A v}" using intervals_decomposition[OF A] H(2) by force have "v \ {a.. A" using next_in_basics(1)[OF A \v \ {a..] by auto have "u < next_in A v" using H(3) u(2) v(2) by auto then have "u \ v" using u(1) next_in_basics(3)[OF A, OF \v \ {a..] by auto consider "v = u" | "v = next_in A u" | "v \ u \ v \ next_in A u" by auto then show ?thesis proof (cases) case 1 have "(1/lambda) * dist x y - 4 * C \ lambda * dist x y - 4 * C" apply (intro mono_intros) by auto also have "... \ lambda * (3/2 * C/lambda) - 3/2 * C" apply (intro mono_intros) using u(2) v(2) unfolding 1 using C gap[OF u(1)] dist_real_def \x < y\ by auto also have "... = 0" using C by auto also have "... \ dist (d x) (d y)" by auto finally show ?thesis by simp next case 2 have "dist x y \ dist x (next_in A u) + dist v y" unfolding 2 by (intro mono_intros) also have "... \ 3/2 * C/lambda + 3/2 * C/lambda" apply (intro mono_intros) unfolding dist_real_def using u(2) v(2) gap[OF u(1)] gap[OF v(1)] by auto finally have *: "dist x y \ 3 * C/lambda" by auto have "(1/lambda) * dist x y - 4 * C \ lambda * dist x y - 4 * C" apply (intro mono_intros) by auto also have "... \ lambda * (3 * C/lambda) - 3 * C" apply (intro mono_intros) using * C by auto also have "... = 0" using C by auto also have "... \ dist (d x) (d y)" by auto finally show ?thesis by simp next case 3 then have "u < v" using \u \ v\ by auto then have *: "next_in A u < v" using v(1) next_in_basics(3)[OF A \u \ {a..] 3 by auto have nu: "next_in A u = u + C/lambda" proof (rule nx) show "u \ A" using u(1) by auto show "u \ b" using u(1) by auto show "u \ b'" proof assume H: "u = b'" have "b < v" using * unfolding H nb' by simp then show False using \v \ {a.. by auto qed show "u \ p N" proof assume H: "u = p N" have "b' < v" using * unfolding H npN by simp then have "next_in A b' \ v" using next_in_basics(3)[OF A \b' \ {a..] v by force then show False unfolding nb' using \v \ {a.. by auto qed qed have nv: "next_in A v \ v + 3/2 * C/lambda" using gap[OF v(1)] by auto have d: "d u = c u" "d (next_in A u) = c (next_in A u)" "d v = c v" "d (next_in A v) = c (next_in A v)" using \u \ A - {b}\ \next_in A u \ A\ \v \ A - {b}\ \next_in A v \ A\ unfolding d_def by auto text \The interval containing $x$ has length $C/\lambda$, while the interval containing $y$ has length at most $\leq 3/2 C/\lambda$. Therefore, $x$ is at proportion $2/5$ of the inner point if $x > u + (3/5) C/\lambda$, and $y$ is at proportion $2/5$ of the inner point if $y < v + (2/5) \cdot 3/2 \cdot C/\lambda = v + (3/5)C/\lambda$.\ consider "x \ u + (3/5) * C/lambda \ y \ v + (3/5) * C/lambda" | "x \ u + (3/5) * C/lambda \ y \ v + (3/5) * C/lambda" | "x \ u + (3/5) * C/lambda \ y \ v + (3/5) * C/lambda" | "x \ u + (3/5) * C/lambda \ y \ v + (3/5) * C/lambda" by linarith then show ?thesis proof (cases) case 1 have "(1/lambda) * dist u v - C \ dist (c u) (c v)" apply (rule quasi_isometry_onD(2)[OF assms(1)]) using \u \ A - {b}\ \v \ A - {b}\ \A \ {a..b}\ by auto also have "... = dist (d u) (d v)" using d by auto also have "... \ dist (d u) (d x) + dist (d x) (d y) + dist (d y) (d v)" by (intro mono_intros) also have "... \ (2 * lambda * dist u x) + dist (d x) (d y) + (2 * lambda * dist y v)" apply (intro mono_intros) apply (rule lipschitz_onD[OF lip[OF u(1)]]) using u(2) apply auto[1] using u(2) apply auto[1] apply (rule lipschitz_onD[OF lip[OF v(1)]]) using v(2) by auto also have "... \ (2 * lambda * (3/5 * C/lambda)) + dist (d x) (d y) + (2 * lambda * (3/5 * C/lambda))" apply (intro mono_intros) unfolding dist_real_def using 1 u v C by auto also have "... = 12/5 * C + dist (d x) (d y)" using C by (auto simp add: algebra_simps divide_simps) finally have *: "(1/lambda) * dist u v \ dist (d x) (d y) + 17/5 * C" by auto have "(1/lambda) * dist x y \ (1/lambda) * (dist u v + dist v y)" apply (intro mono_intros) unfolding dist_real_def using C u(2) v(2) \x < y\ by auto also have "... \ (1/lambda) * (dist u v + 3/5 * C/lambda)" apply (intro mono_intros) unfolding dist_real_def using 1 v(2) C by auto also have "... = (1/lambda) * dist u v + 3/5 * C * (1/(lambda * lambda))" using C by (auto simp add: algebra_simps divide_simps) also have "... \ (1/lambda) * dist u v + 3/5 * C * 1" apply (intro mono_intros) using C by (auto simp add: divide_simps algebra_simps mult_ge1_powers(1)) also have "... \ (dist (d x) (d y) + 17/5 * C) + 3/5 * C * 1" using * by auto finally show ?thesis by auto next case 2 have "(1/lambda) * dist (next_in A u) v - C \ dist (c (next_in A u)) (c v)" apply (rule quasi_isometry_onD(2)[OF assms(1)]) using \next_in A u \ A\ \v \ A - {b}\ \A \ {a..b}\ by auto also have "... = dist (d (next_in A u)) (d v)" using d by auto also have "... \ dist (d (next_in A u)) (d x) + dist (d x) (d y) + dist (d y) (d v)" by (intro mono_intros) also have "... \ (2 * lambda * dist (next_in A u) x) + dist (d x) (d y) + (2 * lambda * dist y v)" apply (intro mono_intros) apply (rule lipschitz_onD[OF lip[OF u(1)]]) using u(2) apply auto[1] using u(2) apply auto[1] apply (rule lipschitz_onD[OF lip[OF v(1)]]) using v(2) by auto also have "... \ (2 * lambda * (2/5 * C/lambda)) + dist (d x) (d y) + (2 * lambda * (3/5 * C/lambda))" apply (intro mono_intros) unfolding dist_real_def using 2 u v C nu by auto also have "... = 2 * C + dist (d x) (d y)" using C by (auto simp add: algebra_simps divide_simps) finally have *: "(1/lambda) * dist (next_in A u) v \ dist (d x) (d y) + 3 * C" by auto have "(1/lambda) * dist x y \ (1/lambda) * (dist x (next_in A u) + dist (next_in A u) v + dist v y)" apply (intro mono_intros) unfolding dist_real_def using C u(2) v(2) \x < y\ by auto also have "... \ (1/lambda) * ((2/5 * C/lambda) + dist (next_in A u) v + (3/5 * C/lambda))" apply (intro mono_intros) unfolding dist_real_def using 2 u(2) v(2) C nu by auto also have "... = (1/lambda) * dist (next_in A u) v + C * (1/(lambda * lambda))" using C by (auto simp add: algebra_simps divide_simps) also have "... \ (1/lambda) * dist (next_in A u) v + C * 1" apply (intro mono_intros) using C by (auto simp add: divide_simps algebra_simps mult_ge1_powers(1)) also have "... \ (dist (d x) (d y) + 3 * C) + C * 1" using * by auto finally show ?thesis by auto next case 3 have "(1/lambda) * dist u (next_in A v) - C \ dist (c u) (c (next_in A v))" apply (rule quasi_isometry_onD(2)[OF assms(1)]) using \u \ A - {b}\ \next_in A v \ A\ \A \ {a..b}\ by auto also have "... = dist (d u) (d (next_in A v))" using d by auto also have "... \ dist (d u) (d x) + dist (d x) (d y) + dist (d y) (d (next_in A v))" by (intro mono_intros) also have "... \ (2 * lambda * dist u x) + dist (d x) (d y) + (2 * lambda * dist y (next_in A v))" apply (intro mono_intros) apply (rule lipschitz_onD[OF lip[OF u(1)]]) using u(2) apply auto[1] using u(2) apply auto[1] apply (rule lipschitz_onD[OF lip[OF v(1)]]) using v(2) by auto also have "... \ (2 * lambda * (3/5 * C/lambda)) + dist (d x) (d y) + (2 * lambda * (9/10 * C/lambda))" apply (intro mono_intros) unfolding dist_real_def using 3 u v C nv by auto also have "... = 3 * C + dist (d x) (d y)" using C by (auto simp add: algebra_simps divide_simps) finally have *: "(1/lambda) * dist u (next_in A v) \ dist (d x) (d y) + 4 * C" by auto have "(1/lambda) * dist x y \ (1/lambda) * dist u (next_in A v)" apply (intro mono_intros) unfolding dist_real_def using C u(2) v(2) \x < y\ by auto also have "... \ dist (d x) (d y) + 4 * C" using * by auto finally show ?thesis by auto next case 4 have "(1/lambda) * dist (next_in A u) (next_in A v) - C \ dist (c (next_in A u)) (c (next_in A v))" apply (rule quasi_isometry_onD(2)[OF assms(1)]) using \next_in A u \ A\ \next_in A v \ A\ \A \ {a..b}\ by auto also have "... = dist (d (next_in A u)) (d (next_in A v))" using d by auto also have "... \ dist (d (next_in A u)) (d x) + dist (d x) (d y) + dist (d y) (d (next_in A v))" by (intro mono_intros) also have "... \ (2 * lambda * dist (next_in A u) x) + dist (d x) (d y) + (2 * lambda * dist y (next_in A v))" apply (intro mono_intros) apply (rule lipschitz_onD[OF lip[OF u(1)]]) using u(2) apply auto[1] using u(2) apply auto[1] apply (rule lipschitz_onD[OF lip[OF v(1)]]) using v(2) by auto also have "... \ (2 * lambda * (2/5 * C/lambda)) + dist (d x) (d y) + (2 * lambda * (9/10 * C/lambda))" apply (intro mono_intros) unfolding dist_real_def using 4 u v C nu nv by auto also have "... = 13/5 * C + dist (d x) (d y)" using C by (auto simp add: algebra_simps divide_simps) finally have *: "(1/lambda) * dist (next_in A u) (next_in A v) \ dist (d x) (d y) + 18/5 * C" by auto have "(1/lambda) * dist x y \ (1/lambda) * (dist x (next_in A u) + dist (next_in A u) (next_in A v))" apply (intro mono_intros) unfolding dist_real_def using C u(2) v(2) \x < y\ by auto also have "... \ (1/lambda) * ((2/5 *C/lambda) + dist (next_in A u) (next_in A v))" apply (intro mono_intros) unfolding dist_real_def using 4 u(2) v(2) C nu by auto also have "... = (1/lambda) * dist (next_in A u) (next_in A v) + 2/5 * C * (1/(lambda * lambda))" using C by (auto simp add: algebra_simps divide_simps) also have "... \ (1/lambda) * dist (next_in A u) (next_in A v) + 2/5 * C * 1" apply (intro mono_intros) using C by (auto simp add: divide_simps algebra_simps mult_ge1_powers(1)) also have "... \ (dist (d x) (d y) + 18/5 * C) + 2/5 * C * 1" using * by auto finally show ?thesis by auto qed qed qed show "dist (d x) (d y) \ (1/lambda) * dist x y - 4 * C" if H: "x \ {a..b}" "y \ {a..b}" for x y proof - consider "x < y" | "x = y" | "x > y" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using J[OF H(1) H(2) 1] by simp next case 2 show ?thesis unfolding 2 using C by auto next case 3 show ?thesis using J[OF H(2) H(1) 3] by (simp add: metric_space_class.dist_commute) qed qed qed text \We have proved that $d$ has all the properties we wanted.\ then have "continuous_on {a..b} d \ d a = c a \ d b = c b \ lambda (4 * C)-quasi_isometry_on {a..b} d \ (\x\{a..b}. dist (c x) (d x) \ 4 *C) \ (2*lambda)-lipschitz_on {a..b} d \ hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C" using dist_c_d \continuous_on {a..b} d\ \d a = c a\ \d b = c b\ \(2*lambda)-lipschitz_on {a..b} d\ \hausdorff_distance (c`{a..b}) (d`{a..b}) \ 2 * C\ by auto then show ?thesis by auto qed qed end (*of theory Isometries*) diff --git a/thys/Gromov_Hyperbolicity/Isometries_Classification.thy b/thys/Gromov_Hyperbolicity/Isometries_Classification.thy --- a/thys/Gromov_Hyperbolicity/Isometries_Classification.thy +++ b/thys/Gromov_Hyperbolicity/Isometries_Classification.thy @@ -1,1115 +1,1117 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Classification of isometries on a Gromov hyperbolic space\ theory Isometries_Classification imports Gromov_Boundary Busemann_Function begin text \Isometries of Gromov hyperbolic spaces are of three types: \begin{itemize} \item Elliptic ones, for which orbits are bounded. \item Parabolic ones, which are not elliptic and have exactly one fixed point at infinity. \item Loxodromic ones, which are not elliptic and have exactly two fixed points at infinity. \end{itemize} In this file, we show that all isometries are indeed of this form, and give further properties for each type. For the definition, we use another characterization in terms of stable translation length: for isometries which are not elliptic, then they are parabolic if the stable translation length is $0$, loxodromic if it is positive. This gives a very efficient definition, and it is clear from this definition that the three categories of isometries are disjoint. All the work is then to go from this general definition to the dynamical properties in terms of fixed points on the boundary. \ subsection \The translation length\ text \The translation length is the minimal translation distance of an isometry. The stable translation length is the limit of the translation length of $f^n$ divided by $n$.\ definition translation_length::"(('a::metric_space) \ 'a) \ real" where "translation_length f = Inf {dist x (f x)|x. True}" lemma translation_length_nonneg [simp, mono_intros]: "translation_length f \ 0" unfolding translation_length_def by (rule cInf_greatest, auto) lemma translation_length_le [mono_intros]: "translation_length f \ dist x (f x)" unfolding translation_length_def apply (rule cInf_lower) by (auto intro: bdd_belowI[of _ 0]) definition stable_translation_length::"(('a::metric_space) \ 'a) \ real" where "stable_translation_length f = Inf {translation_length (f^^n)/n |n. n > 0}" lemma stable_translation_length_nonneg [simp]: "stable_translation_length f \ 0" unfolding stable_translation_length_def by (rule cInf_greatest, auto) lemma stable_translation_length_le_translation_length [mono_intros]: "n * stable_translation_length f \ translation_length (f^^n)" proof - have *: "stable_translation_length f \ translation_length (f^^n)/n" if "n > 0" for n unfolding stable_translation_length_def apply (rule cInf_lower) using that by (auto intro: bdd_belowI[of _ 0]) show ?thesis apply (cases "n = 0") using * by (auto simp add: divide_simps algebra_simps) qed lemma semicontraction_iterates: fixes f::"('a::metric_space) \ 'a" assumes "1-lipschitz_on UNIV f" shows "1-lipschitz_on UNIV (f^^n)" by (induction n, auto intro!: lipschitz_onI lipschitz_on_compose2[of 1 UNIV _ 1 f, simplified] lipschitz_on_subset[OF assms]) text \If $f$ is a semicontraction, then its stable translation length is the limit of $d(x, f^n x)/n$ for any $n$. While it is obvious that the liminf of this quantity is at least the stable translation length (which is defined as an inf over all points and all times), the opposite inequality is more interesting. One may find a point $y$ and a time $k$ for which $d(y, f^k y)/k$ is very close to the stable translation length. By subadditivity of the sequence $n \mapsto f(y, f^n y)$ and Fekete's Lemma, it follows that, for any large $n$, then $d(y, f^n y)/n$ is also very close to the stable translation length. Since this is equal to $d(x, f^n x)/n$ up to $\pm 2 d(x,y)/n$, the result follows.\ proposition stable_translation_length_as_pointwise_limit: assumes "1-lipschitz_on UNIV f" shows "(\n. dist x ((f^^n) x)/n) \ stable_translation_length f" proof - have *: "subadditive (\n. dist y ((f^^n) y))" for y proof (rule subadditiveI) fix m n::nat have "dist y ((f ^^ (m + n)) y) \ dist y ((f^^m) y) + dist ((f^^m) y) ((f^^(m+n)) y)" by (rule dist_triangle) also have "... = dist y ((f^^m) y) + dist ((f^^m) y) ((f^^m) ((f^^n) y))" by (auto simp add: funpow_add) also have "... \ dist y ((f^^m) y) + dist y ((f^^n) y)" using semicontraction_iterates[OF assms, of m] unfolding lipschitz_on_def by auto finally show "dist y ((f ^^ (m + n)) y) \ dist y ((f ^^ m) y) + dist y ((f ^^ n) y)" by simp qed have Ly: "(\n. dist y ((f^^n) y) / n) \ Inf {dist y ((f^^n) y) / n |n. n > 0}" for y by (auto intro!: bdd_belowI[of _ 0] subadditive_converges_bounded'[OF subadditive_imp_eventually_subadditive[OF *]]) have "eventually (\n. dist x ((f^^n) x)/n < l) sequentially" if "stable_translation_length f < l" for l proof - obtain m where m: "stable_translation_length f < m" "m < l" using \stable_translation_length f < l\ dense by auto have "\t \ {translation_length (f^^n)/n |n. n > 0}. t < m" apply (subst cInf_less_iff[symmetric]) using m unfolding stable_translation_length_def by (auto intro!: bdd_belowI[of _ 0]) then obtain k where k: "k > 0" "translation_length (f^^k)/k < m" by auto have "translation_length (f^^k) < k * m" using k by (simp add: divide_simps algebra_simps) then have "\t \ {dist y ((f^^k) y) |y. True}. t < k * m" apply (subst cInf_less_iff[symmetric]) unfolding translation_length_def by (auto intro!: bdd_belowI[of _ 0]) then obtain y where y: "dist y ((f^^k) y) < k * m" by auto have A: "eventually (\n. dist y ((f^^n) y)/n < m) sequentially" apply (auto intro!: order_tendstoD[OF Ly] iffD2[OF cInf_less_iff] bdd_belowI[of _ 0] exI[of _ "dist y ((f^^k) y)/k"]) using y k by (auto simp add: algebra_simps divide_simps) have B: "eventually (\n. dist x y * (1/n) < (l-m)/2) sequentially" apply (intro order_tendstoD[of _ "dist x y * 0"] tendsto_intros) using \m < l\ by simp have C: "dist x ((f^^n) x)/n < l" if "n > 0" "dist y ((f^^n) y)/n < m" "dist x y * (1/n) < (l-m)/2" for n proof - have "dist x ((f^^n) x) \ dist x y + dist y ((f^^n) y) + dist ((f^^n) y) ((f^^n) x)" by (intro mono_intros) also have "... \ dist x y + dist y ((f^^n) y) + dist y x" using semicontraction_iterates[OF assms, of n] unfolding lipschitz_on_def by auto also have "... = 2 * dist x y + dist y ((f^^n) y)" by (simp add: dist_commute) also have "... < 2 * real n * (l-m)/2 + n * m" apply (intro mono_intros) using that by (auto simp add: algebra_simps divide_simps) also have "... = n * l" by (simp add: algebra_simps divide_simps) finally show ?thesis using that by (simp add: algebra_simps divide_simps) qed show "eventually (\n. dist x ((f^^n) x)/n < l) sequentially" by (rule eventually_mono[OF eventually_conj[OF eventually_conj[OF A B] eventually_gt_at_top[of 0]] C], auto) qed moreover have "eventually (\n. dist x ((f^^n) x)/n > l) sequentially" if "stable_translation_length f > l" for l proof - have *: "dist x ((f^^n) x)/n > l" if "n > 0" for n proof - have "n * l < n * stable_translation_length f" using \stable_translation_length f > l\ \n > 0\ by auto also have "... \ translation_length (f^^n)" by (intro mono_intros) also have "... \ dist x ((f^^n) x)" by (intro mono_intros) finally show ?thesis using \n > 0\ by (auto simp add: algebra_simps divide_simps) qed then show ?thesis by (rule eventually_mono[rotated], auto) qed ultimately show ?thesis by (rule order_tendstoI[rotated]) qed text \It follows from the previous proposition that the stable translation length is also the limit of the renormalized translation length of $f^n$.\ proposition stable_translation_length_as_limit: assumes "1-lipschitz_on UNIV f" shows "(\n. translation_length (f^^n) / n) \ stable_translation_length f" proof - obtain x::'a where True by auto show ?thesis proof (rule tendsto_sandwich[of "\n. stable_translation_length f" _ _ "\n. dist x ((f^^n) x)/n"]) have "stable_translation_length f \ translation_length (f ^^ n) / real n" if "n > 0" for n using stable_translation_length_le_translation_length[of n f] that by (simp add: divide_simps algebra_simps) then show "eventually (\n. stable_translation_length f \ translation_length (f ^^ n) / real n) sequentially" by (rule eventually_mono[rotated], auto) have "translation_length (f ^^ n) / real n \ dist x ((f ^^ n) x) / real n" for n using translation_length_le[of "f^^n" x] by (auto simp add: divide_simps) then show "eventually (\n. translation_length (f ^^ n) / real n \ dist x ((f ^^ n) x) / real n) sequentially" by auto qed (auto simp add: stable_translation_length_as_pointwise_limit[OF assms]) qed lemma stable_translation_length_inv: assumes "isometry f" shows "stable_translation_length (inv f) = stable_translation_length f" proof - have *: "dist basepoint (((inv f)^^n) basepoint) = dist basepoint ((f^^n) basepoint)" for n proof - have "basepoint = (f^^n) (((inv f)^^n) basepoint)" by (metis assms comp_apply fn_o_inv_fn_is_id isometry_inverse(2)) then have "dist basepoint ((f^^n) basepoint) = dist ((f^^n) (((inv f)^^n) basepoint)) ((f^^n) basepoint)" by auto also have "... = dist (((inv f)^^n) basepoint) basepoint" unfolding isometryD(2)[OF isometry_iterates[OF assms]] by simp finally show ?thesis by (simp add: dist_commute) qed have "(\n. dist basepoint ((f^^n) basepoint)/n) \ stable_translation_length f" using stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF assms]] by simp moreover have "(\n. dist basepoint ((f^^n) basepoint)/n) \ stable_translation_length (inv f)" unfolding *[symmetric] using stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF isometry_inverse(1)[OF assms]]] by simp ultimately show ?thesis using LIMSEQ_unique by auto qed subsection \The strength of an isometry at a fixed point at infinity\ text \The additive strength of an isometry at a fixed point at infinity is the asymptotic average every point is moved towards the fixed point at each step. It is measured using the Busemann function.\ definition additive_strength::"('a::Gromov_hyperbolic_space \ 'a) \ ('a Gromov_completion) \ real" where "additive_strength f xi = lim (\n. (Busemann_function_at xi ((f^^n) basepoint) basepoint)/n)" text \For additivity reasons, as the Busemann function is a quasi-morphism, the additive strength measures the deplacement even at finite times. It is also uniform in terms of the basepoint. This shows that an isometry sends horoballs centered at a fixed point to horoballs, up to a uniformly bounded error depending only on $\delta$.\ lemma Busemann_function_eq_additive_strength: assumes "isometry f" "Gromov_extension f xi = xi" shows "\Busemann_function_at xi ((f^^n) x) (x::'a::Gromov_hyperbolic_space) - real n * additive_strength f xi\ \ 2 * deltaG(TYPE('a))" proof - define u where "u = (\y n. Busemann_function_at xi ((f^^n) y) y)" have *: "abs(u y (m+n) - u y m - u y n) \ 2 * deltaG(TYPE('a))" for n m y proof - have P: "Gromov_extension (f^^m) xi = xi" unfolding Gromov_extension_isometry_iterates[OF assms(1)] apply (induction m) using assms by auto have *: "u y n = Busemann_function_at xi ((f^^m) ((f^^n) y)) ((f^^m) y)" apply (subst P[symmetric]) unfolding Busemann_function_isometry[OF isometry_iterates[OF \isometry f\]] u_def by auto show ?thesis unfolding * unfolding u_def using Busemann_function_quasi_morphism[of xi "(f^^(m+n)) y" "(f^^m) y" y] unfolding funpow_add comp_def by auto qed define l where "l = (\y. lim (\n. u y n/n))" have A: "abs(u y k - k * l y) \ 2 * deltaG(TYPE('a))" for y k unfolding l_def using almost_additive_converges(2) * by auto then have *: "abs(Busemann_function_at xi ((f^^k) y) y - k * l y) \ 2 * deltaG(TYPE('a))" for y k unfolding u_def by auto have "l basepoint = additive_strength f xi" unfolding l_def u_def additive_strength_def by auto have "abs(k * l basepoint - k * l x) \ 4 * deltaG(TYPE('a)) + 2 * dist basepoint x" for k::nat proof - have "abs(k * l basepoint - k * l x) = abs((Busemann_function_at xi ((f^^k) x) x - k * l x) - (Busemann_function_at xi ((f^^k) basepoint) basepoint - k * l basepoint) + (Busemann_function_at xi ((f^^k) basepoint) basepoint - Busemann_function_at xi ((f^^k) x) x))" by auto also have "... \ abs (Busemann_function_at xi ((f^^k) x) x - k * l x) + abs (Busemann_function_at xi ((f^^k) basepoint) basepoint - k * l basepoint) + abs (Busemann_function_at xi ((f^^k) basepoint) basepoint - Busemann_function_at xi ((f^^k) x) x)" by auto also have "... \ 2 * deltaG(TYPE('a)) + 2 * deltaG(TYPE('a)) + (dist ((f^^k) basepoint) ((f^^k) x) + dist basepoint x)" by (intro mono_intros *) also have "... = 4 * deltaG(TYPE('a)) + 2 * dist basepoint x" unfolding isometryD[OF isometry_iterates[OF assms(1)]] by auto finally show ?thesis by auto qed moreover have "u = v" if H: "\k::nat. abs(k * u - k * v) \ C" for u v C::real proof - have "(\n. abs(u - v)) \ 0" proof (rule tendsto_sandwich[of "\n. 0" _ _ "\n::nat. C/n"], auto) have "(\n. C*(1/n)) \ C * 0" by (intro tendsto_intros) then show "(\n. C/n) \ 0" by auto have "\u - v\ \ C / real n" if "n \ 1" for n using H[of n] that apply (simp add: divide_simps algebra_simps) by (metis H abs_mult abs_of_nat right_diff_distrib') then show "\\<^sub>F n in sequentially. \u - v\ \ C / real n" unfolding eventually_sequentially by auto qed then show ?thesis by (metis LIMSEQ_const_iff abs_0_eq eq_iff_diff_eq_0) qed ultimately have "l basepoint = l x" by auto show ?thesis using A[of x n] unfolding u_def \l basepoint = l x\[symmetric] \l basepoint = additive_strength f xi\ by auto qed lemma additive_strength_as_limit [tendsto_intros]: assumes "isometry f" "Gromov_extension f xi = xi" shows "(\n. Busemann_function_at xi ((f^^n) x) x/n) \ additive_strength f xi" proof - have "(\n. abs(Busemann_function_at xi ((f^^n) x) x/n - additive_strength f xi)) \ 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. (2 * deltaG(TYPE('a))) * (1/real n)"], auto) unfolding eventually_sequentially apply (rule exI[of _ 1]) using Busemann_function_eq_additive_strength[OF assms] apply (simp add: divide_simps algebra_simps) using tendsto_mult[OF _ lim_1_over_n] by auto then show ?thesis using LIM_zero_iff tendsto_rabs_zero_cancel by blast qed text \The additive strength measures the amount of displacement towards a fixed point at infinity. Therefore, the distance from $x$ to $f^n x$ is at least $n$ times the additive strength, but one might think that it might be larger, if there is displacement along the horospheres. It turns out that this is not the case: the displacement along the horospheres is at most logarithmic (this is a classical property of parabolic isometries in hyperbolic spaces), and in fact it is bounded for loxodromic elements. We prove here that the growth is at most logarithmic in all cases, using a small computation based on the hyperbolicity inequality, expressed in Lemma \verb+dist_minus_Busemann_max_ineq+ above. This lemma will be used below to show that the translation length is the absolute value of the additive strength.\ lemma dist_le_additive_strength: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" "n \ 1" shows "dist x ((f^^n) x) \ dist x (f x) + real n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" proof - have A: "\n. n \ {1..2^k} \ dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + k * 16 * deltaG(TYPE('a))" for k proof (induction k) case 0 fix n::nat assume "n \ {1..2^0}" then have "n = 1" by auto then show "dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + real 0 * 16 * deltaG(TYPE('a))" using assms(3) by auto next case (Suc k) fix N::nat assume "N \ {1..2^(Suc k)}" then consider "N \ {1..2^k}" | "N \ {2^k<..2^(Suc k)}" using not_le by auto then show "dist x ((f ^^ N) x) - real N * additive_strength f xi \ dist x (f x) + real (Suc k) * 16 * deltaG TYPE('a)" proof (cases) case 1 show ?thesis by (rule order_trans[OF Suc.IH[OF 1]], auto simp add: algebra_simps) next case 2 define m::nat where "m = N - 2^k" define n::nat where "n = 2^k" have nm: "N = n+m" "m \ {1..2^k}" "n \ {1..2^k}"unfolding m_def n_def using 2 by auto have *: "(f^^(n+m)) x = (f^^n) ((f^^m) x)" unfolding funpow_add comp_def by auto have **: "(f^^(n+m)) x = (f^^m) ((f^^n) x)" apply (subst add.commute) unfolding funpow_add comp_def by auto have "dist x ((f^^N) x) - N * additive_strength f xi - 2 * deltaG(TYPE('a)) \ dist x ((f^^(n+m)) x) - Busemann_function_at xi ((f^^(n+m)) x) x" unfolding nm(1) using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of "n+m" x] by auto also have "... \ max (dist x ((f^^n) x) - Busemann_function_at xi ((f^^n) x) x) (dist ((f^^n) x) ((f^^(n+m)) x) - Busemann_function_at xi ((f^^(n+m)) x) ((f^^n) x) - 2 * Busemann_function_at xi ((f^^n) x) x) + 8 * deltaG(TYPE('a))" using dist_minus_Busemann_max_ineq by auto also have "... \ max (dist x ((f^^n) x) - (n * additive_strength f xi - 2 * deltaG(TYPE('a)))) (dist ((f^^n) x) ((f^^(n+m)) x) - (m * additive_strength f xi - 2 * deltaG(TYPE('a))) - 2 * (n * additive_strength f xi - 2 * deltaG(TYPE('a)))) + 8 * deltaG(TYPE('a))" unfolding ** apply (intro mono_intros) using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of n x] Busemann_function_eq_additive_strength[OF assms(1) assms(2), of m "(f^^n) x"] by auto also have "... \ max (dist x ((f^^n) x) - n * additive_strength f xi + 6 * deltaG(TYPE('a))) (dist x ((f^^m) x) - m * additive_strength f xi + 6 * deltaG(TYPE('a))) + 8 * deltaG(TYPE('a))" unfolding * isometryD(2)[OF isometry_iterates[OF assms(1)], of n] using assms(3) by (intro mono_intros, auto) also have "... = max (dist x ((f^^n) x) - n * additive_strength f xi) (dist x ((f^^m) x) - m * additive_strength f xi) + 14 * deltaG(TYPE('a))" unfolding max_add_distrib_left[symmetric] by auto also have "... \ dist x (f x) + k * 16 * deltaG(TYPE('a)) + 14 * deltaG(TYPE('a))" using nm by (auto intro!: Suc.IH) finally show ?thesis by (auto simp add: algebra_simps) qed qed define k::nat where "k = nat(ceiling (log 2 n))" have "n \ 2^k" unfolding k_def by (meson less_log2_of_power not_le real_nat_ceiling_ge) then have "dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + k * 16 * deltaG(TYPE('a))" using A[of n k] \n \ 1\ by auto moreover have "real (nat \log 2 (real n)\) = real_of_int \log 2 (real n)\" by (metis Transcendental.log_one \n \ 2 ^ k\ assms(4) ceiling_zero int_ops(2) k_def le_antisym nat_eq_iff2 of_int_1 of_int_of_nat_eq order_refl power_0) ultimately show ?thesis unfolding k_def by (auto simp add: algebra_simps) qed text \The strength of the inverse of a map is the opposite of the strength of the map.\ lemma additive_strength_inv: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" shows "additive_strength (inv f) xi = - additive_strength f xi" proof - have *: "(inv f ^^ n) ((f ^^ n) x) = x" for n x by (metis assms(1) comp_apply funpow_code_def inv_fn_o_fn_is_id isometry_inverse(2)) have A: "abs(real n * additive_strength f xi + real n * additive_strength (inv f) xi) \ 6 * deltaG (TYPE('a))" for n::nat and x::'a using Busemann_function_quasi_morphism[of xi x "(f^^n) x" x] Busemann_function_eq_additive_strength[OF assms, of n x] Busemann_function_eq_additive_strength[OF isometry_inverse(1)[OF assms(1)] Gromov_extension_inv_fixed_point[OF assms], of n "(f^^n) x"] unfolding * by auto have B: "abs(additive_strength f xi + additive_strength (inv f) xi) \ 6 * deltaG(TYPE('a)) * (1/n)" if "n \ 1" for n::nat using that A[of n] apply (simp add: divide_simps algebra_simps) unfolding distrib_left[symmetric] by auto have "(\n. abs(additive_strength f xi + additive_strength (inv f) xi)) \ 6 * deltaG(TYPE('a)) * 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. 6 * deltaG(TYPE('a)) * (1/real n)"], simp) unfolding eventually_sequentially apply (rule exI[of _ 1]) using B apply simp by (simp, intro tendsto_intros) then show ?thesis using LIMSEQ_unique mult_zero_right tendsto_const by fastforce qed text \We will now prove that the stable translation length of an isometry is given by the absolute value of its strength at any fixed point. We start with the case where the strength is nonnegative, and then reduce to this case by considering the map or its inverse.\ lemma stable_translation_length_eq_additive_strength_aux: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" shows "stable_translation_length f = additive_strength f xi" proof - have "(\n. dist x ((f^^n) x)/n) \ additive_strength f xi" for x proof (rule tendsto_sandwich[of "\n. (n * additive_strength f xi - 2 * deltaG(TYPE('a)))/real n" _ _ "\n. (dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a)))/ n"]) have "n * additive_strength f xi - 2 * deltaG TYPE('a) \ dist x ((f ^^ n) x)" for n using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of n x] Busemann_function_le_dist[of xi "(f^^n) x" x] by (simp add: dist_commute) then have "(n * additive_strength f xi - 2 * deltaG TYPE('a)) / n \ dist x ((f ^^ n) x) / n" if "n \ 1" for n using that by (simp add: divide_simps) then show "\\<^sub>F n in sequentially. (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n \ dist x ((f ^^ n) x) / real n" unfolding eventually_sequentially by auto have A: "eventually (\n. additive_strength f xi - (2 * deltaG(TYPE('a))) * (1/n) = (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n) sequentially" unfolding eventually_sequentially apply (rule exI[of _ 1]) by (simp add: divide_simps) have B: "(\n. additive_strength f xi - (2 * deltaG(TYPE('a))) * (1/n)) \ additive_strength f xi - (2 * deltaG(TYPE('a))) * 0" by (intro tendsto_intros) show "(\n. (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n) \ additive_strength f xi" using Lim_transform_eventually[OF A B] by simp have "dist x ((f^^n) x) \ dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" if "n \ 1" for n using dist_le_additive_strength[OF assms that] by simp then have "(dist x ((f^^n) x))/n \ (dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a)))/n" if "n \ 1" for n using that by (simp add: divide_simps) then show "\\<^sub>F n in sequentially. dist x ((f ^^ n) x) / real n \ (dist x (f x) + real n * additive_strength f xi + real_of_int (\log 2 (real n)\ * 16) * deltaG TYPE('a)) / real n" unfolding eventually_sequentially by auto have A: "eventually (\n. dist x (f x) * (1/n) + additive_strength f xi + 16 * deltaG TYPE('a) * (\log 2 n\ / n) = (dist x (f x) + real n * additive_strength f xi + real_of_int (\log 2 (real n)\ * 16) * deltaG TYPE('a)) / real n) sequentially" unfolding eventually_sequentially apply (rule exI[of _ 1]) by (simp add: algebra_simps divide_simps) have B: "(\n. dist x (f x) * (1/n) + additive_strength f xi + 16 * deltaG TYPE('a) * (\log 2 n\ / n)) \ dist x (f x) * 0 + additive_strength f xi + 16 * deltaG TYPE('a) * 0" by (intro tendsto_intros) show "(\n. (dist x (f x) + n * additive_strength f xi + real_of_int (\log 2 n\ * 16) * deltaG TYPE('a)) / real n) \ additive_strength f xi" using Lim_transform_eventually[OF A B] by simp qed then show ?thesis using LIMSEQ_unique stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF assms(1)]] by blast qed lemma stable_translation_length_eq_additive_strength: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" shows "stable_translation_length f = abs(additive_strength f xi)" proof (cases "additive_strength f xi \ 0") case True then show ?thesis using stable_translation_length_eq_additive_strength_aux[OF assms] by auto next case False then have *: "abs(additive_strength f xi) = additive_strength (inv f) xi" unfolding additive_strength_inv[OF assms] by auto show ?thesis unfolding * stable_translation_length_inv[OF assms(1), symmetric] using stable_translation_length_eq_additive_strength_aux[OF isometry_inverse(1)[OF assms(1)] Gromov_extension_inv_fixed_point[OF assms]] * by auto qed subsection \Elliptic isometries\ text \Elliptic isometries are the simplest ones: they have bounded orbits.\ definition elliptic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "elliptic_isometry f = (isometry f \ (\x. bounded {(f^^n) x|n. True}))" lemma elliptic_isometryD: assumes "elliptic_isometry f" shows "bounded {(f^^n) x |n. True}" "isometry f" using assms unfolding elliptic_isometry_def by auto lemma elliptic_isometryI [intro]: assumes "bounded {(f^^n) x |n. True}" "isometry f" shows "elliptic_isometry f" proof - have "bounded {(f^^n) y |n. True}" for y proof - obtain c e where c: "\n. dist c ((f^^n) x) \ e" using assms(1) unfolding bounded_def by auto have "dist c ((f^^n) y) \ e + dist x y" for n proof - have "dist c ((f^^n) y) \ dist c ((f^^n) x) + dist ((f^^n) x) ((f^^n) y)" by (intro mono_intros) also have "... \ e + dist x y" using c[of n] isometry_iterates[OF assms(2), of n] by (intro mono_intros, auto simp add: isometryD) finally show ?thesis by simp qed then show ?thesis unfolding bounded_def by auto qed then show ?thesis unfolding elliptic_isometry_def using assms by auto qed text \The inverse of an elliptic isometry is an elliptic isometry.\ lemma elliptic_isometry_inv: assumes "elliptic_isometry f" shows "elliptic_isometry (inv f)" proof - obtain c e where A: "\n. dist c ((f^^n) basepoint) \ e" using elliptic_isometryD(1)[OF assms, of basepoint] unfolding bounded_def by auto have "c = (f^^n) (((inv f)^^n) c)" for n using fn_o_inv_fn_is_id[OF isometry_inverse(2)[OF elliptic_isometryD(2)[OF assms]], of n] unfolding comp_def by metis then have "dist ((f^^n) (((inv f)^^n) c)) ((f^^n) basepoint) \ e" for n using A by auto then have B: "dist basepoint (((inv f)^^n) c) \ e" for n unfolding isometryD(2)[OF isometry_iterates[OF elliptic_isometryD(2)[OF assms]]] by (auto simp add: dist_commute) show ?thesis apply (rule elliptic_isometryI[of _ c]) using isometry_inverse(1)[OF elliptic_isometryD(2)[OF assms]] B unfolding bounded_def by auto qed text \The inverse of a bijective map is an elliptic isometry if and only if the original map is.\ lemma elliptic_isometry_inv_iff: assumes "bij f" shows "elliptic_isometry (inv f) \ elliptic_isometry f" using elliptic_isometry_inv[of f] elliptic_isometry_inv[of "inv f"] inv_inv_eq[OF assms] by auto text \The identity is an elliptic isometry.\ lemma elliptic_isometry_id: "elliptic_isometry id" by (intro elliptic_isometryI isometryI, auto) text \The translation length of an elliptic isometry is $0$.\ lemma elliptic_isometry_stable_translation_length: assumes "elliptic_isometry f" shows "stable_translation_length f = 0" proof - obtain x::'a where True by auto have "bounded {(f^^n) x|n. True}" using elliptic_isometryD[OF assms] by auto then obtain c C where cC: "\n. dist c ((f^^n) x) \ C" unfolding bounded_def by auto have "(\n. dist x ((f^^n) x)/n) \ 0" proof (rule tendsto_sandwich[of "\_. 0" _ sequentially "\n. 2 * C / n"]) have "(\n. 2 * C * (1 / real n)) \ 2 * C * 0" by (intro tendsto_intros) then show "(\n. 2 * C / real n) \ 0" by auto have "dist x ((f ^^ n) x) / real n \ 2 * C / real n" for n using cC[of 0] cC[of n] dist_triangle[of x "(f^^n) x" c] by (auto simp add: algebra_simps divide_simps dist_commute) then show "eventually (\n. dist x ((f ^^ n) x) / real n \ 2 * C / real n) sequentially" by auto qed (auto) moreover have "(\n. dist x ((f^^n) x)/n) \ stable_translation_length f" by (rule stable_translation_length_as_pointwise_limit[OF isometry_on_lipschitz[OF isometryD(1)[OF elliptic_isometryD(2)[OF assms]]]]) ultimately show ?thesis using LIMSEQ_unique by auto qed text \If an isometry has a fixed point, then it is elliptic.\ lemma isometry_with_fixed_point_is_elliptic: assumes "isometry f" "f x = x" shows "elliptic_isometry f" proof - have *: "(f^^n) x = x" for n apply (induction n) using assms(2) by auto show ?thesis apply (rule elliptic_isometryI[of _ x, OF _ assms(1)]) unfolding * by auto qed subsection \Parabolic and loxodromic isometries\ text \An isometry is parabolic if it is not elliptic and if its translation length vanishes.\ definition parabolic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "parabolic_isometry f = (isometry f \ \elliptic_isometry f \ stable_translation_length f = 0)" text \An isometry is loxodromic if it is not elliptic and if its translation length is nonzero.\ definition loxodromic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "loxodromic_isometry f = (isometry f \ \elliptic_isometry f \ stable_translation_length f \ 0)" text \The main features of such isometries are expressed in terms of their fixed points at infinity. We define them now, but proving that the definitions make sense will take some work.\ definition neutral_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "neutral_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi = 0)" definition attracting_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "attracting_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi < 0)" definition repelling_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "repelling_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi > 0)" lemma parabolic_isometryD: assumes "parabolic_isometry f" shows "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f = 0" "\elliptic_isometry f" using assms unfolding parabolic_isometry_def by auto lemma parabolic_isometryI: assumes "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f = 0" shows "parabolic_isometry f" using assms unfolding parabolic_isometry_def elliptic_isometry_def by auto lemma loxodromic_isometryD: assumes "loxodromic_isometry f" shows "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f > 0" "\elliptic_isometry f" using assms unfolding loxodromic_isometry_def by (auto, meson dual_order.antisym not_le stable_translation_length_nonneg) text \To have a loxodromic isometry, it suffices to know that the stable translation length is nonzero, as elliptic isometries have zero translation length.\ lemma loxodromic_isometryI: assumes "isometry f" "stable_translation_length f \ 0" shows "loxodromic_isometry f" using assms elliptic_isometry_stable_translation_length unfolding loxodromic_isometry_def by auto text \Any isometry is elliptic, or parabolic, or loxodromic, and these possibilities are mutually exclusive.\ lemma elliptic_or_parabolic_or_loxodromic: assumes "isometry f" shows "elliptic_isometry f \ parabolic_isometry f \ loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma elliptic_imp_not_parabolic_loxodromic: assumes "elliptic_isometry f" shows "\parabolic_isometry f" "\loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma parabolic_imp_not_elliptic_loxodromic: assumes "parabolic_isometry f" shows "\elliptic_isometry f" "\loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma loxodromic_imp_not_elliptic_parabolic: assumes "loxodromic_isometry f" shows "\elliptic_isometry f" "\parabolic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto text \The inverse of a parabolic isometry is parabolic.\ lemma parabolic_isometry_inv: assumes "parabolic_isometry f" shows "parabolic_isometry (inv f)" unfolding parabolic_isometry_def using isometry_inverse[of f] elliptic_isometry_inv_iff[of f] parabolic_isometryD[OF assms] stable_translation_length_inv[of f] by auto text \The inverse of a loxodromic isometry is loxodromic.\ lemma loxodromic_isometry_inv: assumes "loxodromic_isometry f" shows "loxodromic_isometry (inv f)" unfolding loxodromic_isometry_def using isometry_inverse[of f] elliptic_isometry_inv_iff[of f] loxodromic_isometryD[OF assms] stable_translation_length_inv[of f] by auto text \We will now prove that an isometry which is not elliptic has a fixed point at infinity. This is very easy if the space is proper (ensuring that the Gromov completion is compact), but in fact this holds in general. One constructs it by considering a sequence $r_n$ such that $f^{r_n} 0$ tends to infinity, and additionally $d(f^l 0, 0) < d(f^{r_n} 0, 0)$ for $l < r_n$: this implies the convergence at infinity of $f^{r_n} 0$, by an argument based on a Gromov product computation -- and the limit is a fixed point. Moreover, it has nonpositive additive strength, essentially by construction.\ lemma high_scores: fixes u::"nat \ real" and i::nat and C::real assumes "\(bdd_above (range u))" shows "\n. (\l \ n. u l \ u n) \ u n \ C \ n \ i" proof - define M where "M = max C (Max {u l|l. l < i})" define n where "n = Inf {m. u m > M}" have "\(range u \ {..M})" using assms by (meson bdd_above_Iic bdd_above_mono) then have "{m. u m > M} \ {}" using assms by (simp add: image_subset_iff not_less) then have "n \ {m. u m > M}" unfolding n_def using Inf_nat_def1 by metis then have "u n > M" by simp have "n \ i" proof (rule ccontr) assume "\ i \ n" then have *: "n < i" by simp have "u n \ Max {u l|l. l < i}" apply (rule Max_ge) using * by auto then show False using \u n > M\ M_def by auto qed moreover have "u l \ u n" if "l \ n" for l proof (cases "l = n") case True then show ?thesis by simp next case False then have "l < n" using \l \ n\ by auto then have "l \ {m. u m > M}" unfolding n_def by (meson bdd_below_def cInf_lower not_le zero_le) then show ?thesis using \u n > M\ by auto qed ultimately show ?thesis using \u n > M\ M_def less_eq_real_def by auto qed lemma isometry_not_elliptic_has_attracting_fixed_point: assumes "isometry f" "\(elliptic_isometry f)" shows "\xi \ Gromov_boundary. Gromov_extension f xi = xi \ additive_strength f xi \ 0" proof - define u where "u = (\n. dist basepoint ((f^^n) basepoint))" have NB: "\(bdd_above (range u))" proof assume "bdd_above (range u)" then obtain C where *: "\n. u n \ C" unfolding bdd_above_def by auto have "bounded {(f^^n) basepoint|n. True}" unfolding bounded_def apply (rule exI[of _ basepoint], rule exI[of _ C]) using * unfolding u_def by auto then show False using elliptic_isometryI assms by auto qed have "\r. \n. ((\l \ r n. u l \ u (r n)) \ u (r n) \ 2 * n) \ r (Suc n) \ r n + 1" apply (rule dependent_nat_choice) using high_scores[OF NB] by (auto) blast then obtain r::"nat \ nat" where r: "\n l. l \ r n \ u l \ u (r n)" "\n. u (r n) \ 2 * n" "\n. r (Suc n) \ r n + 1" by auto then have "strict_mono r" by (metis Suc_eq_plus1 Suc_le_lessD strict_monoI_Suc) then have "r n \ n" for n by (simp add: seq_suble) have A: "dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ dist basepoint ((f^^(r n)) basepoint)" if "n \ p" for n p proof - have "r n \ r p" using \n \ p\ \strict_mono r\ by (simp add: strict_mono_less_eq) then have *: "f^^((r n)) = (f^^(r p)) o (f^^(r n - r p))" unfolding funpow_add[symmetric] by auto have "dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) = dist ((f^^(r p)) basepoint) ((f^^(r p)) ((f^^(r n - r p)) basepoint))" unfolding * comp_def by auto also have "... = dist basepoint ((f^^(r n - r p)) basepoint)" using isometry_iterates[OF assms(1), of "r p"] isometryD by auto also have "... \ dist basepoint ((f^^(r n)) basepoint)" using r(1)[of "r n - r p" n] unfolding u_def by auto finally show ?thesis by simp qed have *: "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ p" if "n \ p" for n p proof - have "2 * Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) = dist basepoint ((f^^(r p)) basepoint) + dist basepoint ((f^^(r n)) basepoint) - dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint)" unfolding Gromov_product_at_def by auto also have "... \ dist basepoint ((f^^(r p)) basepoint)" using A[OF that] by auto finally show "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ p" using r(2)[of p] unfolding u_def by auto qed have *: "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ N" if "n \ N" "p \ N" for n p N using *[of n p] *[of p n] that by (auto simp add: Gromov_product_commute intro: le_cases[of n p]) have "Gromov_converging_at_boundary (\n. (f^^(r n)) basepoint)" apply (rule Gromov_converging_at_boundaryI[of basepoint]) using * by (meson dual_order.trans real_arch_simple) with Gromov_converging_at_boundary_converges[OF this] obtain xi where xi: "(\n. to_Gromov_completion ((f^^(r n)) basepoint)) \ xi" "xi \ Gromov_boundary" by auto then have *: "(\n. Gromov_extension f (to_Gromov_completion ((f^^(r n)) basepoint))) \ xi" apply (simp, rule Gromov_converging_at_boundary_bounded_perturbation[of _ _ _ "dist basepoint (f basepoint)"]) by (simp add: assms(1) funpow_swap1 isometryD(2) isometry_iterates) moreover have "(\n. Gromov_extension f (to_Gromov_completion ((f^^(r n)) basepoint))) \ Gromov_extension f xi" using continuous_on_tendsto_compose[OF Gromov_extension_isometry(2)[OF assms(1)] xi(1)] by auto ultimately have fxi: "Gromov_extension f xi = xi" using LIMSEQ_unique by auto have "Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ 0" if "n \ p" for n p unfolding Busemann_function_inner using A[OF that] by auto then have A: "eventually (\n. Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ 0) sequentially" for p unfolding eventually_sequentially by auto have B: "eventually (\n. Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1) sequentially" for p by (rule eventually_mono[OF Busemann_function_inside_approx[OF _ xi(1), of 1 "((f^^(r p)) basepoint)" basepoint, simplified]], simp) have "eventually (\n. Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1 \ 0) sequentially" for p by (rule eventually_mono[OF eventually_conj[OF A[of p] B[of p]]], simp) then have *: "Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1 \ 0" for p by auto then have A: "Busemann_function_at xi ((f^^(r p)) basepoint) basepoint / (r p) - (2 * deltaG(TYPE('a)) + 1) * (1/r p) \ 0" if "p \ 1" for p using order_trans[OF that \p \ r p\] by (auto simp add: algebra_simps divide_simps) have B: "(\p. Busemann_function_at xi ((f^^(p)) basepoint) basepoint / p - (2 * deltaG(TYPE('a)) + 1) * (1/p)) \ additive_strength f xi - (2 * deltaG(TYPE('a)) + 1) * 0" by (intro tendsto_intros assms fxi) have "additive_strength f xi - (2 * deltaG TYPE('a) + 1) * 0 \ 0" apply (rule LIMSEQ_le_const2[OF LIMSEQ_subseq_LIMSEQ[OF B \strict_mono r\]]) using A unfolding comp_def by auto then show ?thesis using xi fxi by auto qed text \Applying the previous result to the inverse map, we deduce that there is also a fixed point with nonnegative strength.\ lemma isometry_not_elliptic_has_repelling_fixed_point: assumes "isometry f" "\(elliptic_isometry f)" shows "\xi \ Gromov_boundary. Gromov_extension f xi = xi \ additive_strength f xi \ 0" proof - have *: "\elliptic_isometry (inv f)" using elliptic_isometry_inv_iff isometry_inverse(2)[OF assms(1)] assms(2) by auto obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension (inv f) xi = xi" "additive_strength (inv f) xi \ 0" using isometry_not_elliptic_has_attracting_fixed_point[OF isometry_inverse(1)[OF assms(1)] *] by auto have *: "Gromov_extension f xi = xi" using Gromov_extension_inv_fixed_point[OF isometry_inverse(1)[OF assms(1)] xi(2)] inv_inv_eq[OF isometry_inverse(2)[OF assms(1)]] by auto moreover have "additive_strength f xi \ 0" using additive_strength_inv[OF assms(1) *] xi(3) by auto ultimately show ?thesis using xi(1) by auto qed subsubsection \Parabolic isometries\ text \We show that a parabolic isometry has (at least) one neutral fixed point at infinity.\ lemma parabolic_fixed_point: assumes "parabolic_isometry f" shows "neutral_fixed_point f \ Gromov_boundary" "Gromov_extension f (neutral_fixed_point f) = neutral_fixed_point f" "additive_strength f (neutral_fixed_point f) = 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" using isometry_not_elliptic_has_attracting_fixed_point parabolic_isometryD[OF assms] by blast moreover have "additive_strength f xi = 0" using stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] xi(2)] parabolic_isometryD(3)[OF assms] by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi = 0" by auto show "neutral_fixed_point f \ Gromov_boundary" "Gromov_extension f (neutral_fixed_point f) = neutral_fixed_point f" "additive_strength f (neutral_fixed_point f) = 0" unfolding neutral_fixed_point_def using someI_ex[OF A] by auto qed text \Parabolic isometries have exactly one fixed point, the neutral fixed point at infinity. The proof goes as follows: if it has another fixed point, then the orbit of a basepoint would stay on the horospheres centered at both fixed points. But the intersection of two horospheres based at different points is a a bounded set. Hence, the map has a bounded orbit, and is therefore elliptic.\ theorem parabolic_unique_fixed_point: assumes "parabolic_isometry f" shows "Gromov_extension f xi = xi \ xi = neutral_fixed_point f" proof (auto simp add: parabolic_fixed_point[OF assms]) assume H: "Gromov_extension f xi = xi" have *: "additive_strength f xi = 0" using stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] H] parabolic_isometryD(3)[OF assms] by auto show "xi = neutral_fixed_point f" proof (rule ccontr) assume N: "xi \ neutral_fixed_point f" define C where "C = 2 * real_of_ereal (extended_Gromov_product_at basepoint xi (neutral_fixed_point f)) + 4 * deltaG(TYPE('a))" have A: "dist basepoint ((f^^n) basepoint) \ C" for n proof - have "dist ((f^^n) basepoint) basepoint - 2 * real_of_ereal (extended_Gromov_product_at basepoint xi (neutral_fixed_point f)) - 2 * deltaG(TYPE('a)) \ max (Busemann_function_at xi ((f^^n) basepoint) basepoint) (Busemann_function_at (neutral_fixed_point f) ((f^^n) basepoint) basepoint)" using dist_le_max_Busemann_functions[OF N] by (simp add: algebra_simps) also have "... \ max (n * additive_strength f xi + 2 * deltaG(TYPE('a))) (n * additive_strength f (neutral_fixed_point f) + 2 * deltaG(TYPE('a)))" apply (intro mono_intros) using Busemann_function_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] H, of n basepoint] Busemann_function_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] parabolic_fixed_point(2)[OF assms], of n basepoint] by auto also have "... = 2 * deltaG(TYPE('a))" unfolding * parabolic_fixed_point[OF assms] by auto finally show ?thesis unfolding C_def by (simp add: algebra_simps dist_commute) qed have "elliptic_isometry f" apply (rule elliptic_isometryI[of _ basepoint]) using parabolic_isometryD(1)[OF assms] A unfolding bounded_def by auto then show False using elliptic_imp_not_parabolic_loxodromic assms by auto qed qed text \When one iterates a parabolic isometry, the distance to the starting point can grow at most logarithmically.\ lemma parabolic_logarithmic_growth: assumes "parabolic_isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "n \ 1" shows "dist x ((f^^n) x) \ dist x (f x) + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" using dist_le_additive_strength[OF parabolic_isometryD(1)[OF assms(1)] parabolic_fixed_point(2)[OF assms(1)] _ assms(2)] parabolic_isometryD(3)[OF assms(1)] stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms(1)] parabolic_fixed_point(2)[OF assms(1)]] by auto text \It follows that there is no parabolic isometry in trees, since the formula in the previous lemma shows that there is no orbit growth as $\delta = 0$, and therefore orbits are bounded, contradicting the parabolicity of the isometry.\ lemma tree_no_parabolic_isometry: assumes "isometry (f::'a::Gromov_hyperbolic_space_0 \ 'a)" shows "elliptic_isometry f \ loxodromic_isometry f" proof - have "\parabolic_isometry f" proof assume P: "parabolic_isometry f" have "dist basepoint ((f^^n) basepoint) \ dist basepoint (f basepoint)" if "n \ 1" for n using parabolic_logarithmic_growth[OF P that, of basepoint] by auto then have *: "dist basepoint ((f^^n) basepoint) \ dist basepoint (f basepoint)" for n by (cases "n \ 1", auto simp add: not_less_eq_eq) have "elliptic_isometry f" apply (rule elliptic_isometryI[OF _ assms, of basepoint]) using * unfolding bounded_def by auto then show False using elliptic_imp_not_parabolic_loxodromic P by auto qed then show ?thesis using elliptic_or_parabolic_or_loxodromic[OF assms] by auto qed subsubsection \Loxodromic isometries\ text \A loxodromic isometry has (at least) two fixed points at infinity, one attracting and one repelling. We have already constructed fixed points with nonnegative and nonpositive strengths. Since the strength is nonzero (its absolute value is the stable translation length), then these fixed points correspond to what we want.\ lemma loxodromic_attracting_fixed_point: assumes "loxodromic_isometry f" shows "attracting_fixed_point f \ Gromov_boundary" "Gromov_extension f (attracting_fixed_point f) = attracting_fixed_point f" "additive_strength f (attracting_fixed_point f) < 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" using isometry_not_elliptic_has_attracting_fixed_point loxodromic_isometryD[OF assms] by blast moreover have "additive_strength f xi < 0" using stable_translation_length_eq_additive_strength[OF loxodromic_isometryD(1)[OF assms] xi(2)] loxodromic_isometryD(3)[OF assms] xi(3) by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi < 0" by auto show "attracting_fixed_point f \ Gromov_boundary" "Gromov_extension f (attracting_fixed_point f) = attracting_fixed_point f" "additive_strength f (attracting_fixed_point f) < 0" unfolding attracting_fixed_point_def using someI_ex[OF A] by auto qed lemma loxodromic_repelling_fixed_point: assumes "loxodromic_isometry f" shows "repelling_fixed_point f \ Gromov_boundary" "Gromov_extension f (repelling_fixed_point f) = repelling_fixed_point f" "additive_strength f (repelling_fixed_point f) > 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" using isometry_not_elliptic_has_repelling_fixed_point loxodromic_isometryD[OF assms] by blast moreover have "additive_strength f xi > 0" using stable_translation_length_eq_additive_strength[OF loxodromic_isometryD(1)[OF assms] xi(2)] loxodromic_isometryD(3)[OF assms] xi(3) by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi > 0" by auto show "repelling_fixed_point f \ Gromov_boundary" "Gromov_extension f (repelling_fixed_point f) = repelling_fixed_point f" "additive_strength f (repelling_fixed_point f) > 0" unfolding repelling_fixed_point_def using someI_ex[OF A] by auto qed text \The attracting and repelling fixed points of a loxodromic isometry are distinct -- precisely since one is attracting and the other is repelling.\ lemma attracting_fixed_point_neq_repelling_fixed_point: assumes "loxodromic_isometry f" shows "attracting_fixed_point f \ repelling_fixed_point f" using loxodromic_repelling_fixed_point[OF assms] loxodromic_attracting_fixed_point[OF assms] by auto text \The attracting fixed point of a loxodromic isometry is indeed attracting. Moreover, the convergence is uniform away from the repelling fixed point. This is expressed in the following proposition, where neighborhoods of the repelling and attracting fixed points are given by the property that the Gromov product with the fixed point is large. The proof goes as follows. First, the Busemann function with respect to the fixed points at infinity evolves like the strength. Therefore, $f^n e$ tends to the repulsive fixed point in negative time, and to the attracting one in positive time. Consider now a general point $x$ with $(\xi^-, x)_e \leq K$. This means that the geodesics from $e$ to $x$ and $\xi^-$ diverge before time $K$. For large $n$, since $f^{-n} e$ is close to $\xi^-$, we also get the inequality $(f^{-n} e, x)_e \leq K$. Applying $f^n$ and using the invariance of the Gromov product under isometries yields $(e, f^n x)_{f^n e} \leq K$. But this Gromov product is equal to $d(e, f^n e) - (f^n e, f^n x)_e$ (this is a general property of Gromov products). In particular, $(f^n e, f^n x) \geq d(e, f^n e) - K$, and moreover $d(e, f^n e)$ is large. Since $f^n e$ is close to $\xi^+$, it follows that $f^n x$ is also close to $\xi^+$, as desired. The real proof requires some more care as everything should be done in ereal, and moreover every inequality is only true up to some multiple of $\delta$. But everything works in the way just described above. \ proposition loxodromic_attracting_fixed_point_attracts_uniformly: assumes "loxodromic_isometry f" shows "\N. \n \ N. \x. extended_Gromov_product_at basepoint x (repelling_fixed_point f) \ ereal K \ extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) x) (attracting_fixed_point f) \ ereal M" proof - have I: "isometry f" using loxodromic_isometryD(1)[OF assms] by simp have J: "\ereal r\ \ \" for r by auto text \We show that $f^n e$ tends to the repelling fixed point in negative time.\ have "(\n. ereal (Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint)) \ - \" proof (rule tendsto_sandwich[of "\n. -\" _ _ "\n. ereal(- real n * additive_strength f (repelling_fixed_point f) + 2 * deltaG(TYPE('a)))", OF _ always_eventually], auto) fix n have "Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint \ real n * additive_strength (inv f) (repelling_fixed_point f) + 2 * deltaG(TYPE('a))" using Busemann_function_eq_additive_strength[OF isometry_inverse(1)[OF I] Gromov_extension_inv_fixed_point[OF I loxodromic_repelling_fixed_point(2)[OF assms]], of n basepoint] by auto then show "Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint \ 2 * deltaG(TYPE('a)) - real n * additive_strength f (repelling_fixed_point f)" by (simp add: I additive_strength_inv assms loxodromic_repelling_fixed_point(2)) next have "(\n. ereal (2 * deltaG TYPE('a)) + ereal (- real n) * additive_strength f (repelling_fixed_point f)) \ ereal (2 * deltaG (TYPE('a))) + (- \) * additive_strength f (repelling_fixed_point f)" apply (intro tendsto_intros) using loxodromic_repelling_fixed_point(3)[OF assms] by auto then show "(\n. ereal (2 * deltaG TYPE('a) - real n * additive_strength f (repelling_fixed_point f))) \ - \" using loxodromic_repelling_fixed_point(3)[OF assms] by auto qed then have "(\n. to_Gromov_completion (((inv f)^^n) basepoint)) \ repelling_fixed_point f" by (rule Busemann_function_minus_infinity_imp_convergent[of _ _ basepoint]) then have "(\n. extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ \" unfolding Gromov_completion_boundary_limit[OF loxodromic_repelling_fixed_point(1)[OF assms]] by simp then obtain Nr where Nr: "\n. n \ Nr \ extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f) \ ereal (K + deltaG(TYPE('a)) + 1)" unfolding Lim_PInfty by auto text \We show that $f^n e$ tends to the attracting fixed point in positive time.\ have "(\n. ereal (Busemann_function_at (attracting_fixed_point f) ((f ^^ n) basepoint) basepoint)) \ - \" proof (rule tendsto_sandwich[of "\n. -\" _ _ "\n. ereal(real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG(TYPE('a)))", OF _ always_eventually], auto) fix n show "Busemann_function_at (attracting_fixed_point f) ((f ^^ n) basepoint) basepoint \ real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG(TYPE('a))" using Busemann_function_eq_additive_strength[OF I loxodromic_attracting_fixed_point(2)[OF assms], of n basepoint] by auto next have "(\n. ereal (2 * deltaG TYPE('a)) + ereal (real n) * additive_strength f (attracting_fixed_point f)) \ ereal (2 * deltaG (TYPE('a))) + (\) * additive_strength f (attracting_fixed_point f)" apply (intro tendsto_intros) using loxodromic_attracting_fixed_point(3)[OF assms] by auto then show "(\n. ereal (real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG TYPE('a))) \ - \" using loxodromic_attracting_fixed_point(3)[OF assms] by (auto simp add: algebra_simps) qed then have *: "(\n. to_Gromov_completion (((f)^^n) basepoint)) \ attracting_fixed_point f" by (rule Busemann_function_minus_infinity_imp_convergent[of _ _ basepoint]) then have "(\n. extended_Gromov_product_at basepoint (to_Gromov_completion (((f)^^n) basepoint)) (attracting_fixed_point f)) \ \" unfolding Gromov_completion_boundary_limit[OF loxodromic_attracting_fixed_point(1)[OF assms]] by simp then obtain Na where Na: "\n. n \ Na \ extended_Gromov_product_at basepoint (to_Gromov_completion (((f)^^n) basepoint)) (attracting_fixed_point f) \ ereal (M + deltaG(TYPE('a)))" unfolding Lim_PInfty by auto text \We show that the distance between $e$ and $f^n e$ tends to infinity.\ have "(\n. extended_Gromov_distance (to_Gromov_completion basepoint) (to_Gromov_completion ((f^^n) basepoint))) \ extended_Gromov_distance (to_Gromov_completion basepoint) (attracting_fixed_point f)" using * extended_Gromov_distance_continuous[of "to_Gromov_completion basepoint"] by (metis UNIV_I continuous_on filterlim_compose tendsto_at_iff_tendsto_nhds) then have "(\n. extended_Gromov_distance (to_Gromov_completion basepoint) (to_Gromov_completion ((f^^n) basepoint))) \ \" using loxodromic_attracting_fixed_point(1)[OF assms] by simp then obtain Nd where Nd: "\n. n \ Nd \ dist basepoint ((f^^n) basepoint) \ M + K + 3 * deltaG(TYPE('a))" unfolding Lim_PInfty by auto text \Now, if $n$ is large enough so that all the convergences to infinity above are almost realized, we show that a point $x$ which is not too close to $\xi^-$ is sent by $f^n$ to a point close to $\xi^+$, uniformly.\ define N where "N = Nr + Na + Nd" have "extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) x) (attracting_fixed_point f) \ ereal M" if H: "extended_Gromov_product_at basepoint x (repelling_fixed_point f) \ K" "n \ N" for n x proof - have n: "n \ Nr" "n \ Na" "n \ Nd" using that unfolding N_def by auto have "min (extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ extended_Gromov_product_at basepoint x (repelling_fixed_point f) + deltaG(TYPE('a))" by (intro mono_intros) also have "... \ ereal K + deltaG(TYPE('a))" apply (intro mono_intros) using H by auto finally have "min (extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ ereal K + deltaG(TYPE('a))" by simp moreover have "extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f) > ereal K + deltaG(TYPE('a))" using Nr[OF n(1)] ereal_le_less by auto ultimately have A: "extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint)) \ ereal K + deltaG(TYPE('a))" using not_le by fastforce have *: "((inv f)^^n) ((f^^n) z) = z" for z by (metis I bij_is_inj inj_fn inv_f_f inv_fn isometry_inverse(2)) have **: "x = Gromov_extension ((inv f)^^n) (Gromov_extension (f^^n) x)" using Gromov_extension_isometry_composition[OF isometry_iterates[OF I] isometry_iterates[OF isometry_inverse(1)[OF I]], of n n] unfolding comp_def * apply auto by meson have "extended_Gromov_product_at (((inv f)^^n) ((f^^n) basepoint)) (Gromov_extension ((inv f)^^n) (Gromov_extension (f^^n) x)) (Gromov_extension (((inv f)^^n)) (to_Gromov_completion basepoint)) \ ereal K + deltaG(TYPE('a))" using A by (simp add: * **[symmetric]) then have B: "extended_Gromov_product_at ((f^^n) basepoint) (Gromov_extension (f^^n) x) (to_Gromov_completion basepoint) \ ereal K + deltaG(TYPE('a))" unfolding Gromov_extension_preserves_extended_Gromov_product[OF isometry_iterates[OF isometry_inverse(1)[OF I]]] by simp have "dist basepoint ((f^^n) basepoint) - deltaG(TYPE('a)) \ extended_Gromov_product_at ((f^^n) basepoint) (Gromov_extension (f^^n) x) (to_Gromov_completion basepoint) + extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))" using extended_Gromov_product_add_ge[of basepoint "(f^^n) basepoint" "Gromov_extension (f^^n) x"] by (auto simp add: algebra_simps) also have "... \ (ereal K + deltaG(TYPE('a))) + extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))" by (intro mono_intros B) finally have "extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint)) \ dist basepoint ((f^^n) basepoint) - (2 * deltaG(TYPE('a)) + K)" - unfolding ereal_minus(1)[symmetric] apply (simp only: ereal_minus_le[OF J]) apply (auto simp add: algebra_simps) - by (metis (no_types, hide_lams) linordered_field_class.sign_simps(1) linordered_field_class.sign_simps(3) mult_2_right plus_ereal.simps(1)) + apply (simp only: ereal_minus_le [OF J] ereal_minus(1) [symmetric]) + apply (auto simp add: algebra_simps) + apply (metis (no_types, hide_lams) add.assoc add.left_commute mult_2_right plus_ereal.simps(1)) + done moreover have "dist basepoint ((f ^^ n) basepoint) - (2 * deltaG TYPE('a) + K) \ M + deltaG(TYPE('a))" using Nd[OF n(3)] by auto ultimately have "extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint)) \ ereal (M + deltaG(TYPE('a)))" using order_trans ereal_le_le by auto then have "ereal (M + deltaG(TYPE('a))) \ min (extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion ((f^^n) basepoint)) (attracting_fixed_point f))" using Na[OF n(2)] by (simp add: extended_Gromov_product_at_commute) also have "... \ extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (attracting_fixed_point f) + deltaG(TYPE('a))" by (intro mono_intros) finally have "ereal M \ extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (attracting_fixed_point f)" unfolding plus_ereal.simps(1)[symmetric] ereal_add_le_add_iff2 by auto then show ?thesis by (simp add: Gromov_extension_isometry_iterates I) qed then show ?thesis by auto qed text \We deduce pointwise convergence from the previous result.\ lemma loxodromic_attracting_fixed_point_attracts: assumes "loxodromic_isometry f" "xi \ repelling_fixed_point f" shows "(\n. ((Gromov_extension f)^^n) xi) \ attracting_fixed_point f" proof - have "(\n. extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) xi) (attracting_fixed_point f)) \ \" unfolding Lim_PInfty using loxodromic_attracting_fixed_point_attracts_uniformly[OF assms(1)] by (metis Gromov_boundary_extended_product_PInf assms(2) ereal_top funpow_code_def infinity_ereal_def linear) then show ?thesis unfolding Gromov_completion_boundary_limit[OF loxodromic_attracting_fixed_point(1)[OF assms(1)]] by simp qed text \Finally, we show that a loxodromic isometry has exactly two fixed points, its attracting and repelling fixed points defined above. Indeed, we already know that these points are fixed. It remains to see that there is no other fixed point. But a fixed point which is not the repelling one is both stationary and attracted to the attracting fixed point by the previous lemma, hence it has to coincide with the attracting fixed point.\ theorem loxodromic_unique_fixed_points: assumes "loxodromic_isometry f" shows "Gromov_extension f xi = xi \ xi = attracting_fixed_point f \ xi = repelling_fixed_point f" proof - have "xi = attracting_fixed_point f" if H: "Gromov_extension f xi = xi" "xi \ repelling_fixed_point f" for xi proof - have "((Gromov_extension f)^^n) xi = xi" for n apply (induction n) using H(1) by auto then have "(\n. ((Gromov_extension f)^^n) xi) \ xi" by auto then show ?thesis using loxodromic_attracting_fixed_point_attracts[OF assms H(2)] LIMSEQ_unique by auto qed then show ?thesis using loxodromic_attracting_fixed_point(2)[OF assms] loxodromic_repelling_fixed_point(2)[OF assms] by auto qed end (*of theory Isometries_Classification*) diff --git a/thys/Gromov_Hyperbolicity/Metric_Completion.thy b/thys/Gromov_Hyperbolicity/Metric_Completion.thy --- a/thys/Gromov_Hyperbolicity/Metric_Completion.thy +++ b/thys/Gromov_Hyperbolicity/Metric_Completion.thy @@ -1,482 +1,482 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \The metric completion of a metric space\ theory Metric_Completion imports Isometries begin text \Any metric space can be completed, by adding the missing limits of Cauchy sequences. Formally, there exists an isometric embedding of the space in a complete space, with dense image. In this paragraph, we construct this metric completion. This is exactly the same construction as the way in which real numbers are constructed from rational numbers.\ subsection \Definition of the metric completion\ quotient_type (overloaded) 'a metric_completion = "nat \ ('a::metric_space)" / partial: "\u v. (Cauchy u) \ (Cauchy v) \ (\n. dist (u n) (v n)) \ 0" unfolding part_equivp_def proof(auto intro!: ext) show "\x. Cauchy x" by (rule exI[of _ "\_. undefined"]) (simp add: convergent_Cauchy convergent_const) fix x y z::"nat \ 'a" assume H: "(\n. dist (x n) (y n)) \ 0" "(\n. dist (x n) (z n)) \ 0" have *: "(\n. dist (x n) (y n) + dist (x n) (z n)) \ 0 + 0" by (rule tendsto_add) (auto simp add: H) show "(\n. dist (y n) (z n)) \ 0" apply (rule tendsto_sandwich[of "\_. 0" _ _ "\n. dist (x n) (y n) + dist (x n) (z n)"]) using * by (auto simp add: dist_triangle3) next fix x y z::"nat \ 'a" assume H: "(\n. dist (x n) (y n)) \ 0" "(\n. dist (y n) (z n)) \ 0" have *: "(\n. dist (x n) (y n) + dist (y n) (z n)) \ 0 + 0" by (rule tendsto_add) (auto simp add: H) show "(\n. dist (x n) (z n)) \ 0" apply (rule tendsto_sandwich[of "\_. 0" _ _ "\n. dist (x n) (y n) + dist (y n) (z n)"]) using * by (auto simp add: dist_triangle) next fix x y::"nat \ 'a" assume H: "Cauchy x" "(\v. Cauchy v \ (\n. dist (x n) (v n)) \ 0) = (\v. Cauchy v \ (\n. dist (y n) (v n)) \ 0)" have "Cauchy x \ (\n. dist (x n) (x n)) \ 0" using H by auto then have "(\n. dist (y n) (x n))\ 0" using H by meson moreover have "dist (x n) (y n) = dist (y n) (x n)" for n using dist_commute by auto ultimately show "(\n. dist (x n) (y n)) \ 0" by auto qed text \We have to show that the metric completion is indeed a metric space, that the original space embeds isometrically into it, and that it is complete. Before we prove these statements, we start with two simple lemmas that will be needed later on.\ lemma convergent_Cauchy_dist: fixes u v::"nat \ ('a::metric_space)" assumes "Cauchy u" "Cauchy v" shows "convergent (\n. dist (u n) (v n))" proof (rule real_Cauchy_convergent, intro CauchyI) fix e::real assume "e > 0" obtain Nu where Nu: "\n \ Nu. \m \ Nu. dist (u n) (u m) < e/2" using assms(1) by (metis \0 < e\ less_divide_eq_numeral1(1) metric_CauchyD mult_zero_left) obtain Nv where Nv: "\n \ Nv. \m \ Nv. dist (v n) (v m) < e/2" using assms(2) by (metis \0 < e\ less_divide_eq_numeral1(1) metric_CauchyD mult_zero_left) define M where "M = max Nu Nv" { fix n m assume H: "n \ M" "m \ M" have *: "dist (u n) (u m) < e/2" "dist (v n) (v m) < e/2" using Nu Nv H unfolding M_def by auto have "dist (u m) (v m) - dist (u n) (v n) \ dist (u m) (u n) + dist (v n) (v m)" - by (metis add.commute add.left_commute add_left_mono dist_commute dist_triangle2 dist_triangle_le linordered_field_class.sign_simps(42)) + by (simp add: algebra_simps) (metis add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) also have "... < e/2 + e/2" using * by (simp add: dist_commute) finally have A: "dist (u m) (v m) - dist (u n) (v n) < e" by simp have "dist (u n) (v n) - dist (u m) (v m) \ dist (u m) (u n) + dist (v n) (v m)" - by (metis add.commute add.left_commute add_left_mono dist_commute dist_triangle2 dist_triangle_le linordered_field_class.sign_simps(42)) + by (simp add: algebra_simps) (metis add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) also have "... < e/2 + e/2" using * by (simp add: dist_commute) finally have "dist (u n) (v n) - dist (u m) (v m) < e" by simp then have "norm(dist (u m) (v m) - dist (u n) (v n)) < e" using A by auto } then show "\M. \m \ M. \n \ M. norm (dist (u m) (v m) - dist (u n) (v n)) < e" by auto qed lemma convergent_add_null: fixes u v::"nat \ ('a::real_normed_vector)" assumes "convergent u" "(\n. v n - u n) \ 0" shows "convergent v" "lim v = lim u" proof - have "(\n. u n + (v n - u n)) \ lim u + 0" apply (rule tendsto_add) using assms convergent_LIMSEQ_iff by auto then have *: "v \ lim u" by auto show "convergent v" using * by (simp add: Lim_def convergentI) show "lim v = lim u" using * by (simp add: limI) qed text \Let us now prove that the metric completion is a metric space: the distance between two Cauchy sequences is the limit of the distances of points in the sequence. The convergence follows from Lemma~\verb+convergent_Cauchy_dist+ above.\ instantiation metric_completion :: (metric_space) metric_space begin lift_definition dist_metric_completion::"('a::metric_space) metric_completion \ 'a metric_completion \ real" is "\x y. lim (\n. dist (x n) (y n))" proof - fix x y z t::"nat \ 'a" assume H: "Cauchy x \ Cauchy y \ (\n. dist (x n) (y n)) \ 0" "Cauchy z \ Cauchy t \ (\n. dist (z n) (t n)) \ 0" show "lim (\n. dist (x n) (z n)) = lim (\n. dist (y n) (t n))" proof (rule convergent_add_null(2)) show "convergent (\n. dist (y n) (t n))" apply (rule convergent_Cauchy_dist) using H by auto have a: "(\n. - dist (t n) (z n) - dist (x n) (y n)) \ -0 -0" apply (intro tendsto_intros) using H by (auto simp add: dist_commute) have b:"(\n. dist (t n) (z n) + dist (x n) (y n)) \ 0 + 0" apply (rule tendsto_add) using H by (auto simp add: dist_commute) have I: "dist (x n) (z n) \ dist (t n) (y n) + (dist (t n) (z n) + dist (x n) (y n))" for n using dist_triangle[of "x n" "z n" "y n"] dist_triangle[of "y n" "z n" "t n"] by (auto simp add: dist_commute add.commute) show "(\n. dist (x n) (z n) - dist (y n) (t n)) \ 0" apply (rule tendsto_sandwich[of "\n. -(dist (x n) (y n) + dist (z n) (t n))" _ _ "\n. dist (x n) (y n) + dist (z n) (t n)"]) apply (auto intro!: always_eventually simp add: algebra_simps dist_commute I) apply (meson add_left_mono dist_triangle3 dist_triangle_le) using a b by auto qed qed lemma dist_metric_completion_limit: fixes x y::"'a metric_completion" shows "(\n. dist (rep_metric_completion x n) (rep_metric_completion y n)) \ dist x y" proof - have C: "Cauchy (rep_metric_completion x)" "Cauchy (rep_metric_completion y)" using Quotient3_metric_completion Quotient3_rep_reflp by fastforce+ show ?thesis unfolding dist_metric_completion_def using C apply auto using convergent_Cauchy_dist[OF C] convergent_LIMSEQ_iff by force qed lemma dist_metric_completion_limit': fixes x y::"nat \ 'a" assumes "Cauchy x" "Cauchy y" shows "(\n. dist (x n) (y n)) \ dist (abs_metric_completion x) (abs_metric_completion y)" apply (subst dist_metric_completion.abs_eq) using assms convergent_Cauchy_dist[OF assms] by (auto simp add: convergent_LIMSEQ_iff) text \To define a metric space in the current library of Isabelle/HOL, one should also introduce a uniformity structure and a topology, as follows (they are prescribed by the distance):\ definition uniformity_metric_completion::"(('a metric_completion) \ ('a metric_completion)) filter" where "uniformity_metric_completion = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_metric_completion :: "'a metric_completion set \ bool" where "open_metric_completion U = (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix x y::"'a metric_completion" have C: "Cauchy (rep_metric_completion x)" "Cauchy (rep_metric_completion y)" using Quotient3_metric_completion Quotient3_rep_reflp by fastforce+ show "(dist x y = 0) = (x = y)" apply (subst Quotient3_rel_rep[OF Quotient3_metric_completion, symmetric]) unfolding dist_metric_completion_def using C apply auto using convergent_Cauchy_dist[OF C] convergent_LIMSEQ_iff apply force by (simp add: limI) next fix x y z::"'a metric_completion" have a: "(\n. dist (rep_metric_completion x n) (rep_metric_completion y n)) \ dist x y" using dist_metric_completion_limit by auto have b: "(\n. dist (rep_metric_completion x n) (rep_metric_completion z n) + dist (rep_metric_completion y n) (rep_metric_completion z n)) \ dist x z + dist y z" apply (rule tendsto_add) using dist_metric_completion_limit by auto show "dist x y \ dist x z + dist y z" by (rule LIMSEQ_le[OF a b], rule exI[of _ 0], auto simp add: dist_triangle2) qed (auto simp add: uniformity_metric_completion_def open_metric_completion_def) end text \Let us now show that the distance thus defined on the metric completion is indeed complete. This is essentially by design.\ instance metric_completion :: (metric_space) complete_space proof fix X::"nat \ 'a metric_completion" assume "Cauchy X" have *: "\N. \n \ N. dist (rep_metric_completion (X k) N) (rep_metric_completion (X k) n) < (1/Suc k)" for k proof - have "Cauchy (rep_metric_completion (X k))" using Quotient3_metric_completion Quotient3_rep_reflp by fastforce+ then have "\N. \m \ N. \n \ N. dist (rep_metric_completion (X k) m) (rep_metric_completion (X k) n) < (1/Suc k)" unfolding Cauchy_def by auto then show ?thesis by auto qed have "\N. \k. \n \ N k. dist (rep_metric_completion (X k) (N k)) (rep_metric_completion (X k) n) < (1/Suc k)" apply (rule choice) using * by auto then obtain N::"nat \ nat" where N: "dist (rep_metric_completion (X k) (N k)) (rep_metric_completion (X k) n) < (1/Suc k)" if "n \ N k" for n k by auto define u where "u = (\k. rep_metric_completion (X k) (N k))" have "Cauchy u" proof (rule metric_CauchyI) fix e::real assume "e > 0" obtain K::nat where "K > 4/e" using reals_Archimedean2 by blast obtain L::nat where L: "\m \ L. \n \ L. dist (X m) (X n) < e/2" using metric_CauchyD[OF \Cauchy X\, of "e/2"] \e > 0\ by auto { fix m n assume "m \ max K L" "n \ max K L" then have "dist (X m) (X n) < e/2" using L by auto then have "eventually (\p. dist (rep_metric_completion (X m) p) (rep_metric_completion (X n) p) < e/2) sequentially" using dist_metric_completion_limit[of "X m" "X n"] by (metis order_tendsto_iff) then obtain p where p: "p \ max (N m) (N n)" "dist (rep_metric_completion (X m) p) (rep_metric_completion (X n) p) < e/2" using eventually_False_sequentially eventually_elim2 eventually_ge_at_top by blast have "dist (u m) (rep_metric_completion (X m) p) < 1 / real (Suc m)" unfolding u_def using N[of m p] p(1) by auto also have "... < e/4" using \m \ max K L\ \K > 4/e\ \e > 0\ apply (auto simp add: divide_simps algebra_simps) by (metis leD le_less_trans less_add_same_cancel2 linear of_nat_le_iff real_mult_le_cancel_iff2) finally have Im: "dist (u m) (rep_metric_completion (X m) p) < e/4" by simp have "dist (u n) (rep_metric_completion (X n) p) < 1 / real (Suc n)" unfolding u_def using N[of n p] p(1) by auto also have "... < e/4" using \n \ max K L\ \K > 4/e\ \e > 0\ apply (auto simp add: divide_simps algebra_simps) by (metis leD le_less_trans less_add_same_cancel2 linear of_nat_le_iff real_mult_le_cancel_iff2) finally have In: "dist (u n) (rep_metric_completion (X n) p) < e/4" by simp have "dist (u m) (u n) \ dist (u m) (rep_metric_completion (X m) p) + dist (rep_metric_completion (X m) p) (rep_metric_completion (X n) p) + dist (rep_metric_completion (X n) p) (u n)" by (metis add.commute add_left_mono dist_commute dist_triangle_le dist_triangle) also have "... < e/4 + e/2 + e/4" using In Im p(2) by (simp add: dist_commute) also have "... = e" by auto finally have "dist (u m) (u n) < e" by auto } then show "\M. \m \ M. \n \ M. dist (u m) (u n) < e" by meson qed have *: "(\n. dist (abs_metric_completion u) (X n)) \ 0" proof (rule order_tendstoI, auto simp add: less_le_trans eventually_sequentially) fix e::real assume "e > 0" obtain K::nat where "K > 4/e" using reals_Archimedean2 by blast obtain L::nat where L: "\m \ L. \n \ L. dist (u m) (u n) < e/4" using metric_CauchyD[OF \Cauchy u\, of "e/4"] \e > 0\ by auto { fix n assume n: "n \ max K L" { fix p assume p: "p \ max (N n) L" have "dist (u n) (rep_metric_completion (X n) p) < 1/(Suc n)" unfolding u_def using N p by simp also have "... < e/4" using \n \ max K L\ \K > 4/e\ \e > 0\ apply (auto simp add: divide_simps algebra_simps) by (metis leD le_less_trans less_add_same_cancel2 linear of_nat_le_iff real_mult_le_cancel_iff2) finally have *: "dist (u n) (rep_metric_completion (X n) p) < e/4" by fastforce have "dist (u p) (rep_metric_completion (X n) p) \ dist (u p) (u n) + dist (u n) (rep_metric_completion (X n) p)" using dist_triangle by auto also have "... < e/4 + e/4" using * L n p by force finally have "dist (u p) (rep_metric_completion (X n) p) \ e/2" by auto } then have A: "eventually (\p. dist (u p) (rep_metric_completion (X n) p) \ e/2) sequentially" using eventually_at_top_linorder by blast have B: "(\p. dist (u p) (rep_metric_completion (X n) p)) \ dist (abs_metric_completion u) (X n)" using dist_metric_completion_limit'[OF \Cauchy u\, of "rep_metric_completion (X n)"] unfolding Quotient3_abs_rep[OF Quotient3_metric_completion, of "X n"] using Quotient3_rep_reflp[OF Quotient3_metric_completion] by auto have "dist (abs_metric_completion u) (X n) \ e/2" apply (rule LIMSEQ_le_const2[OF B]) using A unfolding eventually_sequentially by auto then have "dist (abs_metric_completion u) (X n) < e" using \e > 0\ by auto } then show "\N. \n \ N. dist (abs_metric_completion u) (X n) < e" by blast qed have "X \ abs_metric_completion u" apply (rule tendstoI) using * by (auto simp add: order_tendsto_iff dist_commute) then show "convergent X" unfolding convergent_def by auto qed subsection \Isometric embedding of a space in its metric completion\ text \The canonical embedding of a space into its metric completion is obtained by taking the Cauchy sequence which is constant, equal to the given point. This is indeed an isometric embedding with dense image, as we prove in the lemmas below.\ definition to_metric_completion::"('a::metric_space) \ 'a metric_completion" where "to_metric_completion x = abs_metric_completion (\n. x)" lemma to_metric_completion_isometry: "isometry_on UNIV to_metric_completion" proof (rule isometry_onI) fix x y::'a have "(\n. dist (x) (y)) \ dist (to_metric_completion x) (to_metric_completion y)" unfolding to_metric_completion_def apply (rule dist_metric_completion_limit') unfolding Cauchy_def by auto then show "dist (to_metric_completion x) (to_metric_completion y) = dist x y" by (simp add: LIMSEQ_const_iff) qed lemma to_metric_completion_dense: assumes "open U" "U \ {}" shows "\x. to_metric_completion x \ U" proof - obtain y where "y \ U" using \U \ {}\ by auto obtain e::real where e: "e > 0" "\z. dist z y < e \ z \ U" using \y \ U\ \open U\ by (metis open_dist) have *: "Cauchy (rep_metric_completion y)" using Quotient3_metric_completion Quotient3_rep_reflp by fastforce then obtain N where N: "\n \ N. \m \ N. dist (rep_metric_completion y n) (rep_metric_completion y m) < e/2" using \e > 0\ unfolding Cauchy_def by (meson divide_pos_pos zero_less_numeral) define x where "x = rep_metric_completion y N" have "(\n. dist x (rep_metric_completion y n)) \ dist (to_metric_completion x) (abs_metric_completion (rep_metric_completion y))" unfolding to_metric_completion_def apply (rule dist_metric_completion_limit') using * unfolding Cauchy_def by auto then have "(\n. dist x (rep_metric_completion y n)) \ dist (to_metric_completion x) y" unfolding Quotient3_abs_rep[OF Quotient3_metric_completion] by simp moreover have "eventually (\n. dist x (rep_metric_completion y n) \ e/2) sequentially" unfolding eventually_sequentially x_def apply (rule exI[of _ N]) using N less_imp_le by auto ultimately have "dist (to_metric_completion x) y \ e/2" using LIMSEQ_le_const2 unfolding eventually_sequentially by metis then have "to_metric_completion x \ U" using e by auto then show ?thesis by auto qed lemma to_metric_completion_dense': "closure (range to_metric_completion) = UNIV" apply (auto simp add: closure_iff_nhds_not_empty) using to_metric_completion_dense by fastforce text \The main feature of the completion is that a uniformly continuous function on the original space can be extended to a uniformly continuous function on the completion, i.e., it can be written as the composition of a new function and of the inclusion \verb+to_metric_completion+.\ lemma lift_to_metric_completion: fixes f::"('a::metric_space) \ ('b::complete_space)" assumes "uniformly_continuous_on UNIV f" shows "\g. (uniformly_continuous_on UNIV g) \ (f = g o to_metric_completion) \ (\x \ range to_metric_completion. g x = f (inv to_metric_completion x))" proof - define I::"'a metric_completion \ 'a" where "I = inv to_metric_completion" have "uniformly_continuous_on (range to_metric_completion) I" using isometry_on_uniformly_continuous[OF isometry_on_inverse(1)[OF to_metric_completion_isometry]] I_def by auto then have UC: "uniformly_continuous_on (range to_metric_completion) (\x. f (I x))" using assms uniformly_continuous_on_compose by (metis I_def bij_betw_imp_surj_on bij_betw_inv_into isometry_on_inverse(4) to_metric_completion_isometry) obtain g where g: "uniformly_continuous_on (closure(range to_metric_completion)) g" "\x. x \ range to_metric_completion \ f (I x) = g x" using uniformly_continuous_on_extension_on_closure[OF UC] by metis have "uniformly_continuous_on UNIV g" using to_metric_completion_dense' g(1) by metis moreover have "f x = g (to_metric_completion x)" for x using g(2) by (metis I_def UNIV_I isometry_on_inverse(2) range_eqI to_metric_completion_isometry) moreover have "g x = f (inv to_metric_completion x)" if "x \ range to_metric_completion" for x using I_def g(2) that by auto ultimately show ?thesis unfolding comp_def by auto qed text \When the function is an isometry, the lifted function is also an isometry (and its range is the closure of the range of the original function). This shows that the metric completion is unique, up to isometry:\ lemma lift_to_metric_completion_isometry: fixes f::"('a::metric_space) \ ('b::complete_space)" assumes "isometry_on UNIV f" shows "\g. isometry_on UNIV g \ range g = closure(range f) \ f = g o to_metric_completion \ (\x \ range to_metric_completion. g x = f (inv to_metric_completion x))" proof - have *: "uniformly_continuous_on UNIV f" using assms isometry_on_uniformly_continuous by force obtain g where g: "uniformly_continuous_on UNIV g" "f = g o to_metric_completion" "\x. x \ range to_metric_completion \ g x = f (inv to_metric_completion x)" using lift_to_metric_completion[OF *] by blast have *: "isometry_on (range to_metric_completion) g" apply (rule isometry_on_cong[OF _ g(3)], rule isometry_on_compose[of _ _ f]) using assms isometry_on_inverse[OF to_metric_completion_isometry] isometry_on_subset by (auto) (fastforce) then have "isometry_on UNIV g" unfolding to_metric_completion_dense'[symmetric] apply (rule isometry_on_closure) using continuous_on_subset[OF uniformly_continuous_imp_continuous[OF g(1)]] by auto have "g`(range to_metric_completion) \ range f" using g unfolding comp_def by auto moreover have "g`(closure (range to_metric_completion)) \ closure (g`(range to_metric_completion))" using uniformly_continuous_imp_continuous[OF g(1)] by (meson closed_closure closure_subset continuous_on_subset image_closure_subset top_greatest) ultimately have "range g \ closure (range f)" unfolding to_metric_completion_dense' by (simp add: g(2) image_comp) have "range f \ range g" using g(2) by auto moreover have "closed (range g)" using isometry_on_complete_image[OF \isometry_on UNIV g\] by (simp add: complete_eq_closed) ultimately have "closure (range f) \ range g" by (simp add: closure_minimal) then have "range g = closure (range f)" using \range g \ closure (range f)\ by auto then show ?thesis using \isometry_on UNIV g\ g by metis qed subsection \The metric completion of a second countable space is second countable\ text \We want to show that the metric completion of a second countable space is still second countable. This is most easily expressed using the fact that a metric space is second countable if and only if there exists a dense countable subset. We prove the equivalence in the next lemma, and use it then to prove that the metric completion is still second countable.\ lemma second_countable_iff_dense_countable_subset: "(\B::'a::metric_space set set. countable B \ topological_basis B) \ (\A::'a set. countable A \ closure A = UNIV)" proof assume "\B::'a set set. countable B \ topological_basis B" then obtain B::"'a set set" where "countable B" "topological_basis B" by auto define A where "A = (\U. SOME x. x \ U)`B" have "countable A" unfolding A_def using \countable B\ by auto moreover have "closure A = UNIV" proof (auto simp add: closure_approachable) fix x::'a and e::real assume "e > 0" obtain U where "U \ B" "x \ U" "U \ ball x e" by (rule topological_basisE[OF \topological_basis B\, of "ball x e" x], auto simp add: \e > 0\) define y where "y = (\U. SOME x. x \ U) U" have "y \ U" unfolding y_def using \x \ U\ some_in_eq by fastforce then have "dist y x < e" using \U \ ball x e\ by (metis dist_commute mem_ball subset_iff) moreover have "y \ A" unfolding A_def y_def using \U \ B\ by auto ultimately show "\y\A. dist y x < e" by auto qed ultimately show "\A::'a set. countable A \ closure A = UNIV" by auto next assume "\A::'a set. countable A \ closure A = UNIV" then obtain A::"'a set" where "countable A" "closure A = UNIV" by auto define B where "B = (\(x, (n::nat)). ball x (1/n))`(A \ UNIV)" have "countable B" unfolding B_def using \countable A\ by auto moreover have "topological_basis B" proof (rule topological_basisI) fix x::'a and U assume "x \ U" "open U" then obtain e where "e > 0" "ball x e \ U" using openE by blast obtain n::nat where "n > 2/e" using reals_Archimedean2 by auto then have "n > 0" using \e > 0\ not_less by fastforce then have "1/n > 0" using zero_less_divide_iff by fastforce then obtain y where y: "y \ A" "dist x y < 1/n" by (metis \closure A = UNIV\ UNIV_I closure_approachable dist_commute) then have "ball y (1/n) \ B" unfolding B_def by auto moreover have "x \ ball y (1/n)" using y(2) by (auto simp add: dist_commute) moreover have "ball y (1/n) \ U" proof (auto) fix z assume z: "dist y z < 1/n" have "dist z x \ dist z y + dist y x" using dist_triangle by auto also have "... < 1/n + 1/n" using z y(2) by (auto simp add: dist_commute) also have "... < e" using \n > 2/e\ \e > 0\ \n > 0\ by (auto simp add: divide_simps mult.commute) finally have "z \ ball x e" by (auto simp add: dist_commute) then show "z \ U" using \ball x e \ U\ by auto qed ultimately show "\V\B. x \ V \ V \ U" by metis qed (auto simp add: B_def) ultimately show "\B::'a set set. countable B \ topological_basis B" by auto qed lemma second_countable_metric_dense_subset: "\A::'a::{metric_space, second_countable_topology} set. countable A \ closure A = UNIV" using ex_countable_basis by (auto simp add: second_countable_iff_dense_countable_subset[symmetric]) instance metric_completion::("{metric_space, second_countable_topology}") second_countable_topology proof obtain A::"'a set" where "countable A" "closure A = UNIV" using second_countable_metric_dense_subset by auto define Ab where "Ab = to_metric_completion`A" have "range to_metric_completion \ closure Ab" unfolding Ab_def by (metis \closure A = UNIV\ isometry_on_continuous[OF to_metric_completion_isometry] closed_closure closure_subset image_closure_subset) then have "closure Ab = UNIV" by (metis (no_types) to_metric_completion_dense'[symmetric] \range to_metric_completion \ closure Ab\ closure_closure closure_mono top.extremum_uniqueI) moreover have "countable Ab" unfolding Ab_def using \countable A\ by auto ultimately have "\Ab::'a metric_completion set. countable Ab \ closure Ab = UNIV" by auto then show "\B::'a metric_completion set set. countable B \ open = generate_topology B" using second_countable_iff_dense_countable_subset topological_basis_imp_subbasis by auto qed instance metric_completion::("{metric_space, second_countable_topology}") polish_space by standard end (*of theory Metric_Completion*) diff --git a/thys/IEEE_Floating_Point/IEEE_Properties.thy b/thys/IEEE_Floating_Point/IEEE_Properties.thy --- a/thys/IEEE_Floating_Point/IEEE_Properties.thy +++ b/thys/IEEE_Floating_Point/IEEE_Properties.thy @@ -1,1032 +1,1030 @@ (* Author: Lei Yu, University of Cambridge Author: Fabian Hellauer Fabian Immler *) section \Proofs of Properties about Floating Point Arithmetic\ theory IEEE_Properties imports IEEE begin subsection \Theorems derived from definitions\ lemma valof_eq: "valof x = (if exponent x = 0 then (- 1) ^ sign x * (2 / 2 ^ bias TYPE(('a, 'b) float)) * (real (fraction x) / 2 ^ LENGTH('b)) else (- 1) ^ sign x * (2 ^ exponent x / 2 ^ bias TYPE(('a, 'b) float)) * (1 + real (fraction x) / 2 ^ LENGTH('b)))" for x::"('a, 'b) float" unfolding Let_def by transfer (auto simp: bias_def divide_simps unat_eq_0) lemma exponent_le[simp]: "exponent a \ unat (max_word::'a word)" for a::"('a, _)float" by transfer (auto intro!: unat_mono simp: word_le_nat_alt[symmetric]) lemma max_word_le_exponent_iff[simp]: "unat (max_word::'a word) \ exponent a \ unat (max_word::'a word) = exponent a" for a::"('a, _)float" using le_antisym by fastforce lemma infinity_simps: "sign (plus_infinity::('e, 'f)float) = 0" "sign (minus_infinity::('e, 'f)float) = 1" "exponent (plus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "exponent (minus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "fraction (plus_infinity::('e, 'f)float) = 0" "fraction (minus_infinity::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer (auto simp: emax_def) subgoal by transfer (auto simp: emax_def) subgoal by transfer auto subgoal by transfer auto done lemma zero_simps: "sign (0::('e, 'f)float) = 0" "sign (- 0::('e, 'f)float) = 1" "exponent (0::('e, 'f)float) = 0" "exponent (- 0::('e, 'f)float) = 0" "fraction (0::('e, 'f)float) = 0" "fraction (- 0::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto done lemma topfloat_simps: "sign (topfloat::('e, 'f)float) = 0" "exponent (topfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (topfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" and bottomfloat_simps: "sign (bottomfloat::('e, 'f)float) = 1" "exponent (bottomfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (bottomfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" subgoal by transfer auto subgoal by transfer (auto simp: emax_def unat_sub) subgoal apply transfer apply safe by (metis One_nat_def diff_less lessI minus_one_norm minus_one_word unat_0 unat_lt2p unat_of_nat_len) subgoal by transfer auto subgoal by transfer (auto simp: emax_def unat_sub) subgoal apply transfer apply safe by (metis One_nat_def diff_less lessI minus_one_norm minus_one_word unat_0 unat_lt2p unat_of_nat_len) done lemma emax_eq: "emax x = 2^LENGTH('e) - 1" for x::"('e, 'f)float itself" unfolding bias_def emax_def by (auto simp: max_word_eq unat_minus word_size) lemmas float_defs = is_finite_def is_infinity_def is_zero_def is_nan_def is_normal_def is_denormal_def valof_eq less_eq_float_def less_float_def flt_def fgt_def fle_def fge_def feq_def fcompare_def infinity_simps zero_simps topfloat_simps bottomfloat_simps float_eq_def lemma float_cases: "is_nan a \ is_infinity a \ is_normal a \ is_denormal a \ is_zero a" by (auto simp: emax_def float_defs not_less) lemma float_cases_finite: "is_nan a \ is_infinity a \ is_finite a" by (simp add: float_cases is_finite_def) lemma float_zero1[simp]: "is_zero 0" unfolding float_defs by transfer auto lemma float_zero2[simp]: "is_zero (- x) \ is_zero x" unfolding float_defs by transfer auto lemma emax_pos[simp]: "0 < emax x" "emax x \ 0" by (auto simp: emax_def) text \The types of floating-point numbers are mutually distinct.\ lemma float_distinct: "\ (is_nan a \ is_infinity a)" "\ (is_nan a \ is_normal a)" "\ (is_nan a \ is_denormal a)" "\ (is_nan a \ is_zero a)" "\ (is_infinity a \ is_normal a)" "\ (is_infinity a \ is_denormal a)" "\ (is_infinity a \ is_zero a)" "\ (is_normal a \ is_denormal a)" "\ (is_denormal a \ is_zero a)" by (auto simp: float_defs) lemma denormal_imp_not_zero: "is_denormal f \ \is_zero f" by (simp add: is_denormal_def is_zero_def) lemma normal_imp_not_zero: "is_normal f \ \is_zero f" by (simp add: is_normal_def is_zero_def) lemma normal_imp_not_denormal: "is_normal f \ \is_denormal f" by (simp add: is_normal_def is_denormal_def) lemma denormal_zero[simp]: "\is_denormal 0" "\is_denormal minus_zero" using denormal_imp_not_zero float_zero1 float_zero2 by blast+ lemma normal_zero[simp]: "\is_normal 0" "\is_normal minus_zero" using normal_imp_not_zero float_zero1 float_zero2 by blast+ lemma float_distinct_finite: "\ (is_nan a \ is_finite a)" "\(is_infinity a \ is_finite a)" by (auto simp: float_defs) lemma finite_infinity: "is_finite a \ \ is_infinity a" by (auto simp: float_defs) lemma finite_nan: "is_finite a \ \ is_nan a" by (auto simp: float_defs) text \For every real number, the floating-point numbers closest to it always exist.\ lemma is_closest_exists: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "\a. is_closest v s x a" using finite non_empty proof (induct s rule: finite_induct) case empty then show ?case by simp next case (insert z s) show ?case proof (cases "s = {}") case True then have "is_closest v (insert z s) x z" by (auto simp: is_closest_def) then show ?thesis by metis next case False then obtain a where a: "is_closest v s x a" using insert by metis then show ?thesis proof (cases "\v a - x\" "\v z - x\" rule: le_cases) case le then show ?thesis by (metis insert_iff a is_closest_def) next case ge have "\b. b \ s \ \v a - x\ \ \v b - x\" by (metis a is_closest_def) then have "\b. b \ insert z s \ \v z - x\ \ \v b - x\" by (metis eq_iff ge insert_iff order.trans) then show ?thesis using is_closest_def a by (metis insertI1) qed qed qed lemma closest_is_everything: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "is_closest v s x (closest v p s x) \ ((\b. is_closest v s x b \ p b) \ p (closest v p s x))" unfolding closest_def by (rule someI_ex) (metis assms is_closest_exists [of s v x]) lemma closest_in_set: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "closest v p s x \ s" by (metis assms closest_is_everything is_closest_def) lemma closest_is_closest_finite: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "is_closest v s x (closest v p s x)" by (metis closest_is_everything assms) instance float::(len, len) finite proof qed (transfer, simp) lemma is_finite_nonempty: "{a. is_finite a} \ {}" proof - have "0 \ {a. is_finite a}" unfolding float_defs by transfer auto then show ?thesis by (metis empty_iff) qed lemma closest_is_closest: fixes v :: "('e, 'f)float \ real" assumes "s \ {}" shows "is_closest v s x (closest v p s x)" by (rule closest_is_closest_finite) (auto simp: assms) subsection \Properties about ordering and bounding\ text \Lifting of non-exceptional comparisons.\ lemma float_lt [simp]: assumes "is_finite a" "is_finite b" shows "a < b \ valof a < valof b" proof assume "valof a < valof b" moreover have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) ultimately have "fcompare a b = Lt" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a < b" by (auto simp: float_defs) next assume "a < b" then have lt: "fcompare a b = Lt" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "valof a < valof b" using lt assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_eq [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a = valof b" proof assume *: "valof a = valof b" have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto with * have "fcompare a b = Eq" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a \ b" by (auto simp: float_defs) next assume "a \ b" then have eq: "fcompare a b = Eq" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then show "valof a = valof b" using eq assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_le [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a \ valof b" proof - have "a \ b \ a < b \ a \ b" by (auto simp add: float_defs) then show ?thesis by (auto simp add: assms) qed text \Reflexivity of equality for non-NaNs.\ lemma float_eq_refl [simp]: "a \ a \ \ is_nan a" by (auto simp: float_defs) text \Properties about Ordering.\ lemma float_lt_trans: "is_finite a \ is_finite b \ is_finite c \ a < b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_less_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b \ c \ a \ c" by (auto simp: le_trans) lemma float_le_neg: "is_finite a \ is_finite b \ \ a < b \ b \ a" by auto text \Properties about bounding.\ lemma float_le_infinity [simp]: "\ is_nan a \ a \ plus_infinity" unfolding float_defs by auto lemma zero_le_topfloat[simp]: "0 \ topfloat" "- 0 \ topfloat" - unfolding float_defs - by (auto simp: sign_simps divide_simps power_gt1_lemma) + by (auto simp: float_defs field_simps power_gt1_lemma) lemma LENGTH_contr: "Suc 0 < LENGTH('e) \ 2 ^ LENGTH('e::len) \ Suc (Suc 0) \ False" by (metis le_antisym len_gt_0 n_less_equal_power_2 not_less_eq numeral_2_eq_2 one_le_numeral self_le_power) lemma valof_topfloat: "valof (topfloat::('e, 'f)float) = largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that LENGTH_contr by (auto simp add: emax_eq largest_def divide_simps float_defs ) lemma float_frac_le: "fraction a \ 2^LENGTH('f) - 1" for a::"('e, 'f)float" unfolding float_defs using less_Suc_eq_le by transfer fastforce lemma float_exp_le: "is_finite a \ exponent a \ emax TYPE(('e, 'f)float) - 1" for a::"('e, 'f)float" unfolding float_defs by auto lemma float_sign_le: "(-1::real)^(sign a) = 1 \ (-1::real)^(sign a) = -1" by (metis neg_one_even_power neg_one_odd_power) lemma exp_less: "a \ b \ (2::real)^a \ 2^b" for a b :: nat by auto lemma div_less: "a \ b \ c > 0 \ a/c \ b/c" for a b c :: "'a::linordered_field" by (metis divide_le_cancel less_asym) lemma finite_topfloat: "is_finite topfloat" unfolding float_defs by auto lemmas float_leI = float_le[THEN iffD2] lemma factor_minus: "x * a - x = x * (a - 1)" for x a::"'a::comm_semiring_1_cancel" by (simp add: algebra_simps) lemma real_le_power_numeral_diff: "real a \ numeral b ^ n - 1 \ a \ numeral b ^ n - 1" by (metis (mono_tags, lifting) of_nat_1 of_nat_diff of_nat_le_iff of_nat_numeral one_le_numeral one_le_power semiring_1_class.of_nat_power) definition denormal_exponent::"('e, 'f)float itself \ int" where "denormal_exponent x = 1 - (int (LENGTH('f)) + int (bias TYPE(('e, 'f)float)))" definition normal_exponent::"('e, 'f)float \ int" where "normal_exponent x = int (exponent x) - int (bias TYPE(('e, 'f)float)) - int (LENGTH('f))" definition denormal_mantissa::"('e, 'f)float \ int" where "denormal_mantissa x = (-1::int)^sign x * int (fraction x)" definition normal_mantissa::"('e, 'f)float \ int" where "normal_mantissa x = (-1::int)^sign x * (2^LENGTH('f) + int (fraction x))" lemma unat_one_word_le: "unat a \ Suc 0" for a::"1 word" using unat_lt2p[of a] by auto lemma one_word_le: "a \ 1" for a::"1 word" by (auto simp: word_le_nat_alt unat_one_word_le) lemma sign_cases[case_names pos neg]: obtains "sign x = 0" | "sign x = 1" proof cases assume "sign x = 0" then show ?thesis .. next assume "sign x \ 0" have "sign x \ 1" by transfer (auto simp: unat_one_word_le) then have "sign x = 1" using \sign x \ 0\ by auto then show ?thesis .. qed lemma is_infinity_cases: assumes "is_infinity x" obtains "x = plus_infinity" | "x = minus_infinity" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = plus_infinity" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_0) then show ?thesis .. next assume "sign x = 1" then have "x = minus_infinity" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma is_zero_cases: assumes "is_zero x" obtains "x = 0" | "x = - 0" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = 0" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_0) then show ?thesis .. next assume "sign x = 1" then have "x = minus_zero" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma minus_minus_float[simp]: "- (-f) = f" for f::"('e, 'f)float" by transfer auto lemma sign_minus_float: "sign (-f) = (1 - sign f)" for f::"('e, 'f)float" by transfer (auto simp: unat_eq_1 one_word_le unat_sub) lemma exponent_uminus[simp]: "exponent (- f) = exponent f" by transfer auto lemma fraction_uminus[simp]: "fraction (- f) = fraction f" by transfer auto lemma is_normal_minus_float[simp]: "is_normal (-f) = is_normal f" for f::"('e, 'f)float" by (auto simp: is_normal_def) lemma is_denormal_minus_float[simp]: "is_denormal (-f) = is_denormal f" for f::"('e, 'f)float" by (auto simp: is_denormal_def) lemma bitlen_normal_mantissa: "bitlen (abs (normal_mantissa x)) = Suc LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) - (auto simp: bitlen_le_iff_power bitlen_ge_iff_power nat_add_distrib abs_mult sign_simps - add_nonneg_pos normal_mantissa_def intro!: antisym) + (auto simp: bitlen_le_iff_power bitlen_ge_iff_power nat_add_distrib + normal_mantissa_def intro!: antisym) qed lemma less_int_natI: "x < y" if "0 \ x" "nat x < nat y" using that by arith lemma normal_exponent_bounds_int: "2 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" using that unfolding normal_exponent_def is_normal_def emax_eq bias_def by (auto simp del: zless_nat_conj intro!: less_int_natI simp: nat_add_distrib nat_mult_distrib nat_power_eq power_Suc[symmetric]) lemmas of_int_leI = of_int_le_iff[THEN iffD2] lemma normal_exponent_bounds_real: "2 - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" subgoal by (rule order_trans[OF _ of_int_leI[OF normal_exponent_bounds_int(1)[OF that]]]) auto subgoal by (rule order_trans[OF of_int_leI[OF normal_exponent_bounds_int(2)[OF that]]]) auto done lemma float_eqI: "x = y" if "sign x = sign y" "fraction x = fraction y" "exponent x = exponent y" using that by transfer auto lemma float_induct[induct type:float, case_names normal denormal neg zero infinity nan]: fixes a::"('e, 'f)float" assumes normal: "\x. is_normal x \ valof x = normal_mantissa x * 2 powr normal_exponent x \ P x" assumes denormal: "\x. is_denormal x \ valof x = denormal_mantissa x * 2 powr denormal_exponent TYPE(('e, 'f)float) \ P x" assumes zero: "P 0" "P minus_zero" assumes infty: "P plus_infinity" "P minus_infinity" assumes nan: "\x. is_nan x \ P x" shows "P a" proof - from float_cases[of a] consider "is_nan a" | "is_infinity a" | "is_normal a" | "is_denormal a" | "is_zero a" by blast then show ?thesis proof cases case 1 then show ?thesis by (rule nan) next case 2 then consider "a = plus_infinity" | "a = minus_infinity" by (rule is_infinity_cases) then show ?thesis by cases (auto intro: infty) next case hyps: 3 from hyps have "valof a = normal_mantissa a * 2 powr normal_exponent a" by (cases a rule: sign_cases) (auto simp: valof_eq normal_mantissa_def normal_exponent_def is_normal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule normal) next case hyps: 4 from hyps have "valof a = denormal_mantissa a * 2 powr denormal_exponent TYPE(('e, 'f)float)" by (cases a rule: sign_cases) (auto simp: valof_eq denormal_mantissa_def denormal_exponent_def is_denormal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule denormal) next case 5 then consider "a = 0" | "a = minus_zero" by (rule is_zero_cases) then show ?thesis by cases (auto intro: zero) qed qed lemma infinite_infinity[simp]: "\ is_finite plus_infinity" "\ is_finite minus_infinity" by (auto simp: is_finite_def is_normal_def infinity_simps is_denormal_def is_zero_def) lemma nan_not_finite[simp]: "is_nan x \ \ is_finite x" using float_distinct_finite(1) by blast lemma valof_nonneg: "valof x \ 0" if "sign x = 0" for x::"('e, 'f)float" by (auto simp: valof_eq that) lemma valof_nonpos: "valof x \ 0" if "sign x = 1" for x::"('e, 'f)float" using that by (auto simp: valof_eq is_finite_def) lemma real_le_intI: "x \ y" if "floor x \ floor y" "x \ \" for x y::real using that(2,1) by (induction rule: Ints_induct) (auto elim!: Ints_induct simp: le_floor_iff) lemma real_of_int_le_2_powr_bitlenI: "real_of_int x \ 2 powr n - 1" if "bitlen (abs x) \ m" "m \ n" proof - have "real_of_int x \ abs (real_of_int x)" by simp also have "\ < 2 powr (bitlen (abs x))" by (rule abs_real_le_2_powr_bitlen) finally have "real_of_int x \ 2 powr (bitlen (abs x)) - 1" by (auto simp: powr_real_of_int bitlen_nonneg intro!: real_le_intI) also have "\ \ 2 powr m - 1" by (simp add: that) also have "\ \ 2 powr n - 1" by (simp add: that) finally show ?thesis . qed lemma largest_eq: "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr real_of_int (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" proof - have "2 ^ LENGTH('e) - 1 - 1 = (2::nat) ^ LENGTH('e) - 2" by arith then have "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr (real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f))" unfolding largest_def emax_eq bias_def by (auto simp: largest_def powr_realpow[symmetric] powr_minus divide_simps algebra_simps powr_diff powr_add) also have "2 ^ LENGTH('e) \ (2::nat)" by (simp add: self_le_power) then have "(real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f)) = (real (2 ^ LENGTH('e)) - 2 ^ (LENGTH('e) - 1) - LENGTH('f)) - 1" by auto also have "real (2 ^ LENGTH('e)) = 2 ^ LENGTH('e)" by auto also have "(2 ^ LENGTH('e) - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1) = real_of_int ((2 ^ (LENGTH('e) - 1) - int (LENGTH('f)) - 1))" by (simp, subst power_Suc[symmetric], simp) finally show ?thesis by simp qed lemma bitlen_denormal_mantissa: "bitlen (abs (denormal_mantissa x)) \ LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) - (auto simp: bitlen_le_iff_power nat_add_distrib abs_mult sign_simps - add_nonneg_pos denormal_mantissa_def intro!: antisym) + (auto simp: bitlen_le_iff_power denormal_mantissa_def intro!: antisym) qed lemma float_le_topfloat: fixes a::"('e, 'f)float" assumes "is_finite a" "LENGTH('e) > 1" shows "a \ topfloat" using assms(1) proof (induction a rule: float_induct) case (normal x) note normal(2) also have "real_of_int (normal_mantissa x) * 2 powr real_of_int (normal_exponent x) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using normal_exponent_bounds_real(2)[OF \is_normal x\] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero) also have "\ = largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (auto simp: powr_realpow powr_add) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI normal finite_topfloat) next case (denormal x) note denormal(2) also have "3 \ 2 powr (1 + real (LENGTH('e) - Suc 0))" proof - have "3 \ 2 powr (2::real)" by simp also have "\ \ 2 powr (1 + real (LENGTH('e) - Suc 0))" using assms by (subst powr_le_cancel_iff) auto finally show ?thesis . qed then have "real_of_int (denormal_mantissa x) * 2 powr real_of_int (denormal_exponent TYPE(('e, 'f)float)) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using bitlen_denormal_mantissa[of x] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero denormal_exponent_def bias_def powr_mult_base) also have "\ \ largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (rule mult_mono) (auto simp: powr_realpow powr_add power_Suc[symmetric] simp del: power_Suc) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI denormal finite_topfloat) qed auto lemma float_val_le_largest: "valof a \ largest TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" by (metis that finite_topfloat float_le float_le_topfloat valof_topfloat) lemma float_val_lt_threshold: "valof a < threshold TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" proof - have "valof a \ largest TYPE(('e, 'f)float)" by (rule float_val_le_largest [OF that]) also have "\ < threshold TYPE(('e, 'f)float)" by (auto simp: largest_def threshold_def divide_simps) finally show ?thesis . qed subsection \Algebraic properties about basic arithmetic\ text \Commutativity of addition.\ lemma assumes "is_finite a" "is_finite b" shows float_plus_comm_eq: "a + b = b + a" and float_plus_comm: "is_finite (a + b) \ (a + b) \ (b + a)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "a + b = b + a" by (simp add: float_defs fadd_def plus_float_def add.commute) then show "is_finite (a + b) \ (a + b) \ (b + a)" by (metis float_eq) qed text \The floating-point number a falls into the same category as the negation of \a\.\ lemma is_zero_uminus[simp]: "is_zero (- a) \ is_zero a" by (simp add: is_zero_def) lemma is_infinity_uminus [simp]: "is_infinity (- a) = is_infinity a" by (simp add: is_infinity_def) lemma is_finite_uminus[simp]: "is_finite (- a) \ is_finite a" by (simp add: is_finite_def) lemma is_nan_uminus[simp]: "is_nan (- a) \ is_nan a" by (simp add: is_nan_def) text \The sign of a and the sign of a's negation is different.\ lemma float_neg_sign: "(sign a) \ (sign (- a))" by (cases a rule: sign_cases) (auto simp: sign_minus_float) lemma float_neg_sign1: "sign a = sign (- b) \ sign a \ sign b" by (metis float_neg_sign sign_cases) lemma valof_uminus: assumes "is_finite a" shows "valof (- a) = - valof a" (is "?L = ?R") by (cases a rule: sign_cases) (auto simp: valof_eq sign_minus_float) text \Showing \a + (- b) = a - b\.\ lemma float_neg_add: "is_finite a \ is_finite b \ is_finite (a - b) \ valof a + valof (- b) = valof a - valof b" by (simp add: valof_uminus) lemma float_plus_minus: assumes "is_finite a" "is_finite b" "is_finite (a - b)" shows "(a + - b) \ (a - b)" proof - have nab: "\ is_nan a" "\ is_nan (- b)" "\ is_infinity a" "\ is_infinity (- b)" using assms by (auto simp: finite_nan finite_infinity) have "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then (sign a) else 0) (round To_nearest (valof a - valof b)))" using nab by (auto simp: minus_float_def fsub_def) also have "\ = (zerosign (if is_zero a \ is_zero (- b) \ sign a = sign (- b) then sign a else 0) (round To_nearest (valof a + valof (- b))))" using assms by (simp add: float_neg_sign1 float_neg_add) also have "\ = a + - b" using nab by (auto simp: float_defs fadd_def plus_float_def) finally show ?thesis using assms by (metis float_eq) qed lemma finite_bottomfloat: "is_finite bottomfloat" by (simp add: finite_topfloat) lemma bottomfloat_eq_m_largest: "valof (bottomfloat::('e, 'f)float) = - largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that by (auto simp: valof_topfloat valof_uminus finite_topfloat) lemma float_val_ge_bottomfloat: "valof a \ valof (bottomfloat::('e, 'f)float)" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" proof - have "- a \ topfloat" using that by (auto intro: float_le_topfloat) then show ?thesis using that by (auto simp: valof_uminus finite_topfloat) qed lemma float_ge_bottomfloat: "is_finite a \ a \ bottomfloat" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" by (metis finite_bottomfloat float_le float_val_ge_bottomfloat that) lemma float_val_ge_largest: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a \ - largest TYPE(('e,'f)float)" proof - have "- largest TYPE(('e,'f)float) = valof (bottomfloat::('e,'f)float)" using assms by (simp add: bottomfloat_eq_m_largest) also have "\ \ valof a" using assms by (simp add: float_val_ge_bottomfloat) finally show ?thesis . qed lemma float_val_gt_threshold: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a > - threshold TYPE(('e,'f)float)" proof - have largest: "valof a \ -largest TYPE(('e,'f)float)" using assms by (metis float_val_ge_largest) then have "-largest TYPE(('e,'f)float) > - threshold TYPE(('e,'f)float)" by (auto simp: bias_def threshold_def largest_def divide_simps) then show ?thesis by (metis largest less_le_trans) qed text \Showing \abs (- a) = abs a\.\ lemma float_abs [simp]: "\ is_nan a \ abs (- a) = abs a" by (metis IEEE.abs_float_def float_neg_sign1 minus_minus_float zero_simps(1)) lemma neg_zerosign: "- (zerosign s a) = zerosign (1 - s) (- a)" by (auto simp: zerosign_def) subsection \Properties about Rounding Errors\ definition error :: "('e, 'f)float itself \ real \ real" where "error _ x = valof (round To_nearest x::('e, 'f)float) - x" lemma bound_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" assumes finite: "is_finite a" shows "\valof (round To_nearest x::('e, 'f)float) - x\ \ \valof a - x\" proof - have *: "(round To_nearest x::('e,'f)float) = closest valof (\a. even (fraction a)) {a. is_finite a} x" using threshold finite by auto have "is_closest (valof) {a. is_finite a} x (round To_nearest x::('e,'f)float)" using is_finite_nonempty unfolding * by (intro closest_is_closest) auto then show ?thesis using finite is_closest_def by (metis mem_Collect_eq) qed lemma error_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" and "is_finite a" shows "\error TYPE(('e, 'f)float) x\ \ \valof a - x\ " unfolding error_def by (rule bound_at_worst_lemma; fact) lemma error_is_zero [simp]: fixes a::"('e, 'f)float" assumes "is_finite a" "1 < LENGTH('e)" shows "error TYPE(('e, 'f)float) (valof a) = 0" proof - have "\valof a\ < threshold TYPE(('e, 'f)float)" by (metis abs_less_iff minus_less_iff float_val_gt_threshold float_val_lt_threshold assms) then show ?thesis by (metis abs_le_zero_iff abs_zero diff_self error_at_worst_lemma assms(1)) qed lemma is_finite_zerosign[simp]: "is_finite (zerosign s a) \ is_finite a" by (auto simp: zerosign_def is_finite_def) lemma is_finite_closest: "is_finite (closest (v::_\real) p {a. is_finite a} x)" using closest_is_closest[OF is_finite_nonempty, of v x p] by (auto simp: is_closest_def) lemma defloat_float_zerosign_round_finite: assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" shows "is_finite (zerosign s (round To_nearest x::('e, 'f)float))" proof - have "(round To_nearest x::('e, 'f)float) = (closest valof (\a. even (fraction a)) {a. is_finite a} x)" using threshold by (metis (full_types) abs_less_iff leD le_minus_iff round.simps(1)) then have "is_finite (round To_nearest x::('e, 'f)float)" by (metis is_finite_closest) then show ?thesis using is_finite_zerosign by auto qed lemma valof_zero[simp]: "valof 0 = 0" "valof minus_zero = 0" by (auto simp add: zerosign_def valof_eq zero_simps) lemma signzero_zero: "is_zero a \ valof (zerosign s a) = 0" by (auto simp add: zerosign_def) lemma val_zero: "is_zero a \ valof a = 0" by (cases a rule: is_zero_cases) auto lemma float_add: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a + valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_add: "is_finite (a + b)" and error_float_add: "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "(a + b) = (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)))" using assms by (auto simp add: float_defs fadd_def plus_float_def) then show "is_finite ((a + b))" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a + b) = valof (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof (cases "is_zero (round To_nearest (valof a + valof b)::('e, 'f)float)") case True have "valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b) = valof (round To_nearest (valof a + valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_sub: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a - valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_sub: "is_finite (a - b)" and error_float_sub: "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then have ab: "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)))" using assms by (auto simp add: float_defs fsub_def minus_float_def) then show "is_finite (a - b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a - b) = valof (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof (cases "is_zero (round To_nearest (valof a - valof b)::('e, 'f)float)") case True have "valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b) = valof (round To_nearest (valof a - valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_mul: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a * valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_mul: "is_finite (a * b)" and error_float_mul: "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof - have non: "\ is_nan a" "\ is_nan b" "\ is_infinity a" "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "a * b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" using assms by (auto simp: float_defs fmul_def times_float_def) then show "is_finite (a * b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a * b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" by (auto simp: ab float_defs of_bool_def) show "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof (cases "is_zero (round To_nearest (valof a * valof b)::('e, 'f)float)") case True have "valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b) = valof (round To_nearest (valof a * valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_div: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and not_zero: "\ is_zero b" and threshold: "\valof a / valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_div: "is_finite (a / b)" and error_float_div: "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof - have ab: "a / b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b)))" using assms by (simp add: divide_float_def fdiv_def finite_infinity finite_nan not_zero float_defs [symmetric]) then show "is_finite (a / b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a / b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b))::('e, 'f)float)" by (auto simp: ab float_defs of_bool_def) show "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof (cases "is_zero (round To_nearest (valof a / valof b)::('e, 'f)float)") case True have "valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b) = valof (round To_nearest (valof a / valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma valof_one[simp]: "valof (1::('e, 'f)float) = (if LENGTH('e) \ 1 then 0 else 1)" by transfer (auto simp: bias_def unat_sub word_1_le_power p2_eq_1) end diff --git a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy --- a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy +++ b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy @@ -1,1092 +1,1092 @@ (* File: Irrationality_J_Hancl.thy Authors: Angeliki Koutsoukou-Argyraki and Wenda Li, Computer Laboratory, University of Cambridge, UK. Email: ak2110@cam.ac.uk, wl302@cam.ac.uk *) section \Irrational Rapidly Convergent Series\ theory Irrationality_J_Hancl imports "HOL-Analysis.Analysis" "HOL-Decision_Procs.Approximation" begin text \This is the formalisation of a proof by J. Hancl, in particular of the proof of his Theorem 3 in the paper: Irrational Rapidly Convergent Series, Rend. Sem. Mat. Univ. Padova, Vol 107 (2002). The statement asserts the irrationality of the sum of a series consisting of rational numbers defined using sequences that fulfill certain properties. Even though the statement is number-theoretic, the proof uses only arguments from introductory Analysis.\ text \We show the central result (theorem Hancl3) by proof by contradiction, by making use of some of the auxiliary lemmas. To this end, assuming that the sum is a rational number, for a quantity $\textrm{ALPHA}(n)$ we show that $\textrm{ALPHA}(n) \geq 1$ for all $n \in \mathbb{N}$. After that we show that we can find an $n \in \mathbb{N}$ for which $\textrm{ALPHA}(n) < 1$ which yields a contradiction and we thus conclude that the sum of the series is a rational number. We finally give an immediate application of theorem Hancl3 for a specific series (corollary Hancl3corollary, requiring lemma summable\_ln\_plus) which corresponds to Corollary 2 in the original paper by J. Hancl. \ hide_const floatarith.Max subsection \Misc\ lemma filterlim_sequentially_iff: "filterlim f F1 sequentially \ filterlim (\x. f (x+k)) F1 sequentially" unfolding filterlim_iff apply (subst eventually_sequentially_seg) by auto lemma filterlim_realpow_sequentially_at_top: "(x::real) > 1 \ filterlim ((^) x) at_top sequentially" apply (rule LIMSEQ_divide_realpow_zero[THEN filterlim_inverse_at_top,of _ 1,simplified]) by auto lemma filterlim_at_top_powr_real: fixes g::"'b \ real" assumes "filterlim f at_top F" "(g \ g') F" "g'>1" shows "filterlim (\x. g x powr f x) at_top F" proof - have "filterlim (\x. ((g'+1)/2) powr f x) at_top F" proof (subst filterlim_at_top_gt[of _ _ 1],rule+) fix Z assume "Z>(1::real)" have "\\<^sub>F x in F. ln Z \ ln (((g' + 1) / 2) powr f x)" using assms(1)[unfolded filterlim_at_top,rule_format,of "ln Z / ln ((g' + 1) / 2)"] apply (eventually_elim) using \g'>1\ by (auto simp:ln_powr divide_simps) then show "\\<^sub>F x in F. Z \ ((g' + 1) / 2) powr f x" apply (eventually_elim) apply (subst (asm) ln_le_cancel_iff) using \Z>1\ \g'>1\ by auto qed moreover have "\\<^sub>F x in F. ((g'+1)/2) powr f x \ g x powr f x" proof - have "\\<^sub>F x in F. g x > (g'+1)/2" apply (rule order_tendstoD[OF assms(2)]) using \g'>1\ by auto moreover have "\\<^sub>F x in F. f x>0" using assms(1)[unfolded filterlim_at_top_dense,rule_format,of 0] . ultimately show ?thesis apply eventually_elim using \g'>1\ by (auto intro!: powr_mono2) qed ultimately show ?thesis using filterlim_at_top_mono by fast qed lemma powrfinitesum: fixes a::real and s::nat assumes "s \ n" shows " (\j=s..(n::nat).(a powr (2^j))) = a powr (\j=s..(n::nat).(2^j)) " using \s \ n\ proof(induct n) case 0 then show ?case by auto next case (Suc n) have ?case when "s\n" using Suc.hyps by (metis Suc.prems add.commute linorder_not_le powr_add prod.nat_ivl_Suc' sum.cl_ivl_Suc that) moreover have ?case when "s=Suc n" proof- have "(\j = s..Suc n. a powr 2 ^ j) =(a powr 2 ^(Suc n)) " using \s=Suc n\ by simp also have "(a powr 2 ^(Suc n) ) = a powr sum ((^) 2) {s..Suc n} " using that by auto ultimately show " (\j = s..Suc n. a powr 2 ^ j) = a powr sum ((^) 2) {s..Suc n}" using \s\Suc n\ by linarith qed ultimately show ?case using \s\Suc n\ by linarith qed lemma summable_ratio_test_tendsto: fixes f :: "nat \ 'a::banach" assumes "c < 1" and "\n. f n\0" and "(\n. norm (f (Suc n)) / norm (f n)) \ c" shows "summable f" proof - obtain N where N_dist:"\n\N. dist (norm (f (Suc n)) / norm (f n)) c < (1-c)/2" using assms unfolding tendsto_iff eventually_sequentially by (meson diff_gt_0_iff_gt zero_less_divide_iff zero_less_numeral) have "norm (f (Suc n)) / norm (f n) \ (1+c)/2" when "n\N" for n using N_dist[rule_format,OF that] \c<1\ apply (auto simp add:field_simps dist_norm) by argo then have "norm (f (Suc n)) \ (1+c)/2 * norm (f n)" when "n\N" for n using that assms(2)[rule_format, of n] by (auto simp add:divide_simps) moreover have "(1+c)/2 < 1" using \c<1\ by auto ultimately show ?thesis using summable_ratio_test[of _ N f] by blast qed lemma summable_ln_plus: fixes f::"nat \ real" assumes "summable f" "\n. f n>0" shows "summable (\n. ln (1+f n))" proof (rule summable_comparison_test_ev[OF _ assms(1)]) have "ln (1 + f n) > 0" for n by (simp add: assms(2) ln_gt_zero) moreover have "ln (1 + f n) \ f n" for n apply (rule ln_add_one_self_le_self2) using assms(2)[rule_format,of n] by auto ultimately show "\\<^sub>F n in sequentially. norm (ln (1 + f n)) \ f n" by (auto intro!: eventuallyI simp add:less_imp_le) qed lemma suminf_real_offset_le: fixes f :: "nat \ real" assumes f: "\i. 0 \ f i" and "summable f" shows "(\i. f (i + k)) \ suminf f" proof - have "(\n. \i (\i. f (i + k))" using summable_sums[OF \summable f\] by (simp add: assms(2) summable_LIMSEQ summable_ignore_initial_segment) moreover have "(\n. \i (\i. f i)" using summable_sums[OF \summable f\] by (simp add: sums_def atLeast0LessThan f) then have "(\n. \i (\i. f i)" by (rule LIMSEQ_ignore_initial_segment) ultimately show ?thesis proof (rule LIMSEQ_le, safe intro!: exI[of _ k]) fix n assume "k \ n" have "(\ii (\i. i + k)) i)" by simp also have "\ = (\i\(\i. i + k) ` {.. \ sum f {..i sum f {.. n" shows " (\i=s..n. 2^i) < (2^(n+1) :: real) " using assms proof (induct n) case 0 show ?case by auto next case (Suc n) have ?case when "s=n+1" using that by auto moreover have ?case when "s \ n+1" proof - have"(\i=s..(n+1). 2^i ) = (\i=s..n. 2^i ) + (2::real)^(n+1) " using sum.cl_ivl_Suc \s \ Suc n \ by (auto simp add:add.commute) also have "... < (2) ^ (n +1) + (2)^(n+1)" proof - have "s \n" using that \s \ Suc n \ by auto then show ?thesis using Suc.hyps \s \ n\ by linarith qed also have "... = 2^(n+2)" by simp finally show "(\i=s..(Suc n). 2^i )< (2::real)^(Suc n+1)" by auto qed ultimately show ?case by blast qed subsection \Auxiliary lemmas and the main proof\ lemma showpre7: fixes a b ::"nat\int " and q p::int assumes "q>0" and "p>0"and a: "\n. a n>0" and "\n. b n>0" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*(suminf (\(j::nat). (b (j+n+1)/ a (j+n+1 )))) = ((\j=1..n. of_int( a j)))*(p -q* (\j=1..n. b j / a j)) " proof - define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "(\j. ff (j+n)) = (\n. ff n) - sum ff {..j=1..n. ff (j-1))" proof - have "sum ff {..j=1..n. ff (j-1))" apply (subst sum_bounds_lt_plus1[symmetric]) by simp then show ?thesis unfolding ff_def by auto qed finally have "(\j. ff (j + n)) = p / q - (\j = 1..n. ff (j - 1))" . then have "q*(\j. ff (j + n)) = p - q*(\j = 1..n. ff (j - 1))" using \q>0\ by (auto simp add:field_simps) then have "aa*q*(\j. ff (j + n)) = aa*(p - q*(\j = 1..n. ff (j - 1)))" by auto then show ?thesis unfolding aa_def ff_def by auto qed lemma show7: fixes d::"nat\real" and a b::"nat\int " and q p::int assumes "q \1" and "p \ 1" and a: "\n. a n \ 1" and "\n. b n \ 1" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*( suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1 )))) \ 1 " (is "?L \ _") proof- define LL where "LL=?L" define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "?L > 0" proof - have "aa > 0" unfolding aa_def using a apply (induct n,auto) by (simp add: int_one_le_iff_zero_less prod_pos) moreover have "(\j. ff (j + n)) > 0" proof (rule suminf_pos) have "summable ff" unfolding ff_def using assumerational using summable_def by blast then show " summable (\j. ff (j + n))" using summable_iff_shift[of ff n] by auto show " \i. 0 < ff (i + n)" unfolding ff_def using a assms(4) int_one_le_iff_zero_less by auto qed ultimately show ?thesis unfolding aa_def ff_def using \q\1\ by auto qed moreover have "?L \ \" proof - have "?L = aa *(p -q* ( \j=1..n. b j / a j))" unfolding aa_def apply (rule showpre7) using assms int_one_le_iff_zero_less by auto also have "... = aa * p - q * (\j=1..n. aa * b j / a j)" by (auto simp add:algebra_simps sum_distrib_left) also have "... = prod a {1..n} * p - q * (\j = 1..n. b j * prod a ({1..n} - {j}))" proof - have "(\j=1..n. aa * b j / a j) = (\j=1..n. b j * prod a ({1..n} - {j}))" unfolding of_int_sum proof (rule sum.cong) fix x assume "x \ {1..n}" then have "(\i = 1..n. real_of_int (a i)) = a x * (\i\{1..n} - {x}. real_of_int (a i))" apply (rule_tac prod.remove) by auto then have "aa / real_of_int (a x) = prod a ({1..n} - {x})" unfolding aa_def using a[rule_format,of x] by (auto simp add:field_simps) then show "aa * b x / a x = b x * prod a ({1..n} - {x})" by (metis mult.commute of_int_mult times_divide_eq_right) qed simp moreover have "aa * p = (\j = 1..n. (a j)) * p" unfolding aa_def by auto ultimately show ?thesis by force qed also have "... \ \" using Ints_of_int by blast finally show ?thesis . qed ultimately show ?thesis apply (fold LL_def) by (metis Ints_cases int_one_le_iff_zero_less not_less of_int_0_less_iff of_int_less_1_iff) qed lemma show8: fixes d ::"nat\real " and a :: "nat\int" and s k::nat assumes "A > 1" and d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "convergent_prod d" and assu2: "\n \ s. ( A/ (of_int (a n)) powr(1/of_int (2^n)))> prodinf (\j. d(n +j))" shows "\n\s. prodinf (\j. d(j+n)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n}))" proof(rule,rule) fix n assume "s \ n" define sp where "sp = (\n. prodinf (\j. d(j+n)))" define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have "sp i \ sp n" when "i\n" for i proof - have "(\j. d (j + i)) = (\ia. d (ia + (n - i) + i)) * (\iaconvergent_prod d\ convergent_prod_iff_shift[of d i] by simp subgoal for j using d[rule_format,of "j+i"] by auto done then have "sp i = sp n * (\jn\i\ by (auto simp:algebra_simps) moreover have "sp i>1" "sp n>1" unfolding sp_def using convergent_prod_iff_shift \convergent_prod d\ d by (auto intro!:less_1_prodinf) moreover have "(\j1" apply (rule prod_ge_1) using d less_imp_le by auto ultimately show ?thesis by auto qed moreover have "\j\s. A / ff j > sp j" unfolding ff_def sp_def using assu2 by (auto simp:algebra_simps) ultimately have "\j. s\j \ j\n \ A / ff j > sp n" by force then show "sp n< A / Max (ff ` {s..n})" by (metis (mono_tags, hide_lams) Max_in \n\s\ atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty order_refl) qed lemma auxiliary1_9: fixes d ::"nat\real" and a::"nat\int " and s m::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and "n \ m" and " m \ s" and auxifalse_assu: "\n\m. (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (d (n+1))* (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} ))" shows "(of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (\j=(m+1)..(n+1). d j) * (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..m}))" proof- define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have [simp]:"ff j > 0" for j unfolding ff_def by (metis a less_numeral_extra(3) of_int_0_less_iff powr_gt_zero) have ff_asm:"ff (n+1) < d (n+1) * Max (ff ` {s..n})" when "n\m" for n using auxifalse_assu that unfolding ff_def by simp from \n\m\ have Q: "(Max( ff ` {s..n} )) \ (\j=(m+1)..n. d j)* (Max (ff ` {s..m}))" proof(induct n) case 0 then show ?case using \m\s\ by simp next case (Suc n) have ?case when "m=Suc n" using that by auto moreover have ?case when "m\Suc n" proof - have "m \ n" using that Suc(2) by simp then have IH:"Max (ff ` {s..n}) \ prod d {m + 1..n} * Max (ff ` {s..m})" using Suc(1) by linarith have "Max (ff ` {s..Suc n}) = Max (ff ` {s..n} \ {ff (Suc n)})" using Suc.prems assms(5) atLeastAtMostSuc_conv by auto also have "... = max (Max (ff ` {s..n})) (ff (Suc n))" using Suc.prems assms(5) max_def sup_assoc that by auto also have "... \ max (Max (ff ` {s..n})) (d (n+1) * Max (ff ` {s..n}))" apply (rule max.mono) using ff_asm[of n] \ m \ Suc n\ that \s\m\ by auto also have "... \ Max (ff ` {s..n}) * max 1 (d (n+1))" proof - have "Max (ff ` {s..n}) \0" by (metis (mono_tags, hide_lams) Max_in \\j. 0 < ff j\ \m \ n\ assms(5) atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty less_eq_real_def) then show ?thesis using max_mult_distrib_right by (simp add: max_mult_distrib_right mult.commute) qed also have "... = Max (ff ` {s..n}) * d (n+1)" using d[rule_format, of "n+1"] by auto also have "... \ prod d {m + 1..n} * Max (ff ` {s..m}) * d (n+1)" using IH d[rule_format, of "n+1"] by auto also have "... = prod d {m + 1..n+1} * Max (ff ` {s..m})" using \n\m\ by (simp add:prod.nat_ivl_Suc' algebra_simps) finally show ?case by simp qed ultimately show ?case by blast qed then have "d (n+1) * Max (ff ` {s..n} ) \ (\j=(m+1)..(n+1). d j)* (Max (ff ` {s..m}))" using \m\n\ d[rule_format,of "Suc n"] by (simp add:prod.nat_ivl_Suc') then show ?thesis using ff_asm[of n] \s\m\ \m\n\ unfolding ff_def by auto qed lemma show9: fixes d ::"nat\real " and a :: "nat\int " and s ::nat and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and "convergent_prod d" and 8: "\n\s. prodinf (\j. d( n+j)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n})) " shows "\m \s. \n\m. ( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" proof (rule ccontr) define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" assume assumptioncontra: " \ (\m \s. \n\m. ( (ff (n+1)) \ (d (n+1))* (Max ( ff ` {s..n}))))" then obtain t where "t\s" and ttt: " \n\t. ( (ff (n+1)) < (d (n+1))* (Max ( ff ` {s..n} ) ))" by fastforce define B where "B=prodinf (\j. d(t+1+j))" have "B>0" unfolding B_def apply (rule less_0_prodinf) subgoal using convergent_prod_iff_shift[of d "t+1"] \convergent_prod d\ by (auto simp:algebra_simps) subgoal using d le_less_trans zero_le_one by blast done have "A \ B * Max ( ff ` {s..t})" proof (rule tendsto_le[of sequentially "\n. (\j=(t+1)..(n+1). d j) * Max ( ff ` {s..t})" _ "\n. ff (n+1)"]) show "(\n. ff (n + 1)) \ A" using assu1[folded ff_def] LIMSEQ_ignore_initial_segment by blast have "(\n. prod d {t + 1..n + 1}) \ B" proof - have "(\n. \i\n. d (t + 1 + i)) \ B" apply (rule convergent_prod_LIMSEQ[of "(\j. d(t+1+j))",folded B_def]) using \convergent_prod d\ convergent_prod_iff_shift[of d "t+1"] by (simp add:algebra_simps) then have "(\n. \i\{0..n}. d (i+(t + 1))) \ B" using atLeast0AtMost by (auto simp:algebra_simps) then have "(\n. prod d {(t + 1)..n + (t + 1)}) \ B" apply (subst (asm) prod.shift_bounds_cl_nat_ivl[symmetric]) by simp from seq_offset_neg[OF this,of "t"] show "(\n. prod d {t + 1..n+1}) \ B" apply (elim Lim_transform) apply (rule LIMSEQ_I) apply (rule exI[where x="t+1"]) by auto qed then show "(\n. prod d {t + 1..n + 1} * Max (ff ` {s..t})) \ B * Max (ff ` {s..t})" by (auto intro:tendsto_eq_intros) have "\\<^sub>F n in sequentially. (ff (n+1)) < (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" unfolding eventually_sequentially ff_def using auxiliary1_9[OF d a \s>0\ _ \t\s\ ttt[unfolded ff_def]] by blast then show "\\<^sub>F n in sequentially. (ff (n+1)) \ (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" by (eventually_elim,simp) qed simp also have "... \ B * Max ( ff ` {s..t+1})" proof - have "Max (ff ` {s..t}) \ Max (ff ` {s..t + 1})" apply (rule Max_mono) using \t\s\ by auto then show ?thesis using \B>0\ by auto qed finally have "A \ B * Max (ff ` {s..t + 1})" unfolding B_def . moreover have "B < A / Max (ff ` {s..t + 1})" using 8[rule_format, of "t+1",folded ff_def B_def] \s\t\ by auto moreover have "Max (ff ` {s..t+1})>0" using \A \ B * Max (ff ` {s..t + 1})\ \B>0\ \A>1\ - by (smt linordered_field_class.sign_simps(44)) + zero_less_mult_pos [of B "Max (ff ` {s..Suc t})"] + by simp ultimately show False by (auto simp add:field_simps) qed lemma show10: fixes d ::"nat\real " and a ::"nat\int " and s::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and " s>0" and 9: "\m \s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" shows "\m\s. \n\m. (((d (n+1))powr(2^(n+1))) * (\j=1..n. of_int( a j)) * (1/ (\j=1..s-1. (of_int( a j) )))) \ (a (n+1)) " proof (rule,rule) fix m assume "s \ m" from 9[rule_format,OF this] obtain n where "n\m" and asm_9:"( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" by auto with \s\m\ have "s\n" by auto have prod:"(\j=1..n. real_of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) ))) = (\j=s..n. of_int( a j))" proof - define f where "f= (\j. real_of_int( a j))" have "{Suc 0..n}= {Suc 0..s - Suc 0} \ {s..n}" using \n\s\ \s >0\ by auto then have "(\j=1..n. f j) = (\j=1..s-1. f j) * (\j=s..n. f j)" apply (subst prod.union_disjoint[symmetric]) by auto moreover have "(\j=1..s-1. f j) > 0 " apply (rule linordered_semidom_class.prod_pos) using a unfolding f_def by auto then have "(\j=1..s-1. f j) \ 0" by argo ultimately show ?thesis unfolding f_def by auto qed then have " (((d (n+1))powr(2^(n+1) )) * (\j=1..n. of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) )))) =(((d (n+1))powr(2^(n+1) )) * (\j=s..n. of_int( a j)))" proof - define f where "f= (\j. real_of_int( a j))" define c where "c = (d (n+1))powr(2^(n+1))" show ?thesis using prod apply (fold f_def c_def) by (metis mult.assoc) qed also have "... \ ((d (n+1))powr(2^(n+1) ) * (\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )) powr(2^i)) )" proof (rule mult_left_mono) show "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto show "(\j = s..n. real_of_int (a j)) \ (\i = s..n. Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i)" proof (rule prod_mono) fix i assume i: "i \ {s..n}" have "real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i" unfolding powr_powr by (simp add: a less_eq_real_def) also have "\ \ (Max(( \ (j::nat). ( real_of_int( a j) powr(1 /real_of_int (2^j)))) ` {s..n}))powr(2^i)" proof (rule powr_mono2) show "real_of_int (a i) powr (1 / real_of_int (2 ^ i)) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n})" apply (rule Max_ge) apply auto using i by blast qed simp_all finally have "real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" . then show "0 \ real_of_int (a i) \ real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" using a i by (metis \real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i\ powr_ge_pzero) qed qed also have "... = ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat). ( of_int( a j) powr (1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof- have "((d (n+1))powr(2^(n+1) ))\1 " by (metis Transcendental.log_one d le_powr_iff zero_le_numeral zero_le_power zero_less_one) moreover have "(\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } ) ) powr(2^i)) = (Max(( \ (j::nat). ( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" show ?thesis apply (fold ff_def) using \s\n\ powrfinitesum by auto qed ultimately show ?thesis by auto qed also have "... \ ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat).( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n })) powr(2^(n+1)) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" have " sum ((^) 2) {s..n} < (2::real) ^ (n + 1)" using factt \s\n\ by auto moreover have "1 \ ff" proof - define S where "S=(\(j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n }" have "finite S" unfolding S_def by auto moreover have "S\{}" unfolding S_def using \s\n\ by auto moreover have "\x\S. x\1" proof- have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ 1 " apply (rule ge_one_powr_ge_zero) apply auto by (simp add: a int_one_le_iff_zero_less) moreover have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ S" unfolding S_def apply (rule rev_image_eqI[where x=s]) using \s\n\ by auto ultimately show ?thesis by auto qed ultimately show ?thesis using Max_ge_iff[of S 1] unfolding S_def ff_def by blast qed moreover have "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto ultimately show ?thesis apply(fold ff_def) apply (rule mult_left_mono) apply (rule powr_mono) by auto qed also have "... = ((d (n+1)) * (Max((\j. (of_int( a j) powr(1 /of_int (2^j)))) ` {s..n}))) powr(2^(n+1))" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo moreover have "Max ss \0" proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed ultimately show ?thesis apply (fold ss_def) using powr_mult by auto qed also have "... \ (real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))) powr 2 ^ (n + 1)" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" show ?thesis proof (fold ss_def,rule powr_mono2) have "Max ss \0" \ \NOTE: we are repeating the same proof, so it may be a good idea to put this conclusion in an outer block so that it can be reused (without reproving).\ proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed moreover have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo ultimately show "0 \ (d (n + 1)) * Max ss" by auto show "(d (n + 1)) * Max ss \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" using asm_9[folded ss_def] . qed simp qed also have "... = (of_int (a (n+1)))" by (simp add: a less_eq_real_def pos_add_strict powr_powr) finally show "\n\m. d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / (\j = 1..s - 1. real_of_int (a j))) \ real_of_int (a (n + 1))" using \n\m\ \m\s\ apply (rule_tac x=n in exI) by auto qed lemma lasttoshow: fixes d ::"nat\real " and a b ::"nat\int " and q::int and s::nat assumes d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "q>0" and "A > 1" and b:"\n. b n>0" and 9: "\m\s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" and assu3: "filterlim( \ n. (d n)^(2^n)/ b n) at_top sequentially " and 5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" shows "\n. q*((\j=1..n. real_of_int(a j))) * suminf (\(j::nat). (b (j+n+1)/ a (j+n+1)))<1" proof- define as where "as=(\j = 1..s - 1. real_of_int (a j))" obtain n where "n\s" and n_def1:"real_of_int q * as * 2 * real_of_int (b (n + 1)) / d (n + 1) powr 2 ^ (n + 1) < 1" and n_def2:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" and n_def3:"(\j. (b (n+1 + j)) / (a (n+1 + j))) \ 2 * b (n+1) / a (n+1)" proof - have *:"(\n. real_of_int (b n) / d n ^ 2 ^ n) \ 0" using tendsto_inverse_0_at_top[OF assu3] by auto then have "(\n. real_of_int (b n) / d n powr 2 ^ n) \ 0" proof - have "d n ^ 2 ^ n = d n powr (of_nat (2 ^ n))" for n apply (subst powr_realpow) using d[rule_format, of n] by auto then show ?thesis using * by auto qed from tendsto_mult_right_zero[OF this,of "q * as * 2"] have "(\n. q * as * 2 * b n / d n powr 2 ^ n) \ 0" by auto then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1" by (elim order_tendstoD) simp then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" using 5 by eventually_elim auto then obtain N where N_def:"\n\N. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding eventually_sequentially by auto obtain n where "n\s" and "n\N" and n_def:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" using show10[OF d a \s>0\ 9,folded as_def,rule_format,of "max s N"] by auto with N_def[rule_format, of "n+1"] that[of n] show ?thesis by auto qed define pa where "pa = (\j = 1..n. real_of_int (a j))" define dn where "dn = d (n + 1) powr 2 ^ (n + 1)" have [simp]:"dn >0" "as > 0" subgoal unfolding dn_def by (metis d not_le numeral_One powr_gt_zero zero_le_numeral) subgoal unfolding as_def by (simp add: a prod_pos) done have [simp]:"pa>0" unfolding pa_def using a by (simp add: prod_pos) have K: "q* (\j=1..n. real_of_int (a j)) * suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1))) \q* (\j=1..n. real_of_int (a j)) *2* (b (n+1))/(a( n+1))" apply (fold pa_def) using mult_left_mono[OF n_def3,of "real_of_int q * pa"] \n\s\ \pa>0\ \q>0\ by (auto simp add:algebra_simps) also have KK:"... \ 2*q* (\j=1..n. real_of_int (a j)) * (b(n+1))* ((\j=1..s-1. real_of_int( a j))/((d (n+1))powr(2^(n+1)) * (\j=1..n. real_of_int ( a j))))" proof - have " dn * pa * (1 / as) \ real_of_int (a (n + 1))" using n_def2 unfolding dn_def pa_def . then show ?thesis apply (fold pa_def dn_def as_def) using \pa>0\ \q>0\ a[rule_format, of "Suc n"] b[rule_format, of "Suc n"] by (auto simp add:divide_simps algebra_simps) qed also have KKK: "... = (q* ((\j=1..(s-1). real_of_int( a j)))*2 * (b (n+1)))/ ((d (n+1))powr(2^(n+1)))" apply (fold as_def pa_def dn_def) apply simp using \0 < pa\ by blast also have KKKK: "... < 1" using n_def1 unfolding as_def by simp finally show ?thesis by auto qed lemma fixes d ::"nat\real " and a b ::"nat\int " and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu3: "filterlim ( \ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows issummable: "summable (\j. b j / a j)" and show5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" proof- define c where "c = (\j. b j / a j)" have c_pos:"c j>0" for j using a b unfolding c_def by simp have c_ratio_tendsto:"(\n. c (n+1) / c n ) \ 0" proof - define nn where "nn=(\n. (2::int)^ (Suc n))" define ff where "ff=(\ n. (a n / a (Suc n)) powr(1 /nn n)*(d(Suc n)))" have nn_pos:"nn n>0" and ff_pos:"ff n>0" for n subgoal unfolding nn_def by simp subgoal unfolding ff_def using d[rule_format, of "Suc n"] a[rule_format, of n] a[rule_format, of "Suc n"] by auto done have ff_tendsto:"ff \ 1/ sqrt A" proof - have "(of_int (a n)) powr(1 / (nn n)) = sqrt(of_int (a n) powr(1 /of_int (2^n)))" for n unfolding nn_def using a - apply (simp add: powr_half_sqrt [symmetric]) - by (simp add: powr_half_sqrt [symmetric] linordered_field_class.sign_simps(24) powr_powr) + by (simp add: powr_half_sqrt [symmetric] powr_powr ac_simps) moreover have "(( \ n. sqrt(of_int (a n) powr(1 /of_int (2^n))))\ sqrt A) sequentially " using assu1 tendsto_real_sqrt by blast ultimately have "(( \ n. (of_int (a n)) powr(1 /of_int (nn n)))\ sqrt A) sequentially " by auto from tendsto_divide[OF this assu1[THEN LIMSEQ_ignore_initial_segment[where k=1]]] have "(\n. (a n / a (Suc n)) powr(1 /nn n)) \ 1/sqrt A" using \A>1\ a unfolding nn_def by (auto simp add:powr_divide inverse_eq_divide sqrt_divide_self_eq) moreover have "(\n. d (Suc n))\ 1" apply (rule convergent_prod_imp_LIMSEQ) using convergent_prod_iff_shift[of d 1] \convergent_prod d\ by auto ultimately show ?thesis unfolding ff_def by (auto intro:tendsto_eq_intros) qed have "(\n. (ff n) powr nn n) \ 0" proof - define aa where "aa=(1+1/sqrt A) / 2" have "eventually (\n. ff nA>1\ by (auto simp add:field_simps) moreover have "(\n. aa powr nn n) \ 0" proof - have "(\y. aa ^ (nat \ nn) y) \ 0" apply (rule tendsto_power_zero) subgoal unfolding nn_def comp_def apply (rule filterlim_subseq) by (auto intro:strict_monoI) subgoal unfolding aa_def using \A>1\ by auto done then show ?thesis proof (elim filterlim_mono_eventually) have "aa>0" unfolding aa_def using \A>1\ by (auto simp add:field_simps pos_add_strict) then show " \\<^sub>F x in sequentially. aa ^ (nat \ nn) x = aa powr real_of_int (nn x)" by (auto simp: powr_int order.strict_implies_order[OF nn_pos]) qed auto qed ultimately show ?thesis apply (elim metric_tendsto_imp_tendsto) apply (auto intro!:powr_mono2 elim!:eventually_mono) using nn_pos ff_pos by (meson le_cases not_le)+ qed then have "(\n. (d (Suc n))^(nat (nn n)) * (a n / a (Suc n))) \ 0" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in sequentially. ff x powr (nn x) = d (Suc x) ^ nat (nn x) * (a x / a (Suc x))" apply (rule eventuallyI) subgoal for x unfolding ff_def using a[rule_format,of x] a[rule_format,of "Suc x"] d[rule_format,of "Suc x"] nn_pos[of x] apply (auto simp add:field_simps powr_divide powr_powr powr_mult ) by (simp add: powr_int) done qed auto moreover have "(\n. b (Suc n) / (d (Suc n))^(nat (nn n))) \ 0" using tendsto_inverse_0_at_top[OF assu3,THEN LIMSEQ_ignore_initial_segment[where k=1]] unfolding nn_def by (auto simp add:field_simps nat_mult_distrib nat_power_eq) ultimately have "(\n. b (Suc n) * (a n / a (Suc n))) \ 0" apply - subgoal premises asm using tendsto_mult[OF asm,simplified] apply (elim filterlim_mono_eventually) using d by (auto simp add:algebra_simps,metis (mono_tags, lifting) always_eventually not_one_less_zero) done then have "(\n. (b (Suc n) / b n) * (a n / a (Suc n))) \ 0" apply (elim Lim_transform_bound[rotated]) apply (rule eventuallyI) subgoal for x using a[rule_format, of x] a[rule_format, of "Suc x"] b[rule_format, of x] b[rule_format, of "Suc x"] by (auto simp add:field_simps) done then show ?thesis unfolding c_def by (auto simp add:algebra_simps) qed from c_ratio_tendsto have "(\n. norm (b (Suc n) / a (Suc n)) / norm (b n / a n)) \ 0" unfolding c_def apply (elim Lim_transform_eventually[rotated]) apply (rule eventuallyI) using a b by (auto simp add:divide_simps abs_of_pos) from summable_ratio_test_tendsto[OF _ _ this] a b show "summable c" unfolding c_def apply auto by (metis less_irrefl) have "\\<^sub>F n in sequentially. (\j. c (n + j)) \ 2 * c n" proof - obtain N where N_ratio:"\n\N. c (n + 1) / c n < 1 / 2" proof - have "eventually (\n. c (n+1) / c n < 1/2) sequentially" using c_ratio_tendsto[unfolded tendsto_iff,rule_format, of "1/2",simplified] apply eventually_elim subgoal for n using c_pos[of n] c_pos[of "Suc n"] by auto done then show ?thesis using that unfolding eventually_sequentially by auto qed have "(\j. c (j + n)) \ 2 * c n" when "n\N" for n proof - have "(\j 2*c n * (1 - 1 / 2^(m+1))" for m proof (induct m) case 0 then show ?case using c_pos[of n] by simp next case (Suc m) have "(\ji c n + (\i c (i + n) / 2" for i using N_ratio[rule_format,of "i+n"] \n\N\ c_pos[of "i+n"] by simp then show ?thesis by (auto intro:sum_mono) qed also have "... = c n + (\i c n + c n * (1 - 1 / 2 ^ (m + 1))" using Suc by auto also have "... = 2 * c n * (1 - 1 / 2 ^ (Suc m + 1))" by (auto simp add:field_simps) finally show ?case . qed then have "(\j 2*c n" for m using c_pos[of n] by (smt divide_le_eq_1_pos divide_pos_pos nonzero_mult_div_cancel_left zero_less_power) moreover have "summable (\j. c (j + n))" using \summable c\ by (simp add: summable_iff_shift) ultimately show ?thesis using suminf_le_const[of "\j. c (j+n)" "2*c n"] by auto qed then show ?thesis unfolding eventually_sequentially by (auto simp add:algebra_simps) qed then show "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding c_def by simp qed theorem Hancl3: fixes d ::"nat\real " and a b ::"nat\int " assumes "A > 1" and d:"\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu2: "\n \ s. (A/ (of_int (a n)) powr(1 /of_int (2^n)))> prodinf (\j. d(n +j))" and assu3: "filterlim (\ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows "suminf(\ n. b n / a n ) \ Rats" proof(rule ccontr) assume asm:"\ (suminf(\ n. b n / a n ) \ Rats)" have ab_sum:"summable (\j. b j / a j)" using issummable[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] . obtain p q ::int where "q>0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" proof - from asm have "suminf(\ n. b n / a n ) \ Rats" by auto then have "suminf(\ n. b (n+1) / a (n+1)) \ Rats" apply (subst suminf_minus_initial_segment[OF ab_sum,of 1]) by auto then obtain p' q' ::int where "q'\0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p'/q')" unfolding Rats_eq_int_div_int using summable_ignore_initial_segment[OF ab_sum,of 1,THEN summable_sums] by force define p q where "p=(if q'<0 then - p' else p')" and "q=(if q'<0 then - q' else q')" have "p'/q'=p/q" "q>0" using \q'\0\ unfolding p_def q_def by auto then show ?thesis using that[of q p] pq_def by auto qed define ALPHA where "ALPHA = (\n. (of_int q)*((\j=1..n. of_int(a j)))*(suminf (\ (j::nat). (b (j+n+1)/a (j+n+1)))))" have "ALPHA n \ 1" for n proof - have "suminf(\ n. b (n+1) / a (n+1)) > 0" apply (rule suminf_pos) subgoal using summable_ignore_initial_segment[OF ab_sum,of 1] by auto subgoal using a b by simp done then have "p/q > 0" using sums_unique[OF pq_def,symmetric] by auto then have "q\1" "p\1" using \q>0\ by (auto simp add:divide_simps) moreover have "\n. 1 \ a n" "\n. 1 \ b n" using a b by (auto simp add: int_one_le_iff_zero_less) ultimately show ?thesis unfolding ALPHA_def using show7[OF _ _ _ _ pq_def] by auto qed moreover have "\n. ALPHA n < 1" unfolding ALPHA_def proof (rule lasttoshow[OF d a \s>0\ \q>0\ \A>1\ b _ assu3]) show "\\<^sub>F n in sequentially. (\j. real_of_int (b (n + j)) / real_of_int (a (n + j))) \ real_of_int (2 * b n) / real_of_int (a n)" using show5[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] by simp show "\m\s. \n\m. d (n + 1) * Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" apply (rule show9[OF \A>1\ d a \s>0\ assu1 \convergent_prod d\]) using show8[OF \A>1\ d a \s>0\ \convergent_prod d\ assu2] by (simp add:algebra_simps) qed ultimately show False using not_le by blast qed corollary Hancl3corollary: fixes A::real and a b ::"nat\int " assumes "A > 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "((\ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and asscor2: "\n \6. (of_int (a n)) powr(1 /of_int (2^n ))*(1+ 4*(2/3)^n) \ A \ (b n \2 powr((4/3)^(n-1)) ) " shows "suminf(\ n. b n / a n ) \ Rats" proof- define d::"nat\real" where "d= (\ n. 1+(2/3)^(n+1))" have dgt1:"\n. 1 < d n " unfolding d_def by auto moreover have "convergent_prod d" unfolding d_def apply (rule abs_convergent_prod_imp_convergent_prod) apply (rule summable_imp_abs_convergent_prod) using summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1] by simp moreover have "\n\6. (\j. d (n + j)) < A / real_of_int (a n) powr (1 / real_of_int (2 ^ n))" proof rule+ fix n::nat assume "6 \ n" have d_sum:"summable (\j. ln (d j))" unfolding d_def apply (rule summable_ln_plus) apply (rule summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1]) by simp have "(\j. ln (d (n + j))) < ln (1+4 * (2 / 3) ^ n)" proof - define c::real where "c=(2 / 3) ^ n" have "0n\6\ apply (subst power_add[symmetric]) by auto also have "... \ (2 / 3)^6" by (auto intro:power_le_one) also have "... < 1/8" by (auto simp add:field_simps) finally show "c<1/8" . qed (simp add:c_def) have "(\j. ln (d (n + j))) \ (\j. (2 / 3) ^ (n + j + 1))" apply (rule suminf_le) subgoal unfolding d_def apply (intro allI ln_add_one_self_le_self2 ) apply (rule order.strict_trans[of _ 0]) by auto subgoal using summable_ignore_initial_segment[OF d_sum,of n] by (auto simp add:algebra_simps) subgoal using summable_geometric[THEN summable_ignore_initial_segment,of "2/3" "n+1"] by (auto simp add:algebra_simps) done also have "... = (\j. (2 / 3)^(n+1)*(2 / 3) ^ j)" by (auto simp add:algebra_simps power_add) also have "... = (2 / 3)^(n+1) * (\j. (2 / 3) ^ j)" apply (rule suminf_mult) by (auto intro:summable_geometric) also have "... = 2 * c" unfolding c_def apply (subst suminf_geometric) by auto also have "... < 4 * c - (4 * c)\<^sup>2" using \0 \c<1/8\ by (sos "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1/16 * [1]^2))))") also have "... \ ln (1 + 4 * c)" apply (rule ln_one_plus_pos_lower_bound) using \0 \c<1/8\ by auto finally show ?thesis unfolding c_def by simp qed then have "exp (\j. ln (d (n + j))) < 1 + 4 * (2 / 3) ^ n" by (metis Groups.mult_ac(2) add.right_neutral add_mono_thms_linordered_field(5) divide_inverse divide_less_eq_numeral1(1) divide_pos_pos exp_gt_zero less_eq_real_def ln_exp ln_less_cancel_iff mult_zero_left rel_simps(27) rel_simps(76) zero_less_one zero_less_power) moreover have "exp (\j. ln (d (n + j))) = (\j. d (n + j))" proof (subst exp_suminf_prodinf_real [symmetric]) show "\k. 0 \ ln (d (n + k))" using dgt1 by (simp add: less_imp_le) show "abs_convergent_prod (\na. exp (ln (d (n + na))))" apply (subst exp_ln) subgoal for j using dgt1[rule_format,of "n+j"] by auto subgoal unfolding abs_convergent_prod_def real_norm_def apply (subst abs_of_nonneg) using convergent_prod_iff_shift[of d n] \convergent_prod d\ by (auto simp add: dgt1 less_imp_le algebra_simps) done show "(\na. exp (ln (d (n + na)))) = (\j. d (n + j))" apply (subst exp_ln) subgoal using dgt1 le_less_trans zero_le_one by blast subgoal by simp done qed ultimately have "(\j. d (n + j)) < 1 + 4 * (2 / 3) ^ n" by simp also have "... \ A / (a n) powr (1 / of_int (2 ^ n))" proof - have "a n powr (1 / real_of_int (2 ^ n)) > 0" using a[rule_format,of n] by auto then show ?thesis using asscor2[rule_format,OF \6\n\] by (auto simp add:field_simps) qed finally show "(\j. d (n + j)) < A / real_of_int (a n) powr (1 / of_int (2 ^ n))" . qed moreover have "LIM n sequentially. d n ^ 2 ^ n / real_of_int (b n) :> at_top" proof - have "LIM n sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) :> at_top" proof - define n1 where "n1=(\n. (2::real)* (3/2)^(n-1))" define n2 where "n2=(\n. ((4::real)/3)^(n-1))" have "LIM n sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) :> at_top" proof (rule filterlim_at_top_powr_real[where g'="exp (8/9) / (2::real)"]) define e1 where "e1=exp (8/9) / (2::real)" show "e1>1" unfolding e1_def by (approximation 4) show "(\n. ((1+(8/9)/(n1 n)) powr (n1 n))/2) \ e1" proof - have "(\n. (1+(8/9)/(n1 n)) powr (n1 n)) \ exp (8/9)" apply (rule filterlim_compose[OF tendsto_exp_limit_at_top]) unfolding n1_def by (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_realpow_sequentially_at_top simp:filterlim_sequentially_iff[of "\x. (3 / 2) ^ (x - Suc 0)" _ 1]) then show ?thesis unfolding e1_def by (intro tendsto_intros,auto) qed show "filterlim n2 at_top sequentially" unfolding n2_def apply (subst filterlim_sequentially_iff[of "\n. (4 / 3) ^ (n - 1)" _ 1]) by (auto intro:filterlim_realpow_sequentially_at_top) qed moreover have "\\<^sub>F n in sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) = d n ^ 2 ^ n / 2 powr((4/3)^(n-1))" proof (rule eventually_sequentiallyI[of 1]) fix x assume "x\(1::nat)" have " ((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = (((1 + 8 / 9 / n1 x) powr n1 x) powr n2 x) / 2 powr (4 / 3) ^ (x - 1)" apply (subst powr_divide) apply (simp_all add:n1_def n2_def) by (smt divide_nonneg_nonneg zero_le_power) also have "... = (1 + 8 / 9 / n1 x) powr (n1 x * n2 x) / 2 powr (4 / 3) ^ (x - 1)" apply (subst powr_powr) by simp also have "... = (1 + 8 / 9 / n1 x) powr (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" proof - have "n1 x * n2 x = 2 ^ x" unfolding n1_def n2_def apply (subst mult.assoc) apply (subst power_mult_distrib[symmetric]) using \x\1\ by (auto simp add:power_Suc[symmetric] simp del:power_Suc) then show ?thesis by simp qed also have "... = (1 + 8 / 9 / n1 x) ^ (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" apply (subst (3) powr_realpow[symmetric]) apply (simp_all add: n1_def) by (smt zero_le_divide_iff zero_le_power) also have "... = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1)" proof - define x1 where "x1=x-1" have *:"x=x1+1" unfolding x1_def using \x\1\ by simp have **: "8 / 9 / n1 x = (2 / 3) ^ (x + 1)" unfolding n1_def using \x\1\ apply (fold x1_def *[symmetric]) by (auto simp add:divide_simps) then show ?thesis unfolding d_def apply (subst **) by auto qed finally show "((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1) " . qed ultimately show ?thesis using filterlim_cong by fast qed moreover have "\\<^sub>F n in sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) \ d n ^ 2 ^ n / real_of_int (b n)" apply (rule eventually_sequentiallyI[of 6]) apply (rule divide_left_mono) subgoal for x using asscor2[rule_format,of x] by auto subgoal for x using \\n. 1 < d n\[rule_format, of x] by auto subgoal for x using b by auto done ultimately show ?thesis by (auto elim: filterlim_at_top_mono) qed ultimately show ?thesis using Hancl3[OF \A>1\ _ a b _ assu1,of d 6] by force qed end diff --git a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy --- a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy +++ b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy @@ -1,419 +1,419 @@ (* Author: Benedikt Seidl License: BSD *) section \Instantiation of the LTL to DRA construction\ theory DRA_Instantiation imports DRA_Construction LTL.Equivalence_Relations "HOL-Library.Log_Nat" begin text \We instantiate the construction locale with propositional equivalence and obtain a function converting a formula into an abstract automaton.\ global_interpretation ltl_to_dra: dra_construction "(\\<^sub>P)" rep_ltln\<^sub>P abs_ltln\<^sub>P defines ltl_to_dra = ltl_to_dra.ltl_to_dra and ltl_to_dra_alphabet = ltl_to_dra.ltl_to_dra_alphabet and \' = ltl_to_dra.\' and \\<^sub>1 = ltl_to_dra.\\<^sub>1 and \\<^sub>2 = ltl_to_dra.\\<^sub>2 and \\<^sub>3 = ltl_to_dra.\\<^sub>3 and \\<^sub>\_FG = ltl_to_dra.\\<^sub>\_FG and \\<^sub>\_GF = ltl_to_dra.\\<^sub>\_GF and af_letter\<^sub>G = ltl_to_dra.af_letter\<^sub>G and af_letter\<^sub>F = ltl_to_dra.af_letter\<^sub>F and af_letter\<^sub>G_lifted = ltl_to_dra.af_letter\<^sub>G_lifted and af_letter\<^sub>F_lifted = ltl_to_dra.af_letter\<^sub>F_lifted and af_letter\<^sub>\_lifted = ltl_to_dra.af_letter\<^sub>\_lifted and \ = ltl_to_dra.\ and af_letter\<^sub>\ = ltl_to_dra.af_letter\<^sub>\ by unfold_locales (meson Quotient_abs_rep Quotient_ltln\<^sub>P, simp add: Quotient_abs_rep Quotient_ltln\<^sub>P ltln\<^sub>P.abs_eq_iff) text \We obtain the following theorem:\ thm ltl_to_dra.ltl_to_dra_language text \Furthermore, we verify the size bound of the automaton to be double-exponential.\ lemma prop_equiv_nested_prop_atoms_finite: "finite {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" using prop_equiv_finite'[OF nested_prop_atoms_finite] . lemma prop_equiv_nested_prop_atoms_card: "card {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \} \ 2 ^ 2 ^ card (nested_prop_atoms \)" using prop_equiv_card'[OF nested_prop_atoms_finite] . lemma prop_equiv_nested_prop_atoms\<^sub>\_finite: "finite {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" using prop_equiv_finite'[OF nested_prop_atoms\<^sub>\_finite] by fast lemma prop_equiv_nested_prop_atoms\<^sub>\_card: "card {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X} \ 2 ^ 2 ^ card (nested_prop_atoms \)" (is "?lhs \ ?rhs") proof - have "finite {abs_ltln\<^sub>P \ | \. prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" by (simp add: prop_equiv_nested_prop_atoms\<^sub>\_finite nested_prop_atoms\<^sub>\_finite prop_equiv_finite) then have "?lhs \ card {abs_ltln\<^sub>P \ | \. prop_atoms \ \ (nested_prop_atoms\<^sub>\ \ X)}" using card_mono prop_equiv_subset by blast also have "\ \ 2 ^ 2 ^ card (nested_prop_atoms\<^sub>\ \ X)" using prop_equiv_card[OF nested_prop_atoms\<^sub>\_finite] by fast also have "\ \ ?rhs" using nested_prop_atoms\<^sub>\_card by auto finally show ?thesis . qed lemma \\<^sub>\_GF_nodes_finite: "finite (DBA.nodes (\\<^sub>\_GF \))" using finite_subset[OF ltl_to_dra.\\<^sub>\_GF_nodes prop_equiv_nested_prop_atoms_finite] . lemma \\<^sub>\_FG_nodes_finite: "finite (DCA.nodes (\\<^sub>\_FG \))" using finite_subset[OF ltl_to_dra.\\<^sub>\_FG_nodes prop_equiv_nested_prop_atoms_finite] . lemma \\<^sub>\_GF_nodes_card: "card (DBA.nodes (\\<^sub>\_GF \)) \ 2 ^ 2 ^ card (nested_prop_atoms (F\<^sub>n \))" using le_trans[OF card_mono[OF prop_equiv_nested_prop_atoms_finite ltl_to_dra.\\<^sub>\_GF_nodes] prop_equiv_nested_prop_atoms_card] . lemma \\<^sub>\_FG_nodes_card: "card (DCA.nodes (\\<^sub>\_FG \)) \ 2 ^ 2 ^ card (nested_prop_atoms (G\<^sub>n \))" using le_trans[OF card_mono[OF prop_equiv_nested_prop_atoms_finite ltl_to_dra.\\<^sub>\_FG_nodes] prop_equiv_nested_prop_atoms_card] . lemma \\<^sub>2_nodes_finite_helper: "list_all (finite \ DBA.nodes) (map (\\. \\<^sub>\_GF (\[set ys]\<^sub>\)) xs)" by (auto simp: list.pred_map list_all_iff \\<^sub>\_GF_nodes_finite) lemma \\<^sub>2_nodes_finite: "finite (DBA.nodes (\\<^sub>2 xs ys))" unfolding ltl_to_dra.\\<^sub>2_def using dbail_nodes_finite \\<^sub>2_nodes_finite_helper . lemma \\<^sub>3_nodes_finite_helper: "list_all (finite \ DCA.nodes) (map (\\. \\<^sub>\_FG (\[set xs]\<^sub>\)) ys)" by (auto simp: list.pred_map list_all_iff \\<^sub>\_FG_nodes_finite) lemma \\<^sub>3_nodes_finite: "finite (DCA.nodes (\\<^sub>3 xs ys))" unfolding ltl_to_dra.\\<^sub>3_def using dcail_finite \\<^sub>3_nodes_finite_helper . (* TODO add to HOL/Groups_List.thy *) lemma list_prod_mono: "f \ g \ (\x\xs. f x) \ (\x\xs. g x)" for f g :: "'a \ nat" by (induction xs) (auto simp: le_funD mult_le_mono) (* TODO add to HOL/Groups_List.thy *) lemma list_prod_const: "(\x. x \ set xs \ f x \ c) \ (\x\xs. f x) \ c ^ length xs" for f :: "'a \ nat" by (induction xs) (auto simp: mult_le_mono) (* TODO add to HOL/Finite_Set.thy *) lemma card_insert_Suc: "card (insert x S) \ Suc (card S)" by (metis Suc_n_not_le_n card.infinite card_insert_if finite_insert linear) (* TODO add to HOL/Power.thy *) lemma nat_power_le_imp_le: "0 < a \ a \ b \ x ^ a \ x ^ b" for x :: nat by (metis leD linorder_le_less_linear nat_power_less_imp_less neq0_conv power_eq_0_iff) (* TODO add to HOL/Power.thy *) lemma const_less_power: "n < x ^ n" if "x > 1" using that by (induction n) (auto simp: less_trans_Suc) (* TODO add to HOL-Library/Log_Nat.thy *) lemma floorlog_le_const: "floorlog x n \ n" by (induction n) (simp add: floorlog_eq_zero_iff, metis Suc_lessI floorlog_le_iff le_SucI power_inject_exp) lemma \\<^sub>2_nodes_card: assumes "length xs \ n" and "\\. \ \ set xs \ card (nested_prop_atoms \) \ n" shows "card (DBA.nodes (\\<^sub>2 xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 2)" proof - have 1: "\\. \ \ set xs \ card (nested_prop_atoms (F\<^sub>n \[set ys]\<^sub>\)) \ Suc n" proof - fix \ assume "\ \ set xs" have "card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))) \ Suc (card (nested_prop_atoms (\[set ys]\<^sub>\)))" by (simp add: card_insert_Suc) also have "\ \ Suc (card (nested_prop_atoms \))" by (simp add: FG_advice_nested_prop_atoms_card) also have "\ \ Suc n" by (simp add: assms(2) \\ \ set xs\) finally show "card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))) \ Suc n" . qed have "(\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\)))) \ (\\\xs. 2 ^ 2 ^ card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))))" by (rule list_prod_mono) (insert \\<^sub>\_GF_nodes_card le_fun_def, blast) also have "\ \ (2 ^ 2 ^ Suc n) ^ length xs" by (rule list_prod_const) (metis 1 Suc_leI nat_power_le_imp_le nat_power_eq_Suc_0_iff neq0_conv pos2 zero_less_power) also have "\ \ (2 ^ 2 ^ Suc n) ^ n" using assms(1) nat_power_le_imp_le by fastforce also have "\ = 2 ^ (n * 2 ^ Suc n)" by (metis Groups.mult_ac(2) power_mult) also have "\ \ 2 ^ (2 ^ floorlog 2 n * 2 ^ Suc n)" by (cases "n = 0") (auto simp: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ 2 ^ (Suc n + floorlog 2 n)" by (simp add: power_add) finally have 2: "(\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\)))) \ 2 ^ 2 ^ (Suc n + floorlog 2 n)" . have "card (DBA.nodes (\\<^sub>2 xs ys)) \ max 1 (length xs) * (\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\))))" using dbail_nodes_card[OF \\<^sub>2_nodes_finite_helper] by (auto simp: ltl_to_dra.\\<^sub>2_def comp_def) also have "\ \ max 1 n * 2 ^ 2 ^ (Suc n + floorlog 2 n)" using assms(1) 2 by (simp add: mult_le_mono) also have "\ \ 2 ^ (floorlog 2 n) * 2 ^ 2 ^ (Suc n + floorlog 2 n)" by (cases "n = 0") (auto simp: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ (floorlog 2 n + 2 ^ (Suc n + floorlog 2 n))" by (simp add: power_add) also have "\ \ 2 ^ (n + 2 ^ (Suc n + floorlog 2 n))" by (simp add: floorlog_le_const) also have "\ \ 2 ^ 2 ^ (n + floorlog 2 n + 2)" by simp (metis const_less_power Suc_1 add_Suc_right add_leE lessI less_imp_le_nat power_Suc) finally show ?thesis . qed lemma \\<^sub>3_nodes_card: assumes "length ys \ n" and "\\. \ \ set ys \ card (nested_prop_atoms \) \ n" shows "card (DCA.nodes (\\<^sub>3 xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 1)" proof - have 1: "\\. \ \ set ys \ card (DCA.nodes (\\<^sub>\_FG (\[set xs]\<^sub>\))) \ 2 ^ 2 ^ Suc n" proof - fix \ assume "\ \ set ys" have "card (nested_prop_atoms (G\<^sub>n \[set xs]\<^sub>\)) \ Suc (card (nested_prop_atoms (\[set xs]\<^sub>\)))" by (simp add: card_insert_Suc) also have "\ \ Suc (card (nested_prop_atoms \))" by (simp add: GF_advice_nested_prop_atoms_card) also have "\ \ Suc n" by (simp add: assms(2) \\ \ set ys\) finally have 2: "card (nested_prop_atoms (G\<^sub>n \[set xs]\<^sub>\)) \ Suc n" . then show "?thesis \" by (intro le_trans[OF \\<^sub>\_FG_nodes_card]) (meson one_le_numeral power_increasing) qed have "card (DCA.nodes (\\<^sub>3 xs ys)) \ (\\\ys. card (DCA.nodes (\\<^sub>\_FG (\[set xs]\<^sub>\))))" unfolding ltl_to_dra.\\<^sub>3_def using dcail_nodes_card[OF \\<^sub>3_nodes_finite_helper] by (auto simp: comp_def) also have "\ \ (2 ^ 2 ^ Suc n) ^ length ys" by (rule list_prod_const) (rule 1) also have "\ \ (2 ^ 2 ^ Suc n) ^ n" by (simp add: assms(1) power_increasing) also have "\ \ 2 ^ (n * 2 ^ Suc n)" by (metis le_refl mult.commute power_mult) also have "\ \ 2 ^ (2 ^ floorlog 2 n * 2 ^ Suc n)" by (cases \n > 0\) (simp_all add: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ 2 ^ (n + floorlog 2 n + 1)" by (simp add: power_add) finally show ?thesis . qed lemma \\<^sub>1_nodes_finite: "finite (DCA.nodes (\\<^sub>1 \ xs))" unfolding ltl_to_dra.\\<^sub>1_def by (metis (no_types, lifting) finite_subset ltl_to_dra.\_nodes finite_SigmaI prop_equiv_nested_prop_atoms\<^sub>\_finite prop_equiv_nested_prop_atoms_finite) lemma \\<^sub>1_nodes_card: assumes "card (subfrmlsn \) \ n" shows "card (DCA.nodes (\\<^sub>1 \ xs)) \ 2 ^ 2 ^ (n + 1)" proof - let ?fst = "{abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" let ?snd = "{abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ (set xs)}" have 1: "card (nested_prop_atoms \) \ n" by (meson card_mono[OF subfrmlsn_finite nested_prop_atoms_subfrmlsn] assms le_trans) have "card (DCA.nodes (\\<^sub>1 \ xs)) \ card (?fst \ ?snd)" unfolding ltl_to_dra.\\<^sub>1_def by (rule card_mono) (simp_all add: ltl_to_dra.\_nodes prop_equiv_nested_prop_atoms\<^sub>\_finite prop_equiv_nested_prop_atoms_finite) also have "\ = card ?fst * card ?snd" using prop_equiv_nested_prop_atoms\<^sub>\_finite card_cartesian_product by blast also have "\ \ 2 ^ 2 ^ card (nested_prop_atoms \) * 2 ^ 2 ^ card (nested_prop_atoms \)" using prop_equiv_nested_prop_atoms\<^sub>\_card prop_equiv_nested_prop_atoms_card mult_le_mono by blast also have "\ = 2 ^ 2 ^ (card (nested_prop_atoms \) + 1)" by (simp add: semiring_normalization_rules(36)) also have "\ \ 2 ^ 2 ^ (n + 1)" using assms 1 by simp finally show ?thesis . qed lemma \'_nodes_finite: "finite (DRA.nodes (\' \ xs ys))" unfolding ltl_to_dra.\'_def using dcai_nodes_finite dbcrai_nodes_finite using \\<^sub>1_nodes_finite \\<^sub>2_nodes_finite \\<^sub>3_nodes_finite by fast lemma \'_nodes_card: assumes "length xs \ n" and "\\. \ \ set xs \ card (nested_prop_atoms \) \ n" and "length ys \ n" and "\\. \ \ set ys \ card (nested_prop_atoms \) \ n" and "card (subfrmlsn \) \ n" shows "card (DRA.nodes (\' \ xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 4)" proof - have "n + 1 \ n + floorlog 2 n + 2" by auto then have 1: "(2::nat) ^ (n + 1) \ 2 ^ (n + floorlog 2 n + 2)" using one_le_numeral power_increasing by blast have "card (DRA.nodes (\' \ xs ys)) \ card (DCA.nodes (\\<^sub>1 \ xs)) * card (DBA.nodes (\\<^sub>2 xs ys)) * card (DCA.nodes (\\<^sub>3 xs ys))" (is "?lhs \ ?rhs") proof (unfold ltl_to_dra.\'_def) have "card (DBA.nodes (\\<^sub>2 xs ys)) * card (DCA.nodes (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))) \ ?rhs" by (simp add: dcai_nodes_card[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]) then show "card (DRA.nodes (dbcrai (\\<^sub>2 xs ys) (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys)))) \ ?rhs" by (meson dbcrai_nodes_card[OF \\<^sub>2_nodes_finite dcai_nodes_finite[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]] basic_trans_rules(23)) qed also have "\ \ 2 ^ 2 ^ (n + 1) * 2 ^ 2 ^ (n + floorlog 2 n + 2) * 2 ^ 2 ^ (n + floorlog 2 n + 1)" using \\<^sub>1_nodes_card[OF assms(5)] \\<^sub>2_nodes_card[OF assms(1,2)] \\<^sub>3_nodes_card[OF assms(3,4)] by (metis mult_le_mono) also have "\ = 2 ^ (2 ^ (n + 1) + 2 ^ (n + floorlog 2 n + 2) + 2 ^ (n + floorlog 2 n + 1))" by (metis power_add) also have "\ \ 2 ^ (4 * 2 ^ (n + floorlog 2 n + 2))" using 1 by auto finally show ?thesis by (simp add: numeral.simps(2) power_add) qed lemma subformula_nested_prop_atoms_subfrmlsn: "\ \ subfrmlsn \ \ nested_prop_atoms \ \ subfrmlsn \" using nested_prop_atoms_subfrmlsn subfrmlsn_subset by blast lemma ltl_to_dra_nodes_finite: "finite (DRA.nodes (ltl_to_dra \))" unfolding ltl_to_dra.ltl_to_dra_def apply (rule draul_nodes_finite) apply (simp add: split_def ltl_to_dra.\'_alphabet advice_sets_not_empty) apply (simp add: list.pred_set split_def \'_nodes_finite) done lemma ltl_to_dra_nodes_card: assumes "card (subfrmlsn \) \ n" shows "card (DRA.nodes (ltl_to_dra \)) \ 2 ^ 2 ^ (2 * n + floorlog 2 n + 4)" proof - let ?map = "map (\(x, y). \' \ x y) (advice_sets \)" have 1: "\x::nat. x > 0 \ x ^ length (advice_sets \) \ x ^ 2 ^ card (subfrmlsn \)" by (metis advice_sets_length linorder_not_less nat_power_less_imp_less) have "card (DRA.nodes (ltl_to_dra \)) \ prod_list (map (card \ DRA.nodes) ?map)" unfolding ltl_to_dra.ltl_to_dra_def apply (rule draul_nodes_card) unfolding set_map image_image split_def ltl_to_dra.\'_alphabet apply (simp add: advice_sets_not_empty) unfolding split_def list.pred_set using \'_nodes_finite by auto also have "\ = (\(x, y)\advice_sets \. card (DRA.nodes (\' \ x y)))" by (induction "advice_sets \") (auto, metis (no_types, lifting) comp_apply split_def) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ length (advice_sets \)" proof (rule list_prod_const, unfold split_def, rule \'_nodes_card) show "\x. x \ set (advice_sets \) \ length (fst x) \ n" using advice_sets_element_length assms by fastforce show "\x \. \x \ set (advice_sets \); \ \ set (fst x)\ \ card (nested_prop_atoms \) \ n" using advice_sets_element_subfrmlsn(1) assms subformula_nested_prop_atoms_subfrmlsn subformulas\<^sub>\_subfrmlsn by (metis (no_types, lifting) card_mono subfrmlsn_finite subset_iff sup.absorb_iff2 sup.coboundedI1 surjective_pairing) show "\x. x \ set (advice_sets \) \ length (snd x) \ n" using advice_sets_element_length assms by fastforce show "\x \. \x \ set (advice_sets \); \ \ set (snd x)\ \ card (nested_prop_atoms \) \ n" using advice_sets_element_subfrmlsn(2) assms subformula_nested_prop_atoms_subfrmlsn subformulas\<^sub>\_subfrmlsn by (metis (no_types, lifting) card_mono subfrmlsn_finite subset_iff sup.absorb_iff2 sup.coboundedI1 surjective_pairing) qed (insert assms, blast) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ (2 ^ card (subfrmlsn \))" by (simp add: 1) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ (2 ^ n)" by (simp add: assms power_increasing) also have "\ = 2 ^ (2 ^ n * 2 ^ (n + floorlog 2 n + 4))" - by (metis Rat.sign_simps(5) power_mult) + by (simp add: ac_simps power_mult [symmetric]) also have "\ = 2 ^ 2 ^ (2 * n + floorlog 2 n + 4)" - by (simp add: semiring_normalization_rules(26)) + by (simp add: power_add) (simp add: mult_2 power_add) finally show ?thesis . qed theorem ltl_to_dra_size: "card (DRA.nodes (ltl_to_dra \)) \ 2 ^ 2 ^ (2 * size \ + floorlog 2 (size \) + 4)" using ltl_to_dra_nodes_card subfrmlsn_card by blast end diff --git a/thys/LinearQuantifierElim/Thys/FRE.thy b/thys/LinearQuantifierElim/Thys/FRE.thy --- a/thys/LinearQuantifierElim/Thys/FRE.thy +++ b/thys/LinearQuantifierElim/Thys/FRE.thy @@ -1,204 +1,205 @@ (* Author: Tobias Nipkow, 2007 *) theory FRE imports LinArith begin subsection\Ferrante-Rackoff \label{sec:FRE}\ text\This section formalizes a slight variant of Ferrante and Rackoff's algorithm~\cite{FerranteR-SIAM75}. We consider equalities separately, which improves performance.\ fun between :: "real * real list \ real * real list \ real * real list" where "between (r,cs) (s,ds) = ((r+s)/2, (1/2) *\<^sub>s (cs+ds))" definition FR\<^sub>1 :: "atom fm \ atom fm" where "FR\<^sub>1 \ = (let as = R.atoms\<^sub>0 \; lbs = lbounds as; ubs = ubounds as; ebs = ebounds as; intrs = [subst \ (between l u) . l \ lbs, u \ ubs] in list_disj (inf\<^sub>- \ # inf\<^sub>+ \ # intrs @ map (subst \) ebs))" lemma dense_interval: assumes "finite L" "finite U" "l \ L" "u \ U" "l < x" "x < u" "P(x::real)" and dense: "\y l u. \ \y\{l<.. L; \y\{x<.. U; l \ P y" shows "\l\L.\u\U. l (\y. l y P y)" proof - let ?L = "{l:L. l < x}" let ?U = "{u:U. x < u}" let ?ll = "Max ?L" let ?uu = "Min ?U" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) moreover have "?U \ {}" using \u:U\ \x by (blast intro:order_less_imp_le) ultimately have "\y. ?ll y y \ L" "\y. x y y \ U" using \finite L\ \finite U\ by force+ moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?uu \ U" proof show "?uu \ ?U" using \finite U\ Min_in[OF _ \?U \ {}\] by simp show "?U \ U" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp moreover have "x < ?uu" using \finite U\ \?U \ {}\ by simp moreover have "?ll < ?uu" using \?ll \x by arith ultimately show ?thesis using \l < x\ \x < u\ \?L \ {}\ \?U \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed lemma dense: "\ nqfree f; \y\{l<.. LB f xs; \y\{x<.. UB f xs; l < x; x < u; x \ EQ f xs; R.I f (x#xs); l < y; y < u\ \ R.I f (y#xs)" proof(induct f) case (Atom a) show ?case proof (cases a) case (Less r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Less by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r < c*x + \cs,xs\" using Atom Less by simp { assume "c=0" hence ?thesis using Atom Less Cons by simp } moreover { assume "c<0" hence "x < (r - \cs,xs\)/c" (is "_ < ?u") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?u \ y" using Atom Less Cons \c<0\ by (auto simp add: field_simps) hence "?u < u" using \y by simp with \x show False using Atom Less Cons \c<0\ by(auto simp:depends\<^sub>R_def) qed } moreover { assume "c>0" hence "x > (r - \cs,xs\)/c" (is "_ > ?l") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?l \ y" using Atom Less Cons \c>0\ by (auto simp add: field_simps) hence "?l > l" using \y>l\ by simp with \?l show False using Atom Less Cons \c>0\ by (auto simp:depends\<^sub>R_def) qed } ultimately show ?thesis by force qed next case (Eq r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Eq by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r = c*x + \cs,xs\" using Atom Eq by simp { assume "c=0" hence ?thesis using Atom Eq Cons by simp } moreover { assume "c\0" hence ?thesis using \r = c*x + \cs,xs\\ Atom Eq Cons \l \y by(auto simp: ac_simps depends\<^sub>R_def split:if_splits) } ultimately show ?thesis by force qed qed next case (And f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ next case (Or f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ qed fastforce+ theorem I_FR\<^sub>1: assumes "nqfree \" shows "R.I (FR\<^sub>1 \) xs = (\x. R.I \ (x#xs))" (is "?FR = ?EX") proof assume ?FR { assume "R.I (inf\<^sub>- \) xs" hence ?EX using \?FR\ min_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "R.I (inf\<^sub>+ \) xs" hence ?EX using \?FR\ plus_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "\x \ EQ \ xs. R.I \ (x#xs)" hence ?EX using \?FR\ by(auto simp add:FR\<^sub>1_def) } moreover { assume "\R.I (inf\<^sub>- \) xs \ \R.I (inf\<^sub>+ \) xs \ (\x\EQ \ xs. \R.I \ (x#xs))" with \?FR\ obtain r cs s ds where "R.I (subst \ (between (r,cs) (s,ds))) xs" by(auto simp: FR\<^sub>1_def eval_def diff_divide_distrib set_ebounds I_subst \nqfree \\) blast hence "R.I \ (eval (between (r,cs) (s,ds)) xs # xs)" by(simp add:I_subst \nqfree \\) hence ?EX .. } ultimately show ?EX by blast next assume ?EX then obtain x where x: "R.I \ (x#xs)" .. { assume "R.I (inf\<^sub>- \) xs \ R.I (inf\<^sub>+ \) xs" hence ?FR by(auto simp:FR\<^sub>1_def) } moreover { assume "x \ EQ \ xs" then obtain r cs where "(r,cs) \ set(ebounds(R.atoms\<^sub>0 \)) \ x = r + \cs,xs\" by(force simp:set_ebounds field_simps) moreover hence "R.I (subst \ (r,cs)) xs" using x by(auto simp: I_subst \nqfree \\ eval_def) ultimately have ?FR by(force simp:FR\<^sub>1_def) } moreover { assume "\ R.I (inf\<^sub>- \) xs" and "\ R.I (inf\<^sub>+ \) xs" and "x \ EQ \ xs" obtain l where "l \ LB \ xs" "l < x" using LBex[OF \nqfree \\ x \\ R.I (inf\<^sub>- \) xs\ \x \ EQ \ xs\] .. obtain u where "u \ UB \ xs" "x < u" using UBex[OF \nqfree \\ x \\ R.I (inf\<^sub>+ \) xs\ \x \ EQ \ xs\] .. have "\l\LB \ xs. \u\UB \ xs. l (\y. l < y \ y < u \ R.I \ (y#xs))" using dense_interval[where P = "\x. R.I \ (x#xs)", OF finite_LB finite_UB \l:LB \ xs\ \u:UB \ xs\ \l \x x] x dense[OF \nqfree \\ _ _ _ _ \x \ EQ \ xs\] by simp then obtain r c cs s d ds where "Less r (c # cs) \ set (R.atoms\<^sub>0 \)" "Less s (d # ds) \ set (R.atoms\<^sub>0 \)" "\y. (r - \cs,xs\) / c < y \ y < (s - \ds,xs\) / d \ R.I \ (y # xs)" and *: "c > 0" "d < 0" "(r - \cs,xs\) / c < (s - \ds,xs\) / d" by blast moreover have "(r - \cs,xs\) / c < eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs" (is ?P) and "eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs < (s - \ds,xs\) / d" (is ?Q) proof - - from * have [simp]: "c * (c * (d * (d * 4))) > 0" by (auto simp add: sign_simps) + from * have [simp]: "c * (c * (d * (d * 4))) > 0" + by (simp add: sign_simps) from * have "c * s + d * \cs,xs\ < d * r + c * \ds,xs\" by (simp add: field_simps) with * have "(2 * c * c * d) * (d * r + c * \ds,xs\) < (2 * c * c * d) * (c * s + d * \cs,xs\)" and "(2 * c * d * d) * (c * s + d * \cs,xs\) < (2 * c * d * d) * (d * r + c * \ds,xs\)" by simp_all with * show ?P and ?Q by (auto simp add: field_simps eval_def iprod_left_add_distrib) qed ultimately have ?FR by (fastforce simp: FR\<^sub>1_def bex_Un set_lbounds set_ubounds set_ebounds I_subst \nqfree \\) } ultimately show ?FR by blast qed definition "FR = R.lift_nnf_qe FR\<^sub>1" lemma qfree_FR\<^sub>1: "nqfree \ \ qfree (FR\<^sub>1 \)" apply(simp add:FR\<^sub>1_def) apply(rule qfree_list_disj) apply(auto simp:qfree_min_inf qfree_plus_inf set_ubounds set_lbounds set_ebounds image_def qfree_map_fm) done theorem I_FR: "R.I (FR \) xs = R.I \ xs" by(simp add:I_FR\<^sub>1 FR_def R.I_lift_nnf_qe qfree_FR\<^sub>1) theorem qfree_FR: "qfree (FR \)" by(simp add:FR_def R.qfree_lift_nnf_qe qfree_FR\<^sub>1) end diff --git a/thys/Multi_Party_Computation/Secure_Multiplication.thy b/thys/Multi_Party_Computation/Secure_Multiplication.thy --- a/thys/Multi_Party_Computation/Secure_Multiplication.thy +++ b/thys/Multi_Party_Computation/Secure_Multiplication.thy @@ -1,799 +1,800 @@ subsection \Secure multiplication protocol\ theory Secure_Multiplication imports CryptHOL.Cyclic_Group_SPMF Uniform_Sampling Semi_Honest_Def begin locale secure_mult = fixes q :: nat assumes q_gt_0: "q > 0" and "prime q" begin type_synonym real_view = "nat \ nat \ ((nat \ nat \ nat \ nat) \ nat \ nat) spmf" type_synonym sim = "nat \ nat \ ((nat \ nat \ nat \ nat) \ nat \ nat) spmf" lemma samp_uni_set_spmf [simp]: "set_spmf (sample_uniform q) = {..< q}" by(simp add: sample_uniform_def) definition funct :: "nat \ nat \ (nat \ nat) spmf" where "funct x y = do { s \ sample_uniform q; return_spmf (s, (x*y + (q - s)) mod q)}" definition TI :: "((nat \ nat) \ (nat \ nat)) spmf" where "TI = do { a \ sample_uniform q; b \ sample_uniform q; r \ sample_uniform q; return_spmf ((a, r), (b, ((a*b + (q - r)) mod q)))}" definition out :: "nat \ nat \ (nat \ nat) spmf" where "out x y = do { ((c1,d1),(c2,d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; return_spmf (((x*e1 + (q - d1)) mod q), ((e2 * c2 + (q - d2)) mod q))}" definition R1 :: "real_view" where "R1 x y = do { ((c1, d1), (c2, d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((x, c1, d1, e1), s1, s2)}" definition S1 :: "nat \ nat \ (nat \ nat \ nat \ nat) spmf" where "S1 x s1 = do { a :: nat \ sample_uniform q; e1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf (x, a, d1, e1)}" definition Out1 :: "nat \ nat \ nat \ (nat \ nat) spmf" where "Out1 x y s1 = do { let s2 = (x*y + (q - s1)) mod q; return_spmf (s1,s2)}" definition R2 :: "real_view" where "R2 x y = do { ((c1, d1), (c2, d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((y, c2, d2, e2), s1, s2)}" definition S2 :: "nat \ nat \ (nat \ nat \ nat \ nat) spmf" where "S2 y s2 = do { b \ sample_uniform q; e2 \ sample_uniform q; let d2 = (e2*b + (q - s2)) mod q; return_spmf (y, b, d2, e2)}" definition Out2 :: "nat \ nat \ nat \ (nat \ nat) spmf" where "Out2 y x s2 = do { let s1 = (x*y + (q - s2)) mod q; return_spmf (s1,s2)}" definition Ideal2 :: "nat \ nat \ nat \ ((nat \ nat \ nat \ nat) \ (nat \ nat)) spmf" where "Ideal2 y x out2 = do { view2 :: (nat \ nat \ nat \ nat) \ S2 y out2; out2 \ Out2 y x out2; return_spmf (view2, out2)}" sublocale sim_non_det_def: sim_non_det_def R1 S1 Out1 R2 S2 Out2 funct . lemma minus_mod: assumes "a > b" shows "[a - b mod q = a - b] (mod q)" by(metis cong_diff_nat cong_def le_trans less_or_eq_imp_le assms mod_less_eq_dividend mod_mod_trivial) lemma q_cong:"[a = q + a] (mod q)" by (simp add: cong_def) lemma q_cong_reverse: "[q + a = a] (mod q)" by (simp add: cong_def) lemma qq_cong: "[a = q*q + a] (mod q)" by (simp add: cong_def) lemma minus_q_mult_cancel: assumes "[a = e + b - q * c - d] (mod q)" and "e + b - d > 0" and "e + b - q * c - d > 0" shows "[a = e + b - d] (mod q)" proof- have "a mod q = (e + b - q * c - d) mod q" using assms(1) cong_def by blast then have "a mod q = (e + b - d) mod q" by (metis (no_types) add_cancel_left_left assms(3) diff_commute diff_is_0_eq' linordered_semidom_class.add_diff_inverse mod_add_left_eq mod_mult_self1_is_0 nat_less_le) then show ?thesis using cong_def by blast qed lemma s1_s2: assumes "x < q" "a < q" "b < q" and r:"r < q" "y < q" shows "((x + a) mod q * b + q - (a * b + q - r) mod q) mod q = (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q" proof- have s: "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q = ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" proof- have lhs: "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q = (x*b + r) mod q" proof- let ?h = "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q" have "[?h = x * y + q - (x * ((y + (q - b)) mod q) + (q - r)) mod q] (mod q)" by(simp add: assms(1) cong_def q_gt_0) then have "[?h = x * y + q - (x * (y + (q - b)) + (q - r)) mod q] (mod q)" by (metis mod_add_left_eq mod_mult_right_eq) then have no_qq: "[?h = x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by(metis distrib_left) then have "[?h = q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" proof- have "[x * y + q - (x * y + x * (q - b) + (q - r)) mod q = q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by (smt qq_cong add.assoc cong_diff_nat cong_def le_add2 le_trans mod_le_divisor q_gt_0) then show ?thesis using cong_trans no_qq by blast qed then have mod: "[?h = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def add.assoc add.commute le_add2 le_trans mod_add_self2 mod_le_divisor q_gt_0) then have "[?h = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r))] (mod q)" proof- have 1: "q \ q - b" using assms by simp then have "q*q \ x*(q-b)" "q \ q - r" using 1 assms apply (auto simp add: mult_strict_mono) by (simp add: mult_le_mono) then have "q + q*q + x * y + q > x * y + x * (q - b) + (q - r)" using assms(5) by linarith then have "[q + q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r))] (mod q)" using minus_mod by blast then show ?thesis using mod using cong_trans by blast qed then have "[?h = q + q*q + x * y + q - (x * y + (x * q - x*b) + (q - r))] (mod q)" by (simp add: right_diff_distrib') then have "[?h = q + q*q + x * y + q - x * y - (x * q - x*b) - (q - r)] (mod q)" by simp then have mod': "[?h = q + q*q + q - (x * q - x*b) - (q - r)] (mod q)" by(simp add: add.commute) then have neg: "[?h = q + q*q + q - x * q + x*b - (q - r)] (mod q)" proof- have "[q + q*q + q - (x * q - x*b) - (q - r) = q + q*q + q - x * q + x*b - (q - r)] (mod q)" proof(cases "x = 0") case True then show ?thesis by simp next case False have "x * q - x*b > 0" using False assms by simp also have "q + q*q + q - x * q > 0" by (metis assms(1) add.commute diff_mult_distrib2 less_Suc_eq mult.commute mult_Suc_right nat_0_less_mult_iff q_gt_0 zero_less_diff) ultimately show ?thesis by simp qed then show ?thesis using mod' cong_trans by blast qed then have "[?h = q + q*q + q + x*b - (q - r)] (mod q)" proof- have "[q + q*q + q - x * q + x*b - (q - r) = q + q*q + q + x*b - (q - r)] (mod q)" proof(cases "x = 0") case True then show ?thesis by simp next case False have "q*q > x*q" using False assms by (simp add: mult_strict_mono) then have 1: "q + q*q + q - x * q + x*b - (q - r) > 0" by linarith then have 2: "q + q*q + q + x*b - (q - r) > 0" by simp then show ?thesis by (smt 1 2 Nat.add_diff_assoc2 \x * q < q * q\ add_cancel_left_left add_diff_inverse_nat le_add1 le_add2 le_trans less_imp_add_positive less_numeral_extra(3) minus_mod minus_q_mult_cancel mod_if mult.commute q_gt_0) qed then show ?thesis using cong_trans neg by blast qed then have "[?h = q + q*q + q + x*b - q + r] (mod q)" by (metis r(1) Nat.add_diff_assoc2 Nat.diff_diff_right le_add2 less_imp_le_nat semiring_normalization_rules(23)) then have "[?h = q + q*q + q + x*b + r] (mod q)" apply(simp add: cong_def) by (metis (no_types, lifting) add.assoc add.commute add_diff_cancel_right' diff_is_0_eq' mod_if mod_le_divisor q_gt_0) then have "[?h = x*b + r] (mod q)" apply(simp add: cong_def) by (metis mod_add_cong mod_add_self1 mod_mult_self1) then show ?thesis by (simp add: cong_def assms) qed also have rhs: "((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q = (x*b + r) mod q" proof- let ?h = "((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" have "[?h = (x + a) mod q * b + q - (a * b + (q - r)) mod q] (mod q)" by (simp add: q_gt_0 assms(1) cong_def) then have "[?h = (x + a) * b + q - (a * b + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def mod_add_cong mod_le_divisor mod_mult_left_eq q_gt_0 assms) then have "[?h = x*b + a*b + q - (a * b + (q - r)) mod q] (mod q)" by(metis distrib_right) then have mod: "[?h = q + x*b + a*b + q - (a * b + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def add.assoc add.commute le_add2 le_trans mod_add_self2 mod_le_divisor q_gt_0) then have "[?h = q + x*b + a*b + q - (a * b + (q - r))] (mod q)" using q_cong assms(1) proof- have ge: "q + x*b + a*b + q > (a * b + (q - r))" using assms by simp then have "[ q + x*b + a*b + q - (a * b + (q - r)) mod q = q + x*b + a*b + q - (a * b + (q - r))] (mod q)" using Divides.mod_less_eq_dividend cong_diff_nat cong_def le_trans less_not_refl2 less_or_eq_imp_le q_gt_0 minus_mod by presburger then show ?thesis using mod cong_trans by blast qed then have "[?h = q + x*b + q - (q - r)] (mod q)" by (simp add: add.commute) then have "[?h = q + x*b + q - q + r] (mod q)" by (metis Nat.add_diff_assoc2 Nat.diff_diff_right r(1) le_add2 less_imp_le_nat) then have "[?h = q + x*b + r] (mod q)" by simp then have "[?h = q + (x*b + r)] (mod q)" using add.assoc by metis then have "[?h = x*b + r] (mod q)" using cong_def q_cong_reverse by metis then show ?thesis by (simp add: cong_def assms(1)) qed ultimately show ?thesis by simp qed have lhs:"((x + a) mod q * b + q - (a * b + q - r) mod q) mod q = ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" using assms by simp have rhs: "(x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q = (x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q " using assms by simp have " ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q = (x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q" using assms s[symmetric] by blast then show ?thesis using lhs rhs by metis qed lemma s1_s2_P2: assumes "x < q" "xa < q" "xb < q" "xc < q" "y < q" shows "((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, ((x + xb) mod q * xa + q - (xb * xa + q - xc) mod q) mod q) = ((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q)" using assms s1_s2 by metis lemma c1: assumes "e2 = (x + c1) mod q" and "x < q" "c1 < q" shows "c1 = (e2 + q - x) mod q" proof- have "[e2 + q = x + c1] (mod q)" by(simp add: assms cong_def) then have "[e2 + q - x = c1] (mod q)" proof- have "e2 + q \ x" using assms by simp then show ?thesis by (metis \[e2 + q = x + c1] (mod q)\ cong_add_lcancel_nat le_add_diff_inverse) qed then show ?thesis by(simp add: cong_def assms) qed lemma c1_P2: assumes "xb < q" "xa < q" "xc < q" "x < q" shows "((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q) = ((y, xa, (((x + xb) mod q + q - x) mod q * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q)" proof- have "(xb * xa + q - xc) mod q = (((x + xb) mod q + q - x) mod q * xa + q - xc) mod q" using assms c1 by simp then show ?thesis using assms by metis qed lemma minus_mod_cancel: assumes "a - b > 0" "a - b mod q > 0" shows "[a - b + c = a - b mod q + c] (mod q)" proof- have "(b - b mod q + (a - b)) mod q = (0 + (a - b)) mod q" using cong_def mod_add_cong neq0_conv q_gt_0 by (simp add: minus_mod_eq_mult_div) then show ?thesis by (metis (no_types) Divides.mod_less_eq_dividend Nat.add_diff_assoc2 add_diff_inverse_nat assms(1) cong_def diff_is_0_eq' less_or_eq_imp_le mod_add_cong neq0_conv) qed lemma d2: assumes d2: "d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q" and s1: "s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q" and s2: "s2 = (x*y + (q - s1)) mod q" and x: "x < q" and y: "y < q" and r: "r < q" and b: "b < q" and e2: "e2 < q" shows "d2 = (e2*b + (q - s2)) mod q" proof- have s1_le_q: "s1 < q" using s1 q_gt_0 by simp have s2_le_q: "s2 < q" using s2 q_gt_0 by simp have xb: "(x*b) mod q = (s2 + (q - r)) mod q" proof- have "s1 = (x*(y + (q - b)) + (q - r)) mod q" using s1 b by (metis mod_add_left_eq mod_mult_right_eq) then have s1_dist: "s1 = (x*y + x*(q - b) + (q - r)) mod q" by(metis distrib_left) then have "s1 = (x*y + x*q - x*b + (q - r)) mod q" using distrib_left b diff_mult_distrib2 by auto then have "[s1 = x*y + x*q - x*b + (q - r)] (mod q)" by(simp add: cong_def) then have "[s1 + x * b = x*y + x*q - x*b + x*b + (q - r)] (mod q)" by (metis add.commute add.left_commute cong_add_lcancel_nat) then have "[s1 + x*b = x*y + x*q + (q - r)] (mod q)" - by (metis add.commute add_lessD1 add_mult_distrib b left_diff_distrib' less_imp_le linordered_field_class.sign_simps(18) linordered_semidom_class.add_diff_inverse mult.commute not_less) + using b by (simp add: algebra_simps) + (metis add_diff_inverse_nat diff_diff_left diff_mult_distrib2 less_imp_add_positive mult.commute not_add_less1 zero_less_diff) then have s1_xb: "[s1 + x*b = q + x*y + x*q + (q - r)] (mod q)" by (smt mod_add_cong mod_add_self1 cong_def) then have "[x*b = q + x*y + x*q + (q - r) - s1] (mod q)" proof- have "q + x*y + x*q + (q - r) - s1 > 0" using s1_le_q by simp then show ?thesis by (metis add_diff_inverse_nat less_numeral_extra(3) s1_xb cong_add_lcancel_nat nat_diff_split) qed then have "[x*b = x*y + x*q + (q - r) + q - s1] (mod q)" by (metis add.assoc add.commute) then have "[x*b = x*y + (q - r) + q - s1] (mod q)" by (smt Nat.add_diff_assoc cong_def less_imp_le_nat mod_mult_self1 s1_le_q semiring_normalization_rules(23)) then have "(x*b) mod q = (x*y + (q - r) + q - s1) mod q" by(simp add: cong_def) then have "(x*b) mod q = (x*y + (q - r) + (q - s1)) mod q" using add.assoc s1_le_q by auto then have "(x*b) mod q = (x*y + (q - s1) + (q - r)) mod q" using add.commute by presburger then show ?thesis using s2 by presburger qed have "d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q" using d2 by simp then have "d2 = (((e2 + q - x))*b + (q - r)) mod q" using mod_add_cong mod_mult_left_eq by blast then have "d2 = (e2*b + q*b - x*b + (q - r)) mod q" by (simp add: distrib_right diff_mult_distrib) then have a: "[d2 = e2*b + q*b - x*b + (q - r)] (mod q)" by(simp add: cong_def) then have b:"[d2 = q + q + e2*b + q*b - x*b + (q - r)] (mod q)" proof- have "[e2*b + q*b - x*b + (q - r) = q + q + e2*b + q*b - x*b + (q - r)] (mod q)" by (smt assms Nat.add_diff_assoc add.commute cong_def less_imp_le_nat mod_add_self2 mult.commute nat_mult_le_cancel_disj semiring_normalization_rules(23)) then show ?thesis using cong_trans a by blast qed then have "[d2 = q + q + e2*b + q*b - (x*b) mod q + (q - r)] (mod q)" proof- have "[q + q + e2*b + q*b - (x*b) + (q - r) = q + q + e2*b + q*b - (x*b) mod q + (q - r)] (mod q)" proof(cases "b = 0") case True then show ?thesis by simp next case False have "q*b - (x*b) > 0" using False x by simp then have 1: "q + q + e2*b + q*b - (x*b) > 0" by linarith then have 2:"q + q + e2*b + q*b - (x*b) mod q > 0" by (simp add: q_gt_0 trans_less_add1) then show ?thesis using 1 2 minus_mod_cancel by simp qed then show ?thesis using cong_trans b by blast qed then have c: "[d2 = q + q + e2*b + q*b - (s2 + (q - r)) mod q + (q - r)] (mod q)" using xb by simp then have "[d2 = q + q + e2*b + q*b - (s2 + (q - r)) + (q - r)] (mod q)" proof- have "[q + q + e2*b + q*b - (s2 + (q - r)) mod q + (q - r) = q + q + e2*b + q*b - (s2 + (q - r)) + (q - r)] (mod q)" proof- have "q + q + e2*b + q*b - (s2 + (q - r)) mod q > 0" by (metis diff_is_0_eq gr0I le_less_trans mod_less_divisor not_add_less1 q_gt_0 semiring_normalization_rules(23) trans_less_add2) moreover have"q + q + e2*b + q*b - (s2 + (q - r)) > 0" using s2_le_q by simp ultimately show ?thesis using minus_mod_cancel cong_sym by blast qed then show ?thesis using cong_trans c by blast qed then have d: "[d2 = q + q + e2*b + q*b - s2 - (q - r) + (q - r)] (mod q)" by simp then have "[d2 = q + q + e2*b + q*b - s2 ] (mod q)" proof- have "q + q + e2*b + q*b - s2 - (q - r) > 0" using s2_le_q by simp then show ?thesis using d cong_trans by simp qed then have "[d2 = q + q + e2*b - s2] (mod q)" by (smt Nat.add_diff_assoc2 cong_def less_imp_le_nat mod_mult_self1 mult.commute s2_le_q semiring_normalization_rules(23) trans_less_add2) then have "[d2 = q + e2*b + q - s2] (mod q)" by(simp add: add.commute add.assoc) then have "[d2 = e2*b + q - s2] (mod q)" by (smt Nat.add_diff_assoc2 add.commute cong_def less_imp_le_nat mod_add_self2 s2_le_q trans_less_add2) then have "[d2 = e2*b + (q - s2)] (mod q)" by (simp add: less_imp_le_nat s2_le_q) then show ?thesis by(simp add: cong_def d2) qed lemma d2_P2: assumes x: "x < q" and y: "y < q" and r: "b < q" and b: "e2 < q" and e2: "r < q" shows "((y, b, ((e2 + q - x) mod q * b + q - r) mod q, e2), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) = ((y, b, (e2 * b + q - (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) mod q, e2), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q)" proof- have "((e2 + q - x) mod q * b + q - r) mod q = (e2 * b + q - (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) mod q" (is "?lhs = ?rhs") proof- have d2: "(((e2 + q - x) mod q)*b + (q - r)) mod q = (e2*b + (q - ((x*y + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q))) mod q" using assms d2 by blast have "?lhs = (((e2 + q - x) mod q)*b + (q - r)) mod q" using assms by simp also have "?rhs = (e2*b + (q - ((x*y + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q))) mod q" using assms by simp ultimately show ?thesis using assms d2 by metis qed then show ?thesis using assms by metis qed lemma s1: assumes s2: "s2 = (x*y + q - s1) mod q" and x: "x < q" and y: "y < q" and s1: "s1 < q" shows "s1 = (x*y + q - s2) mod q" proof- have s2_le_q:"s2 < q" using s2 q_gt_0 by simp have "[s2 = x*y + q - s1] (mod q)" by(simp add: cong_def s2) then have "[s2 = x*y + q - s1] (mod q)" using add.assoc by (simp add: less_imp_le_nat s1) then have s1_s2: "[s2 + s1 = x*y + q] (mod q)" by (metis (mono_tags, lifting) cong_def le_add2 le_add_diff_inverse2 le_trans mod_add_left_eq order.strict_implies_order s1) then have "[s1 = x*y + q - s2] (mod q)" proof- have "x*y + q - s2 > 0" using s2_le_q by simp then show ?thesis by (metis s1_s2 add_diff_cancel_left' cong_diff_nat cong_def le_add1 less_imp_le_nat zero_less_diff) qed then show ?thesis by(simp add: cong_def s1) qed lemma s1_P2: assumes x: "x < q" and y: "y < q" and "b < q" and "e2 < q" and "r < q" and "s1 < q" shows "((y, b, (e2 * b + q - (x * y + q - r) mod q) mod q, e2), r, (x * y + q - r) mod q) = ((y, b, (e2 * b + q - (x * y + q - r) mod q) mod q, e2), (x * y + q - (x * y + q - r) mod q) mod q, (x * y + q - r) mod q)" proof- have "s1 = (x*y + q - ((x*y + q - s1) mod q)) mod q" using assms secure_mult.s1 secure_mult_axioms by blast then show ?thesis using assms s1 by blast qed theorem P2_security: assumes "x < q" "y < q" shows "sim_non_det_def.inf_theoretic_P2 x y" including monad_normalisation proof- have "((funct x y) \ (\ (s1',s2'). (sim_non_det_def.Ideal2 y x s2'))) = R2 x y" proof- have "R2 x y = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let d2 = ((a*b + (q - r)) mod q); let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - r)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((y, c2, d2, e2), s1, s2)}" by(simp add: R2_def TI_def Let_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let e2 = (x + c1) mod q; let d2 = ((((e2 + q - x) mod q)*b + (q - r)) mod q); let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), (s1, (x*y + (q - s1)) mod q))}" by(simp add: Let_def s1_s2_P2 assms c1_P2 cong: bind_spmf_cong_simp) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let d1 = r; let c2 = b; e2 \ map_spmf (\ c1. (x + c1) mod q) (sample_uniform q); let d2 = ((((e2 + q - x) mod q)*b + (q - r)) mod q); let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let d1 = r; let c2 = b; e2 \ sample_uniform q; let d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q; let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_plus_one_time_pad) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; e2 \ sample_uniform q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let s2 = (x*y + (q - s1)) mod q; let d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; e2 \ sample_uniform q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: d2_P2 assms Let_def cong: bind_spmf_cong_simp) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ map_spmf (\ r. (x*((y + (q - b)) mod q) + (q - r)) mod q) (sample_uniform q); let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ sample_uniform q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: samp_uni_minus_one_time_pad) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ sample_uniform q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), (x*y + (q - s2)) mod q, s2)}" by(simp add: s1_P2 assms Let_def cong: bind_spmf_cong_simp) ultimately show ?thesis by(simp add: funct_def Let_def sim_non_det_def.Ideal2_def Out2_def S2_def R2_def) qed then show ?thesis by(simp add: sim_non_det_def.inf_theoretic_P2_def) qed lemma s1_s2_P1: assumes "x < q" "xa < q" "xb < q" "xc < q" "y < q" shows "((x, xa, xb, (y + q - xc) mod q), (x * ((y + q - xc) mod q) + q - xb) mod q, ((x + xa) mod q * xc + q - (xa * xc + q - xb) mod q) mod q) = ((x, xa, xb, (y + q - xc) mod q), (x * ((y + q - xc) mod q) + q - xb) mod q, (x * y + q - (x * ((y + q - xc) mod q) + q - xb) mod q) mod q)" using assms s1_s2 by metis lemma mod_minus: assumes "a - b > 0" and "c - d > 0" shows "(a - b + (c - d mod q)) mod q = (a - b + (c - d)) mod q" - using assms + using assms by (metis cong_def minus_mod mod_add_right_eq zero_less_diff) lemma r: assumes e1: "e1 = (y + (q - b)) mod q" and s1: "s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q" and b: "b < q" and x: "x < q" and y: "y < q" and r: "r < q" shows "r = (x*e1 + (q - s1)) mod q" (is "?lhs = ?rhs") proof- have "s1 = (x*((y + (q - b))) + (q - r)) mod q" using s1 b by (metis mod_add_left_eq mod_mult_right_eq) then have s1_dist: "s1 = (x*y + x*(q - b) + (q - r)) mod q" by(metis distrib_left) then have "?rhs = (x*((y + (q - b)) mod q) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" using e1 by simp then have "?rhs = (x*((y + (q - b))) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" by (metis mod_add_left_eq mod_mult_right_eq) then have "?rhs = (x*y + x*(q - b) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" by(metis distrib_left) then have a: "?rhs = (x*y + x*q - x*b + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" using distrib_left b diff_mult_distrib2 by auto then have b: "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" proof - have "(x*y + x*q - x*b + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" proof - have f1: "\n na nb nc nd. (n::nat) mod na \ nb mod na \ nc mod na \ nd mod na \ (n + nc) mod na = (nb + nd) mod na" by (meson mod_add_cong) then have "(q - (x * y + x*(q - b) + (q - r)) mod q) mod q = (q * q + q * q + q - (x * y + x*(q - b) + (q - r)) mod q) mod q" by (metis Nat.diff_add_assoc mod_le_divisor q_gt_0 mod_mult_self4) then show ?thesis using f1 by blast qed then show ?thesis using a by simp qed then have "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)))) mod q" proof- have "(x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b)+ (q - r)))) mod q" proof(cases "x = 0") case True then show ?thesis by (metis (no_types, lifting) assms(2) assms(4) True Nat.add_diff_assoc add.left_neutral cong_def diff_le_self minus_mod mult_is_0 not_gr_zero zero_eq_add_iff_both_eq_0 zero_less_diff) next case False have qb: "q - b \ q" using b by simp then have qb':"x*(q - b) < q*q" using x by (metis mult_less_le_imp_less nat_0_less_mult_iff nat_less_le neq0_conv) have a: "x*y + x*(q - b) > 0" using False assms by simp have 1: "q*q > x*y" using False by (simp add: mult_strict_mono q_gt_0 x y) have 2: "q*q > x*q" using False by (simp add: mult_strict_mono q_gt_0 x y) have b: "(q*q + q*q + q - (x*y + x*(q - b) + (q - r))) > 0" using 1 qb' by simp then show ?thesis using a b mod_minus[of "x*y + x*q" "x*b" "q*q + q*q + q" "x*y + x*(q - b) + (q - r)"] by (smt add.left_neutral cong_def gr0I minus_mod zero_less_diff) qed then show ?thesis using b by simp qed then have d: "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - x*y - x*(q - b) - (q - r))) mod q" by simp then have e: "?rhs = (x*y + x*q - x*b + q*q + q*q + q - x*y - x*(q - b) - (q - r)) mod q" proof(cases "x = 0") case True then show ?thesis using True d by simp next case False have qb: "q - b \ q" using b by simp then have qb':"x*(q - b) < q*q" using x by (metis mult_less_le_imp_less nat_0_less_mult_iff nat_less_le neq0_conv) have a: "x*y + x*(q - b) > 0" using False assms by simp have 1: "q*q > x*y" using False by (simp add: mult_strict_mono q_gt_0 x y) have 2: "q*q > x*q" using False by (simp add: mult_strict_mono q_gt_0 x y) have b: "q*q + q*q + q - x*y - x*(q - b) - (q - r) > 0" using 1 qb' by simp then show ?thesis using b d by (smt Nat.add_diff_assoc add.assoc diff_diff_left less_imp_le_nat zero_less_diff) qed then have "?rhs = (x*q - x*b + q*q + q*q + q - x*(q - b) - (q - r)) mod q" proof- have "(x*y + x*q - x*b + q*q + q*q + q - x*y - x*(q - b) - (q - r)) mod q = (x*q - x*b + q*q + q*q + q - x*(q - b) - (q - r)) mod q" proof - have 1: "q + n - b = q - b + n" for n by (simp add: assms(3) less_imp_le) have 2: "(c::nat) * b + (c * a + n) = c * (b + a) + n" for n a b c by (simp add: distrib_left) have "(c::nat) + (b + a) - (n + a) = c + b - n" for n a b c by auto then have "(q + (q * q + (q * q + x * (q + y - b))) - (q - r + x * (y + (q - b)))) mod q = (q + (q * q + (q * q + x * (q - b))) - (q - r + x * (q - b))) mod q" by (metis (no_types) add.commute 1 2) then show ?thesis by (simp add: add.commute diff_mult_distrib2 distrib_left) qed then show ?thesis using e by simp qed then have "?rhs = (x*(q - b) + q*q + q*q + q - x*(q - b) - (q - r)) mod q" by(metis diff_mult_distrib2) then have "?rhs = (q*q + q*q + q - (q - r)) mod q" using assms(6) by simp then have "?rhs = (q*q + q*q + q - q + r) mod q" using assms(6) by(simp add: Nat.diff_add_assoc2 less_imp_le) then have "?rhs = (q*q + q*q + r) mod q" by simp then have "?rhs = r mod q" by (metis add.commute distrib_right mod_mult_self1) then show ?thesis using assms(6) by simp qed lemma r_P2: assumes b: "b < q" and x: "x < q" and y: "y < q" and r: "r < q" shows "((x, a, r, (y + q - b) mod q), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) = ((x, a, (x * ((y + q - b) mod q) + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q, (y + q - b) mod q), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q)" proof- have "(x * ((y + q - b) mod q) + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q = r" (is "?lhs = ?rhs") proof- have 1:"r = (x*((y + (q - b)) mod q) + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q" using assms secure_mult.r secure_mult_axioms by blast also have "?rhs = (x*((y + (q - b)) mod q) + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q" using assms 1 by blast ultimately show ?thesis using assms 1 by simp qed then show ?thesis using assms by simp qed theorem P1_security: assumes "x < q" "y < q" shows "sim_non_det_def.inf_theoretic_P1 x y" including monad_normalisation proof- have "(funct x y) \ (\ (s1',s2'). (sim_non_det_def.Ideal1 x y s1')) = R1 x y" proof- have "R1 x y = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let d2 = ((a*b + (q - r)) mod q); let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((x, c1, d1, e1), s1, s2)}" by(simp add: R1_def TI_def Let_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: Let_def assms s1_s2_P1 r_P2 cong: bind_spmf_cong_simp) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; s1 \ map_spmf (\ r. (x*((y + (q - b)) mod q) + (q - r)) mod q) (sample_uniform q); let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf Let_def o_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_minus_one_time_pad) also have "... = do { a :: nat \ sample_uniform q; let c1 = a; e1 \ map_spmf (\b. (y + (q - b)) mod q) (sample_uniform q); s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf Let_def o_def) also have "... = do { a :: nat \ sample_uniform q; let c1 = a; e1 \ sample_uniform q; s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_minus_one_time_pad) ultimately show ?thesis by(simp add: funct_def sim_non_det_def.Ideal1_def Let_def R1_def TI_def Out1_def S1_def) qed thus ?thesis by(simp add: sim_non_det_def.inf_theoretic_P1_def) qed end locale secure_mult_asymp = fixes q :: "nat \ nat" assumes "\ n. secure_mult (q n)" begin sublocale secure_mult "q n" for n using secure_mult_asymp_axioms secure_mult_asymp_def by blast theorem P1_secure: assumes "x < q n" "y < q n" shows "sim_non_det_def.inf_theoretic_P1 n x y" by (metis P1_security assms) theorem P2_secure: assumes "x < q n" "y < q n" shows "sim_non_det_def.inf_theoretic_P2 n x y" by (metis P2_security assms) end end \ No newline at end of file diff --git a/thys/Neumann_Morgenstern_Utility/PMF_Composition.thy b/thys/Neumann_Morgenstern_Utility/PMF_Composition.thy --- a/thys/Neumann_Morgenstern_Utility/PMF_Composition.thy +++ b/thys/Neumann_Morgenstern_Utility/PMF_Composition.thy @@ -1,448 +1,448 @@ (* License: LGPL *) (* Author: Julian Parsert *) theory PMF_Composition imports "HOL-Probability.Probability" begin section \ Composition of Probability Mass functions \ definition mix_pmf :: "real \ 'a pmf \ 'a pmf \ 'a pmf" where "mix_pmf \ p q = (bernoulli_pmf \)\ (\X. if X then p else q)" lemma pmf_mix: "a \ {0..1} \ pmf (mix_pmf a p q) x = a * pmf p x + (1 - a) * pmf q x" by (simp add: mix_pmf_def pmf_bind) lemma pmf_mix_deeper: "a \ {0..1} \ pmf (mix_pmf a p q) x = a * pmf p x + pmf q x - a * pmf q x" by (simp add: left_diff_distrib' pmf_mix) lemma bernoulli_pmf_0 [simp]: "bernoulli_pmf 0 = return_pmf False" by (intro pmf_eqI) (auto simp: bernoulli_pmf.rep_eq) lemma bernoulli_pmf_1 [simp]: "bernoulli_pmf 1 = return_pmf True" by (intro pmf_eqI) (auto simp: bernoulli_pmf.rep_eq) lemma pmf_mix_0 [simp]: "mix_pmf 0 p q = q" by (simp add: mix_pmf_def bind_return_pmf) lemma pmf_mix_1 [simp]: "mix_pmf 1 p q = p" by (simp add: mix_pmf_def bind_return_pmf) lemma set_pmf_mix: "a \ {0<..<1} \ set_pmf (mix_pmf a p q) = set_pmf p \ set_pmf q" by (auto simp add: mix_pmf_def split: if_splits) lemma set_pmf_mix_eq: "a \ {0..1} \ mix_pmf a p p = p" by (simp add: mix_pmf_def) lemma pmf_equiv_intro[intro]: assumes "\e. e \ set_pmf p \ pmf p e = pmf q e" assumes "\e. e \ set_pmf q \ pmf q e = pmf p e" shows "p = q" by (metis assms(2) less_irrefl pmf_neq_exists_less pmf_not_neg set_pmf_iff) lemma pmf_equiv_intro1[intro]: assumes "\e. e \ set_pmf p \ pmf p e = pmf q e" shows "p = q" by (standard, auto simp: assms, metis assms set_pmf_iff assms linorder_not_le order_refl pmf_neq_exists_less pmf_not_neg set_pmf_iff) lemma pmf_inverse_switch_eqals: assumes "a \ {0..1}" shows "mix_pmf a p q = mix_pmf (1-a) q p" proof - have fst: "\x \ set_pmf p. pmf (mix_pmf a p q) x = pmf (mix_pmf (1-a) q p) x" proof fix x assume "x \ set_pmf p" have "pmf (mix_pmf a p q) x = a * pmf p x + (1 - a) * pmf q x" using pmf_mix[of a p q x] assms by blast also have "... = a * pmf p x + pmf q x - a * pmf q x" by (simp add: left_diff_distrib) from pmf_mix[of "1-a" q p x] assms have "pmf (mix_pmf (1 - a) q p) x = (1 - a) * pmf q x + (1 - (1 - a)) * pmf p x" by auto then show "pmf (mix_pmf a p q) x = pmf (mix_pmf (1 - a) q p) x" using calculation by auto qed have "\x \ set_pmf q. pmf (mix_pmf a p q) x = pmf (mix_pmf (1-a) q p) x" proof fix x assume "x \ set_pmf q" have "pmf (mix_pmf a p q) x = a * pmf p x + (1 - a) * pmf q x" using pmf_mix[of a p q x] assms by blast also have "... = a * pmf p x + pmf q x - a * pmf q x" by (simp add: left_diff_distrib) from pmf_mix[of "1-a" q p x] assms show "pmf (mix_pmf a p q) x = pmf (mix_pmf (1 - a) q p) x" using calculation by auto qed then have "\x \ set_pmf (mix_pmf a p q). pmf (mix_pmf a p q) x = pmf (mix_pmf (1-a) q p) x" by (metis (no_types) fst add_0_left assms mult_eq_0_iff pmf_mix set_pmf_iff) thus ?thesis by (simp add: pmf_equiv_intro1) qed lemma mix_pmf_comp_left_div: assumes "\ \ {0..(1::real)}" and "\ \ {0..(1::real)}" assumes "\ > \" shows "pmf (mix_pmf (\/\) (mix_pmf \ p q) q) e = \ * pmf p e + pmf q e - \ * pmf q e" proof- let ?l = "(mix_pmf (\/\) (mix_pmf \ p q) q)" have fst: "pmf ?l e = (\/\) * pmf (mix_pmf \ p q) e + (1-\/\) * pmf q e" by (meson assms(1) assms(2) assms(3) atLeastAtMost_iff less_divide_eq_1 less_eq_real_def not_less pmf_mix zero_le_divide_iff) then have "pmf (mix_pmf \ p q) e = \ * pmf p e + (1 - \) * pmf q e" using pmf_mix[of "\" p q] assms(2) assms(3) assms(1) by blast have "pmf ?l e = (\/\) * (\ * pmf p e + (1 - \) * pmf q e) + (1-\/\) * pmf q e" using fst assms(1) pmf_mix by fastforce then have "pmf ?l e = ((\/\) *\ * pmf p e + (\/\) *(1 - \) * pmf q e) + (1-\/\) * pmf q e" using fst assms(1) by (metis mult.assoc ring_class.ring_distribs(1)) then have *: "pmf ?l e = (\ * pmf p e + (\/\) *(1 - \) * pmf q e) + (1-\/\) * pmf q e" using fst assms(1) assms(2) assms(3) by auto then have "pmf ?l e = (\ * pmf p e + ((\/\) - (\/\)*\) * pmf q e) + (1-\/\) * pmf q e" using fst assms(1) assms(2) assms(3) by (simp add: * diff_divide_distrib right_diff_distrib') then have "pmf ?l e = (\ * pmf p e + ((\/\) - \) * pmf q e) + (1-\/\) * pmf q e" using fst assms(1) assms(2) assms(3) by auto then have "pmf ?l e = (\ * pmf p e + (\/\) * pmf q e - \ * pmf q e) + 1* pmf q e-\/\* pmf q e" by (simp add: left_diff_distrib) thus ?thesis by linarith qed lemma mix_pmf_comp_with_dif_equiv: assumes "\ \ {0..(1::real)}" and "\ \ {0..(1::real)}" assumes "\ > \" shows "mix_pmf (\/\) (mix_pmf \ p q) q = mix_pmf \ p q" (is "?l = ?r") proof (rule pmf_equiv_intro1[symmetric]) fix e assume a: "e \ set_pmf ?r" have "e \ set_pmf ?l" using a pmf_mix_deeper by (metis assms(1) assms(2) assms(3) mix_pmf_comp_left_div pmf_eq_0_set_pmf) then have "pmf ?l e = \ * pmf p e - \ * pmf q e + pmf q e" using pmf_mix_deeper[of "\/\" p q e] mix_pmf_comp_left_div[of \ \ p q e] assms by auto then show "pmf (mix_pmf \ p q) e = pmf (mix_pmf (\ / \) (mix_pmf \ p q) q) e" by (metis (full_types) assms(1) assms(2) assms(3) mix_pmf_comp_left_div pmf_mix_deeper) qed lemma product_mix_pmf_prob_distrib: assumes "a \ {0..1}" and "b \ {0..1}" shows "mix_pmf a (mix_pmf b p q) q = mix_pmf (a*b) p q" proof - define \ where g: "\ = (a * b)" define l where l: "l = (mix_pmf b p q)" define r where r: "r = mix_pmf (a*b) p q" have y: "\ \ {0..1}" using assms(2) mult_le_one assms g by auto have alt_: "\e \ set_pmf l. pmf r e = \ * pmf p e + pmf q e - \ * pmf q e" proof fix e have "pmf r e = \ * pmf p e + (1-\) * pmf q e" using \\ \ {0..1}\ g pmf_mix r by fastforce moreover have "... = \ * pmf p e + 1 * pmf q e - \ * pmf q e" - by (simp add: linordered_field_class.sign_simps(37)) + by (simp add: algebra_simps) moreover have "... = pmf (mix_pmf \ p q) e" using calculation g r by auto moreover have "... = \ * pmf p e + pmf q e - \ * pmf q e" using calculation by auto ultimately show "pmf r e = \ * pmf p e + pmf q e - \ * pmf q e" by auto qed have "\e \ set_pmf r. pmf l e = b * pmf p e + pmf q e - b * pmf q e" using allI pmf_mix_deeper assms(2) l by fastforce have "mix_pmf a (mix_pmf b p q) q = mix_pmf (a * b) p q" proof (rule ccontr) assume neg:"\mix_pmf a (mix_pmf b p q) q = mix_pmf (a * b) p q" then have b: "b \ 0" by (metis (no_types) assms(1) mult_cancel_right2 pmf_mix_0 set_pmf_mix_eq) have f3: "b - (a * b) > 0 \ mix_pmf a (mix_pmf b p q) q = mix_pmf (a * b) p q" by (metis assms(2) diff_le_0_iff_le g mix_pmf_comp_with_dif_equiv mult_eq_0_iff nonzero_mult_div_cancel_right not_le order_refl y) thus False using b neg assms(1) assms(2) by auto qed then show ?thesis by auto qed lemma mix_pmf_subset_of_original: assumes "a \ {0..1}" shows "(set_pmf (mix_pmf a p q)) \ set_pmf p \ set_pmf q" proof - have "a \ {0<..<1} \ ?thesis" by (simp add: set_pmf_mix) moreover have "a = 1 \ ?thesis" by simp moreover have "a = 0 \ ?thesis" by simp ultimately show ?thesis using assms less_eq_real_def by auto qed lemma mix_pmf_preserves_finite_support: assumes "a \ {0..1}" assumes "finite (set_pmf p)" and "finite (set_pmf q)" shows "finite (set_pmf (mix_pmf a p q))" by (meson assms(1) assms(2) assms(3) finite_Un finite_subset mix_pmf_subset_of_original) lemma ex_certain_iff_singleton_support: shows "(\x. pmf p x = 1) \ card (set_pmf p) = 1" proof (rule iffI, goal_cases) case 1 show ?case proof (rule ccontr) assume neg: "\ card (set_pmf p) = 1" then have "card (set_pmf p) \ 1" by blast have "finite (set_pmf p)" by (metis "1" empty_iff finite.emptyI finite_insert insert_iff not_le pmf_le_1 pmf_neq_exists_less pmf_nonneg set_pmf_iff set_return_pmf) then have sumeq_1: "(\i \ set_pmf p. pmf p i) = 1" using sum_pmf_eq_1[of "set_pmf p" p] by auto have set_pmf_nemtpy: "set_pmf p \ {}" by (simp add: set_pmf_not_empty) then have g1: "card (set_pmf p) > 1" by (metis card_0_eq less_one nat_neq_iff neg sum.infinite sumeq_1 zero_neq_one) have "card (set_pmf p) > 1 \ (\i \ set_pmf p. pmf p i) > 1" proof assume "card (set_pmf p) > 1" have "\x y. pmf p x = 1 \ y \ x \ y \ set_pmf p" using set_pmf_nemtpy is_singletonI' is_singleton_altdef by (metis "1" neg) then show "(\i \ set_pmf p. pmf p i) > 1" by (metis AE_measure_pmf_iff UNIV_I empty_iff insert_iff measure_pmf.prob_eq_1 pmf.rep_eq sets_measure_pmf) qed then have "card (set_pmf p) < 1" using sumeq_1 neg by linarith then show False using g1 by linarith qed qed (metis card_1_singletonE less_numeral_extra(1) pmf.rep_eq subset_eq sum_pmf_eq_1[of "set_pmf p" p] card_gt_0_iff[of "set_pmf p"] measure_measure_pmf_finite[of "set_pmf p"]) text \ We thank Manuel Eberl for suggesting the following two lemmas. \ lemma mix_pmf_partition: fixes p :: "'a pmf" assumes "y \ set_pmf p" "set_pmf p - {y} \ {}" obtains a q where "a \ {0<..<1}" "set_pmf q = set_pmf p - {y}" "p = mix_pmf a q (return_pmf y)" proof - from assms obtain x where x: "x \ set_pmf p - {y}" by auto define a where "a = 1 - pmf p y" have a_n1:"a \ 1" by (simp add: a_def assms(1) pmf_eq_0_set_pmf) have "pmf p y \ 1" using ex_certain_iff_singleton_support by (metis (full_types) Diff_cancel assms(1) assms(2) card_1_singletonE singletonD) hence y: "pmf p y < 1" using pmf_le_1[of p y] unfolding a_def by linarith hence a: "a > 0" by (simp add: a_def) define q where "q = embed_pmf (\z. if z = y then 0 else pmf p z / a)" have q: "pmf q z = (if z = y then 0 else pmf p z / a)" for z unfolding q_def proof (rule pmf_embed_pmf) have "1 = (\\<^sup>+ x. ennreal (pmf p x) \count_space UNIV)" by (rule nn_integral_pmf_eq_1 [symmetric]) also have "\ = (\\<^sup>+ x. ennreal (pmf p x) * indicator {y} x + ennreal (pmf p x) * indicator (-{y}) x \count_space UNIV)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+ x. ennreal (pmf p x) * indicator {y} x \count_space UNIV) + (\\<^sup>+ x. ennreal (pmf p x) * indicator (-{y}) x \count_space UNIV)" by (subst nn_integral_add) auto also have "(\\<^sup>+ x. ennreal (pmf p x) * indicator {y} x \count_space UNIV) = pmf p y" by (subst nn_integral_indicator_finite) auto also have "ennreal (pmf p y) + (\\<^sup>+ x. ennreal (pmf p x) * indicator (-{y}) x \count_space UNIV) - ennreal (pmf p y) = (\\<^sup>+ x. ennreal (pmf p x) * indicator (-{y}) x \count_space UNIV)" by simp also have "1 - ennreal (pmf p y) = ennreal (1 - pmf p y)" by (subst ennreal_1 [symmetric], subst ennreal_minus) auto finally have eq: "(\\<^sup>+x\-{y}. ennreal (pmf p x)\count_space UNIV) = 1 - pmf p y" .. have "(\\<^sup>+ x. ennreal (if x = y then 0 else pmf p x / a) \count_space UNIV) = (\\<^sup>+ x. inverse a * (ennreal (pmf p x) * indicator (-{y}) x) \count_space UNIV)" using a by (intro nn_integral_cong) (auto simp: divide_simps ennreal_mult' [symmetric]) also have "\ = inverse a * (\\<^sup>+ x\-{y}. ennreal (pmf p x) \count_space UNIV)" using a by (subst nn_integral_cmult [symmetric]) (auto simp: ennreal_mult') also note eq also have "ennreal (inverse a) * ennreal (1 - pmf p y) = ennreal ((1 - pmf p y) / a)" using a by (subst ennreal_mult' [symmetric]) (auto simp: field_simps) also have "(1 - pmf p y) / a = 1" using y by (simp add: a_def) finally show "(\\<^sup>+ x. ennreal (if x = y then 0 else pmf p x / a) \count_space UNIV) = 1" by simp qed (insert a, auto) have "mix_pmf (1 - pmf p y) q (return_pmf y) = p" using y by (intro pmf_eqI) (auto simp: q pmf_mix pmf_le_1 a_def) moreover have "set_pmf q = set_pmf p - {y}" using y by (auto simp: q set_pmf_eq a_def) ultimately show ?thesis using that[of "1 - pmf p y" q] y assms by (auto simp: set_pmf_eq) qed lemma pmf_mix_induct [consumes 2, case_names degenerate mix]: assumes "finite A" "set_pmf p \ A" assumes degenerate: "\x. x \ A \ P (return_pmf x)" assumes mix: "\p a y. set_pmf p \ A \ a \ {0<..<1} \ y \ A \ P p \ P (mix_pmf a p (return_pmf y))" shows "P p" proof - have "finite (set_pmf p)" "set_pmf p \ {}" "set_pmf p \ A" using assms(1,2) by (auto simp: set_pmf_not_empty dest: finite_subset) thus ?thesis proof (induction "set_pmf p" arbitrary: p rule: finite_ne_induct) case (singleton x p) hence "p = return_pmf x" using set_pmf_subset_singleton[of p x] by auto thus ?case using singleton by (auto intro: degenerate) next case (insert x B p) from insert.hyps have "x \ set_pmf p" "set_pmf p - {x} \ {}" by auto from mix_pmf_partition[OF this] obtain a q where decomp: "a \ {0<..<1}" "set_pmf q = set_pmf p - {x}" "p = mix_pmf a q (return_pmf x)" by blast have "P (mix_pmf a q (return_pmf x))" using insert.prems decomp(1) insert.hyps by (intro mix insert) (auto simp: decomp(2)) with decomp(3) show ?case by simp qed qed lemma pmf_mix_induct' [consumes 2, case_names degenerate mix]: assumes "finite A" "set_pmf p \ A" assumes degenerate: "\x. x \ A \ P (return_pmf x)" assumes mix: "\p q a. set_pmf p \ A \ set_pmf q \ A \ a \ {0<..<1} \ P p \ P q \ P (mix_pmf a p q)" shows "P p" using assms by (induct p rule: pmf_mix_induct)(auto)+ lemma finite_sum_distribute_mix_pmf: assumes "finite (set_pmf (mix_pmf a p q))" assumes "finite (set_pmf p)" assumes "finite (set_pmf q)" shows "(\i \ set_pmf (mix_pmf a p q). pmf (mix_pmf a p q) i) = (\i\set_pmf p. a * pmf p i) + (\i\set_pmf q. (1-a) * pmf q i)" proof - have fst: "(\i \ set_pmf (mix_pmf a p q). pmf (mix_pmf a p q) i) = 1" using sum_pmf_eq_1 assms by auto have "(\i\set_pmf p. a * pmf p i) = a * (\i\set_pmf p. pmf p i)" by (simp add: sum_distrib_left) also have "... = a * 1" using assms sum_pmf_eq_1 by (simp add: sum_pmf_eq_1) then show ?thesis by (metis fst add.assoc add_diff_cancel_left' add_uminus_conv_diff assms(3) mult.right_neutral order_refl sum_distrib_left sum_pmf_eq_1) qed lemma distribute_alpha_over_sum: shows "(\i\set_pmf T. a * pmf p i * f i) = a * (\i\set_pmf T. pmf p i * f i)" by (metis (mono_tags, lifting) semiring_normalization_rules(18) sum.cong sum_distrib_left) lemma sum_over_subset_pmf_support: assumes "finite T" assumes "set_pmf p \ T" shows "(\i\T. a * pmf p i * f i) = (\i\set_pmf p. a * pmf p i * f i)" proof - consider (eq) "set_pmf p = T" | (sub) "set_pmf p \ T" using assms by blast then show ?thesis proof (cases) next case sub define A where "A = T - (set_pmf p)" have "finite (set_pmf p)" using assms(1) assms(2) finite_subset by auto moreover have "finite A" using A_def assms(1) by blast moreover have "A \ set_pmf p = {}" using A_def assms(1) by blast ultimately have *: "(\i\T. a * pmf p i * f i) = (\i\set_pmf p. a * pmf p i * f i) + (\i\A. a * pmf p i * f i)" using sum.union_disjoint by (metis (no_types) A_def Un_Diff_cancel2 Un_absorb2 assms(2) inf.commute inf_sup_aci(5) sum.union_disjoint) have "\e \ A. pmf p e = 0" by (simp add: A_def pmf_eq_0_set_pmf) hence "(\i\A. a * pmf p i * f i) = 0" by simp then show ?thesis by (simp add: "*") qed (auto) qed lemma expected_value_mix_pmf_distrib: assumes "finite (set_pmf p)" and "finite (set_pmf q)" assumes "a \ {0<..<1}" shows "measure_pmf.expectation (mix_pmf a p q) f = a * measure_pmf.expectation p f + (1-a) * measure_pmf.expectation q f" proof - have fn: "finite (set_pmf (mix_pmf a p q))" using mix_pmf_preserves_finite_support assms less_eq_real_def by auto have subsets: "set_pmf p \ set_pmf (mix_pmf a p q)" "set_pmf q \ set_pmf (mix_pmf a p q)" using assms assms set_pmf_mix by(fastforce)+ have *: "(\i \ set_pmf (mix_pmf a p q). a * pmf p i * f i) = a * (\i \ set_pmf (mix_pmf a p q). pmf p i * f i)" by (metis (mono_tags, lifting) mult.assoc sum.cong sum_distrib_left) have **: "(\i \ set_pmf (mix_pmf a p q). (1-a) * pmf q i * f i) = (1-a) * (\i \ set_pmf (mix_pmf a p q). pmf q i * f i)" using distribute_alpha_over_sum[of "(1 - a)" q f "(mix_pmf a p q)"] by auto have ***: "measure_pmf.expectation (mix_pmf a p q) f = (\i \ set_pmf (mix_pmf a p q). pmf (mix_pmf a p q) i * f i)" by (metis fn pmf_integral_code_unfold pmf_integral_pmf_set_integral pmf_set_integral_code_alt_finite) also have g: "... = (\i \ set_pmf (mix_pmf a p q). (a * pmf p i + (1-a) * pmf q i) * f i)" using pmf_mix[of a p q] assms(3) by auto also have ****: "... = (\i \ set_pmf (mix_pmf a p q). a * pmf p i * f i + (1-a) * pmf q i * f i)" by (simp add: mult.commute ring_class.ring_distribs(1)) also have f: "... = (\i \ set_pmf (mix_pmf a p q). a * pmf p i * f i) + (\i \ set_pmf (mix_pmf a p q). (1-a) * pmf q i * f i)" by (simp add: sum.distrib) also have "... = a * (\i \ set_pmf (mix_pmf a p q). pmf p i * f i) + (1-a) * (\i \ set_pmf (mix_pmf a p q). pmf q i * f i)" using * ** by simp also have h: "... = a * (\i \ set_pmf p. pmf p i * f i) + (1-a) * (\i \ set_pmf q. pmf q i * f i)" proof - have "(\i \ set_pmf (mix_pmf a p q). pmf p i * f i) = (\i \ set_pmf p. pmf p i * f i)" using subsets sum_over_subset_pmf_support[of "(mix_pmf a p q)" p 1 f] fn by auto moreover have "(\i \ set_pmf (mix_pmf a p q). pmf q i * f i) = (\i \ set_pmf q. pmf q i * f i)" using subsets sum_over_subset_pmf_support[of "(mix_pmf a p q)" q 1 f] fn by auto ultimately show ?thesis by (simp) qed finally show ?thesis proof - have "(\i\set_pmf q. pmf q i * f i) = measure_pmf.expectation q f" by (metis (full_types) assms(2) pmf_integral_code_unfold pmf_integral_pmf_set_integral pmf_set_integral_code_alt_finite) moreover have "(\i\set_pmf p. pmf p i * f i) = measure_pmf.expectation p f" by (metis (full_types) assms(1) pmf_integral_code_unfold pmf_integral_pmf_set_integral pmf_set_integral_code_alt_finite) ultimately show ?thesis by (simp add: "*" "**" "***" "****" f g h) qed qed lemma expected_value_mix_pmf: assumes "finite (set_pmf p)" and "finite (set_pmf q)" assumes "a \ {0..1}" shows "measure_pmf.expectation (mix_pmf a p q) f = a * measure_pmf.expectation p f + (1-a) * measure_pmf.expectation q f" proof - consider (0) "a = 0" | (b) "a \ {0<..<1}" | (1) "a = 1" using assms(3) less_eq_real_def by auto then show ?thesis proof (cases) case 0 have "(mix_pmf a p q) = q" using 0 pmf_mix_0 by blast have "measure_pmf.expectation q f = (1-a) * measure_pmf.expectation q f" by (simp add: 0) show ?thesis using 0 by auto next case b show ?thesis using assms(1) assms(2) b expected_value_mix_pmf_distrib by blast next case 1 have "(mix_pmf a p q) = p" using 1 pmf_mix_0 by simp then show ?thesis by (simp add: "1") qed qed end diff --git a/thys/Ordinary_Differential_Equations/IVP/Cones.thy b/thys/Ordinary_Differential_Equations/IVP/Cones.thy --- a/thys/Ordinary_Differential_Equations/IVP/Cones.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Cones.thy @@ -1,694 +1,694 @@ theory Cones imports "HOL-Analysis.Analysis" Triangle.Triangle "../ODE_Auxiliarities" begin lemma arcsin_eq_zero_iff[simp]: "-1 \ x \ x \ 1 \ arcsin x = 0 \ x = 0" using sin_arcsin by fastforce definition conemem :: "'a::real_vector \ 'a \ real \ 'a" where "conemem u v t = cos t *\<^sub>R u + sin t *\<^sub>R v" definition "conesegment u v = conemem u v ` {0.. pi / 2}" lemma bounded_linear_image_conemem: assumes "bounded_linear F" shows "F (conemem u v t) = conemem (F u) (F v) t" proof - from assms interpret bounded_linear F . show ?thesis by (auto simp: conemem_def[abs_def] cone_hull_expl closed_segment_def add scaleR) qed lemma bounded_linear_image_conesegment: assumes "bounded_linear F" shows "F ` conesegment u v = conesegment (F u) (F v)" proof - from assms interpret bounded_linear F . show ?thesis apply (auto simp: conesegment_def conemem_def[abs_def] cone_hull_expl closed_segment_def add scaleR) apply (auto simp: add[symmetric] scaleR[symmetric]) done qed (* This is vangle in $AFP/Triangles/Angles *) lemma discriminant: "a * x\<^sup>2 + b * x + c = (0::real) \ 0 \ b\<^sup>2 - 4 * a * c" by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))") lemma quadratic_eq_factoring: assumes D: "D = b\<^sup>2 - 4 * a * c" assumes nn: "0 \ D" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a: "a \ 0" shows "a * x\<^sup>2 + b * x + c = a * (x - x\<^sub>1) * (x - x\<^sub>2)" using nn by (simp add: D x1 x2) (simp add: assms algebra_simps power2_eq_square power3_eq_cube divide_simps) lemma quadratic_eq_zeroes_iff: assumes D: "D = b\<^sup>2 - 4 * a * c" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a: "a \ 0" shows "a * x\<^sup>2 + b * x + c = 0 \ (D \ 0 \ (x = x\<^sub>1 \ x = x\<^sub>2))" (is "?z \ _") using quadratic_eq_factoring[OF D _ x1 x2 a, of x] discriminant[of a x b c] a by (auto simp: D) lemma quadratic_ex_zero_iff: "(\x. a * x\<^sup>2 + b * x + c = 0) \ (a \ 0 \ b\<^sup>2 - 4 * a * c \ 0 \ a = 0 \ (b = 0 \ c = 0))" for a b c::real apply (cases "a = 0") subgoal by (auto simp: intro: exI[where x="- c / b"]) subgoal by (subst quadratic_eq_zeroes_iff[OF refl refl refl]) auto done lemma Cauchy_Schwarz_eq_iff: shows "(inner x y)\<^sup>2 = inner x x * inner y y \ ((\k. x = k *\<^sub>R y) \ y = 0)" proof safe assume eq: "(x \ y)\<^sup>2 = x \ x * (y \ y)" and "y \ 0" define f where "f \ \l. inner (x - l *\<^sub>R y) (x - l *\<^sub>R y)" have f_quadratic: "f l = inner y y * l\<^sup>2 + - 2 * inner x y * l + inner x x" for l by (auto simp: f_def algebra_simps power2_eq_square inner_commute) have "\l. f l = 0" unfolding f_quadratic quadratic_ex_zero_iff using \y \ 0\ by (auto simp: eq) then show "(\k. x = k *\<^sub>R y)" by (auto simp: f_def) qed (auto simp: power2_eq_square) lemma Cauchy_Schwarz_strict_ineq: "(inner x y)\<^sup>2 < inner x x * inner y y" if "y \ 0" "\k. x \ k *\<^sub>R y" apply (rule neq_le_trans) subgoal using that unfolding Cauchy_Schwarz_eq_iff by auto subgoal by (rule Cauchy_Schwarz_ineq) done lemma Cauchy_Schwarz_eq2_iff: "\inner x y\ = norm x * norm y \ ((\k. x = k *\<^sub>R y) \ y = 0)" using Cauchy_Schwarz_eq_iff[of x y] by (subst power_eq_iff_eq_base[symmetric, where n = 2]) (simp_all add: dot_square_norm power_mult_distrib) lemma Cauchy_Schwarz_strict_ineq2: "\inner x y\ < norm x * norm y" if "y \ 0" "\k. x \ k *\<^sub>R y" apply (rule neq_le_trans) subgoal using that unfolding Cauchy_Schwarz_eq2_iff by auto subgoal by (rule Cauchy_Schwarz_ineq2) done lemma gt_minus_one_absI: "abs k < 1 \ - 1 < k" for k::real by auto lemma gt_one_absI: "abs k < 1 \ k < 1" for k::real by auto lemma abs_impossible: "\y1\ < x1 \ \y2\ < x2 \ x1 * x2 + y1 * y2 \ 0" for x1 x2::real proof goal_cases case 1 have "- y1 * y2 \ abs y1 * abs y2" by (metis abs_ge_minus_self abs_mult mult.commute mult_minus_right) also have "\ < x1 * x2" apply (rule mult_strict_mono) using 1 by auto finally show ?case by auto qed lemma vangle_eq_arctan_minus:\ \TODO: generalize?!\ assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes xy1: "\y1\ < x1" assumes xy2: "\y2\ < x2" assumes less: "y2 / x2 > y1 / x1" shows "vangle (x1 *\<^sub>R i + y1 *\<^sub>R j) (x2 *\<^sub>R i + y2 *\<^sub>R j) = arctan (y2 / x2) - arctan (y1 / x1)" (is "vangle ?u ?v = _") proof - from assms have less2: "x2 * y1 - x1 * y2 < 0" by (auto simp: divide_simps abs_real_def algebra_simps split: if_splits) have norm_eucl: "norm (x *\<^sub>R i + y *\<^sub>R j) = sqrt ((norm x)\<^sup>2 + (norm y)\<^sup>2)" for x y apply (subst norm_eq_sqrt_inner) using ij ij_neq by (auto simp: inner_simps inner_Basis power2_eq_square) have nonzeroes: "x1 *\<^sub>R i + y1 *\<^sub>R j \ 0" "x2 *\<^sub>R i + y2 *\<^sub>R j \ 0" apply (auto simp: euclidean_eq_iff[where 'a='a] inner_simps intro!: bexI[where x=i]) using assms by (auto simp: inner_Basis) have indep: "x1 *\<^sub>R i + y1 *\<^sub>R j \ k *\<^sub>R (x2 *\<^sub>R i + y2 *\<^sub>R j)" for k proof assume "x1 *\<^sub>R i + y1 *\<^sub>R j = k *\<^sub>R (x2 *\<^sub>R i + y2 *\<^sub>R j)" then have "x1 / x2 = k" "y1 = k * y2" using ij ij_neq xy1 xy2 apply (auto simp: abs_real_def divide_simps algebra_simps euclidean_eq_iff[where 'a='a] inner_simps split: if_splits) by (auto simp: inner_Basis split: if_splits) then have "y1 = x1 / x2 * y2" by simp with less show False using xy1 by (auto split: if_splits) qed have "((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - ((x1 *\<^sub>R i + y1 *\<^sub>R j) \ (x2 *\<^sub>R i + y2 *\<^sub>R j))\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2)))) = ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - (x1 * x2 + y1 * y2)\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2))))" using ij_neq ij by (auto simp: algebra_simps divide_simps inner_simps inner_Basis) also have "\ = (x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) - (x1 * x2 + y1 * y2)\<^sup>2" unfolding right_diff_distrib by simp also have "\ = (x2 * y1 - x1 * y2)^2" by (auto simp: algebra_simps power2_eq_square) also have "sqrt \ = \x2 * y1 - x1 * y2\" by simp also have "\ = x1 * y2 - x2 * y1" using less2 by (simp add: abs_real_def) finally have sqrt_eq: "sqrt ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - ((x1 *\<^sub>R i + y1 *\<^sub>R j) \ (x2 *\<^sub>R i + y2 *\<^sub>R j))\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2)))) = x1 * y2 - x2 * y1" . show ?thesis using ij xy1 xy2 unfolding vangle_def apply (subst arccos_arctan) subgoal apply (rule gt_minus_one_absI) apply (simp add: ) apply (subst pos_divide_less_eq) subgoal apply (rule mult_pos_pos) using nonzeroes by auto subgoal apply simp apply (rule Cauchy_Schwarz_strict_ineq2) using nonzeroes indep by auto done subgoal apply (rule gt_one_absI) apply (simp add: ) apply (subst pos_divide_less_eq) subgoal apply (rule mult_pos_pos) using nonzeroes by auto subgoal apply simp apply (rule Cauchy_Schwarz_strict_ineq2) using nonzeroes indep by auto done subgoal apply (auto simp: nonzeroes) apply (subst (3) diff_conv_add_uminus) apply (subst arctan_minus[symmetric]) apply (subst arctan_add) apply force apply force apply (subst arctan_inverse[symmetric]) subgoal apply (rule divide_pos_pos) subgoal apply (auto simp add: inner_simps inner_Basis algebra_simps ) apply (thin_tac "_ \ Basis")+ apply (thin_tac "j = i") apply (sos "((((A<0 * (A<1 * (A<2 * A<3))) * R<1) + ((A<=0 * (A<0 * (A<2 * R<1))) * (R<1 * [1]^2))))") apply (thin_tac "_ \ Basis")+ apply (thin_tac "j \ i") by (sos "((((A<0 * (A<1 * (A<2 * A<3))) * R<1) + (((A<2 * (A<3 * R<1)) * (R<1/3 * [y1]^2)) + (((A<1 * (A<3 * R<1)) * ((R<1/12 * [x2 + y1]^2) + (R<1/12 * [x1 + y2]^2))) + (((A<1 * (A<2 * R<1)) * (R<1/12 * [~1*x1 + x2 + y1 + y2]^2)) + (((A<0 * (A<3 * R<1)) * (R<1/12 * [~1*x1 + x2 + ~1*y1 + ~1*y2]^2)) + (((A<0 * (A<2 * R<1)) * ((R<1/12 * [x2 + ~1*y1]^2) + (R<1/12 * [~1*x1 + y2]^2))) + (((A<0 * (A<1 * R<1)) * (R<1/3 * [y2]^2)) + ((A<=0 * R<1) * (R<1/3 * [x1 + x2]^2))))))))))") subgoal apply (intro mult_pos_pos) using nonzeroes indep apply auto apply (rule gt_one_absI) apply (simp add: power_divide power_mult_distrib power2_norm_eq_inner) apply (rule Cauchy_Schwarz_strict_ineq) apply auto done done subgoal apply (rule arg_cong[where f=arctan]) using nonzeroes ij_neq apply (auto simp: norm_eucl) apply (subst real_sqrt_mult[symmetric]) apply (subst real_sqrt_mult[symmetric]) apply (subst real_sqrt_mult[symmetric]) apply (subst power_divide) apply (subst real_sqrt_pow2) apply simp apply (subst nonzero_divide_eq_eq) subgoal apply (auto simp: algebra_simps inner_simps inner_Basis) by (auto simp: algebra_simps divide_simps abs_real_def abs_impossible) apply (subst sqrt_eq) apply (auto simp: algebra_simps inner_simps inner_Basis) apply (auto simp: algebra_simps divide_simps abs_real_def abs_impossible) by (auto split: if_splits) done done qed lemma vangle_le_pi2: "0 \ u \ v \ vangle u v \ pi/2" unfolding vangle_def atLeastAtMost_iff apply (simp del: le_divide_eq_numeral1) apply (intro impI arccos_le_pi2 arccos_lbound) using Cauchy_Schwarz_ineq2[of u v] - by (auto simp: divide_simps sign_simps) + by (auto simp: algebra_simps) lemma inner_eq_vangle: "u \ v = cos (vangle u v) * (norm u * norm v)" by (simp add: cos_vangle) lemma vangle_scaleR_self: "vangle (k *\<^sub>R v) v = (if k = 0 \ v = 0 then pi / 2 else if k > 0 then 0 else pi)" "vangle v (k *\<^sub>R v) = (if k = 0 \ v = 0 then pi / 2 else if k > 0 then 0 else pi)" by (auto simp: vangle_def dot_square_norm power2_eq_square) lemma vangle_scaleR: "vangle (k *\<^sub>R v) w = vangle v w" "vangle w (k *\<^sub>R v) = vangle w v" if "k > 0" using that by (auto simp: vangle_def) lemma cos_vangle_eq_zero_iff_vangle: "cos (vangle u v) = 0 \ (u = 0 \ v = 0 \ u \ v = 0)" using Cauchy_Schwarz_ineq2[of u v] by (auto simp: vangle_def divide_simps sign_simps split: if_splits) lemma ortho_imp_angle_pi_half: "u \ v = 0 \ vangle u v = pi / 2" using orthogonal_iff_vangle[of u v] by (auto simp: orthogonal_def) lemma arccos_eq_zero_iff: "arccos x = 0 \ x = 1" if "-1 \ x" "x \ 1" using that apply auto using cos_arccos by fastforce lemma vangle_eq_zeroD: "vangle u v = 0 \ (\k. v = k *\<^sub>R u)" apply (auto simp: vangle_def split: if_splits) apply (subst (asm) arccos_eq_zero_iff) apply (auto simp: divide_simps mult_less_0_iff split: if_splits) apply (metis Real_Vector_Spaces.norm_minus_cancel inner_minus_left minus_le_iff norm_cauchy_schwarz) apply (metis norm_cauchy_schwarz) by (metis Cauchy_Schwarz_eq2_iff abs_of_pos inner_commute mult.commute mult_sign_intros(5) zero_less_norm_iff) lemma less_one_multI:\ \TODO: also in AA!\ fixes e x::real shows "e \ 1 \ 0 < x \ x < 1 \ e * x < 1" by (metis (erased, hide_lams) less_eq_real_def monoid_mult_class.mult.left_neutral mult_strict_mono zero_less_one) lemma conemem_expansion_estimate: fixes u v u' v'::"'a::euclidean_space" assumes "t \ {0 .. pi / 2}" assumes angle_pos: "0 < vangle u v" "vangle u v < pi / 2" assumes angle_le: "(vangle u' v') \ (vangle u v)" assumes "norm u = 1" "norm v = 1" shows "norm (conemem u' v' t) \ min (norm u') (norm v') * norm (conemem u v t)" proof - define e_pre where "e_pre = min (norm u') (norm v')" let ?w = "conemem u v" let ?w' = "conemem u' v'" have cos_angle_le: "cos (vangle u' v') \ cos (vangle u v)" using angle_pos vangle_bounds by (auto intro!: cos_monotone_0_pi_le angle_le) have e_pre_le: "e_pre\<^sup>2 \ norm u' * norm v'" by (auto simp: e_pre_def min_def power2_eq_square intro: mult_left_mono mult_right_mono) have lt: "0 < 1 + 2 * (u \ v) * sin t * cos t" proof - have "\u \ v\ < norm u * norm v" apply (rule Cauchy_Schwarz_strict_ineq2) using assms apply auto apply (subst (asm) vangle_scaleR_self)+ by (auto simp: split: if_splits) then have "abs (u \ v * sin (2 * t)) < 1" using assms apply (auto simp add: abs_mult) apply (subst mult.commute) apply (rule less_one_multI) apply (auto simp add: abs_mult inner_eq_vangle ) by (auto simp: cos_vangle_eq_zero_iff_vangle dest!: ortho_imp_angle_pi_half) then show ?thesis by (subst mult.assoc sin_times_cos)+ auto qed have le: "0 \ 1 + 2 * (u \ v) * sin t * cos t" proof - have "\u \ v\ \ norm u * norm v" by (rule Cauchy_Schwarz_ineq2) then have "abs (u \ v * sin (2 * t)) \ 1" by (auto simp add: abs_mult assms intro!: mult_le_one) then show ?thesis by (subst mult.assoc sin_times_cos)+ auto qed have "(norm (?w t))\<^sup>2 = (cos t)\<^sup>2 *\<^sub>R (norm u)\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v)\<^sup>2 + 2 * (u \ v) * sin t * cos t" by (auto simp: conemem_def algebra_simps power2_norm_eq_inner) (auto simp: power2_eq_square inner_commute) also have "\ = 1 + 2 * (u \ v) * sin t * cos t" by (auto simp: sin_squared_eq algebra_simps assms) finally have "(norm (conemem u v t))\<^sup>2 = 1 + 2 * (u \ v) * sin t * cos t" by simp moreover have "(norm (?w' t))\<^sup>2 = (cos t)\<^sup>2 *\<^sub>R (norm u')\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v')\<^sup>2 + 2 * (u' \ v') * sin t * cos t" by (auto simp: conemem_def algebra_simps power2_norm_eq_inner) (auto simp: power2_eq_square inner_commute) ultimately have "(norm (?w' t) / norm (?w t))\<^sup>2 = ((cos t)\<^sup>2 *\<^sub>R (norm u')\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v')\<^sup>2 + 2 * (u' \ v') * sin t * cos t) / (1 + 2 * (u \ v) * sin t * cos t)" (is "_ = (?a + ?b) / ?c") by (auto simp: divide_inverse power_mult_distrib) (auto simp: inverse_eq_divide power2_eq_square) also have "\ \ (e_pre\<^sup>2 + ?b) / ?c" apply (rule divide_right_mono) apply (rule add_right_mono) subgoal using assms e_pre_def apply (auto simp: min_def) subgoal by (auto simp: algebra_simps cos_squared_eq intro!: mult_right_mono power_mono) subgoal by (auto simp: algebra_simps sin_squared_eq intro!: mult_right_mono power_mono) done subgoal by (rule le) done also (xtrans) have inner_nonneg: "u' \ v' \ 0" using angle_le(1) angle_pos vangle_bounds[of u' v'] by (auto simp: inner_eq_vangle intro!: mult_nonneg_nonneg cos_ge_zero) from vangle_bounds[of u' v'] vangle_le_pi2[OF this] have u'v'e_pre: "u' \ v' \ cos (vangle u' v') * e_pre\<^sup>2" apply (subst inner_eq_vangle) apply (rule mult_left_mono) apply (rule e_pre_le) apply (rule cos_ge_zero) by auto have "(e_pre\<^sup>2 + ?b) / ?c \ (e_pre\<^sup>2 + 2 * (cos (vangle u' v') * e_pre\<^sup>2) * sin t * cos t) / ?c" (is "_ \ ?ddd") apply (intro divide_right_mono add_left_mono mult_right_mono mult_left_mono u'v'e_pre) using \t \ _\ by (auto intro!: mult_right_mono sin_ge_zero divide_right_mono le cos_ge_zero simp: sin_times_cos u'v'e_pre) also (xtrans) have "?ddd = e_pre\<^sup>2 * ((1 + 2 * cos (vangle u' v') * sin t * cos t) / ?c)" (is "_ = ?ddd") by (auto simp add: divide_simps algebra_simps) also (xtrans) have sc_ge_0: "0 \ sin t * cos t" using \t \ _\ by (auto simp: assms cos_angle_le intro!: mult_nonneg_nonneg sin_ge_zero cos_ge_zero) have "?ddd \ e_pre\<^sup>2" apply (subst mult_le_cancel_left1) apply (auto simp add: divide_simps split: if_splits) apply (rule mult_right_mono) using lt by (auto simp: assms inner_eq_vangle intro!: mult_right_mono sc_ge_0 cos_angle_le) finally (xtrans) have "(norm (conemem u' v' t))\<^sup>2 \ (e_pre * norm (conemem u v t))\<^sup>2" by (simp add: divide_simps power_mult_distrib split: if_splits) then show "norm (conemem u' v' t) \ e_pre * norm (conemem u v t)" using norm_imp_pos_and_ge power2_le_imp_le by blast qed lemma conemem_commute: "conemem a b t = conemem b a (pi / 2 - t)" if "0 \ t" "t \ pi / 2" using that by (auto simp: conemem_def cos_sin_eq algebra_simps) lemma conesegment_commute: "conesegment a b = conesegment b a" apply (auto simp: conesegment_def ) apply (subst conemem_commute) apply auto apply (subst conemem_commute) apply auto done definition "conefield u v = cone hull (conesegment u v)" lemma conefield_alt_def: "conefield u v = cone hull {u--v}" apply (auto simp: conesegment_def conefield_def cone_hull_expl in_segment) subgoal premises prems for c t proof - from prems have sc_pos: "sin t + cos t > 0" apply (cases "t = 0") subgoal by (rule add_nonneg_pos) auto subgoal by (auto intro!: add_pos_nonneg sin_gt_zero cos_ge_zero) done then have 1: "(sin t / (sin t + cos t) + cos t / (sin t + cos t)) = 1" by (auto simp: divide_simps) have "\c x. c > 0 \ 0 \ x \ x \ 1 \ c *\<^sub>R conemem u v t = (1 - x) *\<^sub>R u + x *\<^sub>R v" apply (auto simp: algebra_simps conemem_def) apply (rule exI[where x="1 / (sin t + cos t)"]) using prems by (auto intro!: exI[where x="(1 / (sin t + cos t) * sin t)"] sc_pos divide_nonneg_nonneg sin_ge_zero add_nonneg_nonneg cos_ge_zero simp: scaleR_add_left[symmetric] 1 divide_le_eq_1) then obtain d x where dx: "d > 0" "conemem u v t = (1 / d) *\<^sub>R ((1 - x) *\<^sub>R u + x *\<^sub>R v)" "0 \ x" "x \ 1" by (auto simp: eq_vector_fraction_iff) show ?thesis apply (rule exI[where x="c / d"]) using dx by (auto simp: intro!: divide_nonneg_nonneg prems ) qed subgoal premises prems for c t proof - let ?x = "arctan (t / (1 - t))" let ?s = "t / sin ?x" have *: "c *\<^sub>R ((1 - t) *\<^sub>R u + t *\<^sub>R v) = (c * ?s) *\<^sub>R (cos ?x *\<^sub>R u + sin ?x *\<^sub>R v)" if "0 < t" "t < 1" using that by (auto simp: scaleR_add_right sin_arctan cos_arctan divide_simps) show ?thesis apply (cases "t = 0") subgoal apply simp apply (rule exI[where x=c]) apply (rule exI[where x=u]) using prems by (auto simp: conemem_def[abs_def] intro!: image_eqI[where x=0]) subgoal apply (cases "t = 1") subgoal apply simp apply (rule exI[where x=c]) apply (rule exI[where x=v]) using prems by (auto simp: conemem_def[abs_def] intro!: image_eqI[where x="pi/2"]) subgoal apply (rule exI[where x="(c * ?s)"]) apply (rule exI[where x="(cos ?x *\<^sub>R u + sin ?x *\<^sub>R v)"]) using prems * arctan_ubound[of "t / (1 - t)"] apply (auto simp: conemem_def[abs_def] intro!: imageI) by (auto simp: scaleR_add_right sin_arctan) done done qed done lemma bounded_linear_image_cone_hull: assumes "bounded_linear F" shows "F ` (cone hull T) = cone hull (F ` T)" proof - from assms interpret bounded_linear F . show ?thesis apply (auto simp: conefield_def cone_hull_expl closed_segment_def add scaleR) apply (auto simp: ) apply (auto simp: add[symmetric] scaleR[symmetric]) done qed lemma bounded_linear_image_conefield: assumes "bounded_linear F" shows "F ` conefield u v = conefield (F u) (F v)" unfolding conefield_def using assms by (auto simp: bounded_linear_image_conesegment bounded_linear_image_cone_hull) lemma conefield_commute: "conefield x y = conefield y x" by (auto simp: conefield_def conesegment_commute) lemma convex_conefield: "convex (conefield x y)" by (auto simp: conefield_alt_def convex_cone_hull) lemma conefield_scaleRI: "v \ conefield (r *\<^sub>R x) y" if "v \ conefield x y" "r > 0" using that using \r > 0\ unfolding conefield_alt_def cone_hull_expl apply (auto simp: in_segment) proof goal_cases case (1 c u) let ?d = "c * (1 - u) / r + c * u" let ?t = "c * u / ?d" have "c * (1 - u) = ?d * (1 - ?t) * r" if "0 < u" using \0 < r\ that(1) 1(3,5) mult_pos_pos by (force simp: divide_simps ac_simps ring_distribs[symmetric]) then have eq1: "(c * (1 - u)) *\<^sub>R x = (?d * (1 - ?t) * r) *\<^sub>R x" if "0 < u" using that by simp have "c * u = ?d * ?t" if "u < 1" using \0 < r\ that(1) 1(3,4,5) mult_pos_pos apply (auto simp: divide_simps ac_simps ring_distribs[symmetric]) proof - assume "0 \ u" "0 < r" "1 - u + r * u = 0" "u < 1" then have False by (sos "((((A<0 * A<1) * R<1) + (([~1*r] * A=0) + ((A<=0 * R<1) * (R<1 * [r]^2)))))") then show "u = 0" by metis qed then have eq2: "(c * u) *\<^sub>R y = (?d * ?t) *\<^sub>R y" if "u < 1" using that by simp have *: "c *\<^sub>R ((1 - u) *\<^sub>R x + u *\<^sub>R y) = ?d *\<^sub>R ((1 - ?t) *\<^sub>R r *\<^sub>R x + ?t *\<^sub>R y)" if "0 < u" "u < 1" using that eq1 eq2 by (auto simp: algebra_simps) show ?case apply (cases "u = 0") subgoal using 1 by (intro exI[where x="c / r"] exI[where x="r *\<^sub>R x"]) auto apply (cases "u = 1") subgoal using 1 by (intro exI[where x="c"] exI[where x="y"]) (auto intro!: exI[where x=1]) subgoal apply (rule exI[where x="?d"]) apply (rule exI[where x="((1 - ?t) *\<^sub>R r *\<^sub>R x + ?t *\<^sub>R y)"]) apply (subst *) using 1 apply (auto intro!: exI[where x = ?t]) apply (auto simp: algebra_simps divide_simps) defer proof - assume a1: "c + c * (r * u) < c * u" assume a2: "0 \ c" assume a3: "0 \ u" assume a4: "u \ 0" assume a5: "0 < r" have "c + c * (r * u) \ c * u" using a1 less_eq_real_def by blast then show "c \ c * u" using a5 a4 a3 a2 by (metis (no_types) less_add_same_cancel1 less_eq_real_def mult_pos_pos order_trans real_scaleR_def real_vector.scale_zero_left) next assume a1: "0 \ c" assume a2: "u \ 1" have f3: "\x0. ((x0::real) < 1) = (\ 1 \ x0)" by auto have f4: "\x0. ((1::real) < x0) = (\ x0 \ 1)" by fastforce have "\x0 x1. ((x1::real) < x1 * x0) = (\ 0 \ x1 + - 1 * (x1 * x0))" by auto then have "(\r ra. ((r::real) < r * ra) = ((0 \ r \ 1 < ra) \ (r \ 0 \ ra < 1))) = (\r ra. (\ (0::real) \ r + - 1 * (r * ra)) = ((\ 0 \ r \ \ ra \ 1) \ (\ r \ 0 \ \ 1 \ ra)))" using f4 f3 by presburger then have "0 \ c + - 1 * (c * u)" using a2 a1 mult_less_cancel_left1 by blast then show "c * u \ c" by auto qed done qed lemma conefield_scaleRD: "v \ conefield x y" if "v \ conefield (r *\<^sub>R x) y" "r > 0" using conefield_scaleRI[OF that(1) positive_imp_inverse_positive[OF that(2)]] that(2) by auto lemma conefield_scaleR: "conefield (r *\<^sub>R x) y = conefield x y" if "r > 0" using conefield_scaleRD conefield_scaleRI that by blast lemma conefield_expansion_estimate: fixes u v::"'a::euclidean_space" and F::"'a \ 'a" assumes "t \ {0 .. pi / 2}" assumes angle_pos: "0 < vangle u v" "vangle u v < pi / 2" assumes angle_le: "vangle (F u) (F v) \ vangle u v" assumes "bounded_linear F" assumes "x \ conefield u v" shows "norm (F x) \ min (norm (F u)/norm u) (norm (F v)/norm v) * norm x" proof cases assume [simp]: "x \ 0" from assms have [simp]: "u \ 0" "v \ 0" by auto interpret bounded_linear F by fact define u1 where "u1 = u /\<^sub>R norm u" define v1 where "v1 = v /\<^sub>R norm v" note \x \ conefield u v\ also have \conefield u v = conefield u1 v1\ by (auto simp: u1_def v1_def conefield_scaleR conefield_commute[of u]) finally obtain c t where x: "x = c *\<^sub>R conemem u1 v1 t" "t \ {0 .. pi / 2}" "c \ 0" by (auto simp: conefield_def cone_hull_expl conesegment_def) then have xc: "x /\<^sub>R c = conemem u1 v1 t" by (auto simp: divide_simps) also have "F \ = conemem (F u1) (F v1) t" by (simp add: bounded_linear_image_conemem assms) also have "norm \ \ min (norm (F u1)) (norm (F v1)) * norm (conemem u1 v1 t)" apply (rule conemem_expansion_estimate) subgoal by fact subgoal using angle_pos by (simp add: u1_def v1_def vangle_scaleR) subgoal using angle_pos by (simp add: u1_def v1_def vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) done finally show "norm (F x) \ min (norm (F u)/norm u) (norm (F v)/norm v) * norm x" unfolding xc[symmetric] scaleR u1_def v1_def norm_scaleR x using \c \ 0\ by (simp add: divide_simps split: if_splits) qed simp lemma conefield_rightI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "y \ {y1 .. y2}" shows "(i + y *\<^sub>R j) \ conefield (i + y1 *\<^sub>R j) (i + y2 *\<^sub>R j)" unfolding conefield_alt_def apply (rule hull_inc) using assms by (auto simp: in_segment divide_simps inner_Basis algebra_simps intro!: exI[where x="(y - y1) / (y2 - y1)"] euclidean_eqI[where 'a='a] ) lemma conefield_right_vangleI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "y \ {y1 .. y2}" "y1 < y2" shows "(i + y *\<^sub>R j) \ conefield (i + y1 *\<^sub>R j) (i + y2 *\<^sub>R j)" unfolding conefield_alt_def apply (rule hull_inc) using assms by (auto simp: in_segment divide_simps inner_Basis algebra_simps intro!: exI[where x="(y - y1) / (y2 - y1)"] euclidean_eqI[where 'a='a] ) lemma cone_conefield[intro, simp]: "cone (conefield x y)" unfolding conefield_def by (rule cone_cone_hull) lemma conefield_mk_rightI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "(i + (y / x) *\<^sub>R j) \ conefield (i + (y1 / x1) *\<^sub>R j) (i + (y2 / x2) *\<^sub>R j)" assumes "x > 0" "x1 > 0" "x2 > 0" shows "(x *\<^sub>R i + y *\<^sub>R j) \ conefield (x1 *\<^sub>R i + y1 *\<^sub>R j) (x2 *\<^sub>R i + y2 *\<^sub>R j)" proof - have rescale: "(x *\<^sub>R i + y *\<^sub>R j) = x *\<^sub>R (i + (y / x) *\<^sub>R j)" if "x > 0" for x y using that by (auto simp: algebra_simps) show ?thesis unfolding rescale[OF \x > 0\] rescale[OF \x1 > 0\] rescale[OF \x2 > 0\] conefield_scaleR[OF \x1 > 0\] apply (subst conefield_commute) unfolding conefield_scaleR[OF \x2 > 0\] apply (rule mem_cone) apply simp apply (subst conefield_commute) by (auto intro!: assms less_imp_le) qed lemma conefield_prod3I: assumes "x > 0" "x1 > 0" "x2 > 0" assumes "y1 / x1 \ y / x" "y / x \ y2 / x2" shows "(x, y, 0) \ (conefield (x1, y1, 0) (x2, y2, 0)::(real*real*real) set)" proof - have "(x *\<^sub>R (1, 0, 0) + y *\<^sub>R (0, 1, 0)) \ (conefield (x1 *\<^sub>R (1, 0, 0) + y1 *\<^sub>R (0, 1, 0)) (x2 *\<^sub>R (1, 0, 0) + y2 *\<^sub>R (0, 1, 0))::(real*real*real) set)" apply (rule conefield_mk_rightI) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal using assms by (intro conefield_rightI) (auto simp: Basis_prod_def zero_prod_def) by (auto intro: assms) then show ?thesis by simp qed end diff --git a/thys/Planarity_Certificates/Planarity/Digraph_Map_Impl.thy b/thys/Planarity_Certificates/Planarity/Digraph_Map_Impl.thy --- a/thys/Planarity_Certificates/Planarity/Digraph_Map_Impl.thy +++ b/thys/Planarity_Certificates/Planarity/Digraph_Map_Impl.thy @@ -1,412 +1,412 @@ theory Digraph_Map_Impl imports Graph_Genus Executable_Permutations "Transitive-Closure.Transitive_Closure_Impl" begin section \Enumerating Maps\ definition grouped_by_fst :: "('a \ 'b) list \ ('a \ 'b) list list" where "grouped_by_fst xs = map (\u. filter (\x. fst x = u) xs) (remdups (map fst xs))" fun grouped_out_arcs :: "'a list \ ('a \ 'a) list \ ('a \ 'a) list list" where "grouped_out_arcs (vs,as) = grouped_by_fst as" definition all_maps_list :: "('a list \ ('a \ 'a) list) \ ('a \ 'a) list list list" where "all_maps_list G_list = (cyc_permutationss o grouped_out_arcs) G_list" definition "list_digraph_ext ext G_list \ \ pverts = set (fst G_list), parcs = set (snd G_list), \ = ext \" abbreviation "list_digraph \ list_digraph_ext ()" code_datatype list_digraph_ext lemma list_digraph_simps: "pverts (list_digraph G_list) = set (fst G_list)" "parcs (list_digraph G_list) = set (snd G_list)" by (auto simp: list_digraph_ext_def) lemma union_grouped_by_fst: "(\xs \ set (grouped_by_fst ys). set xs) = set ys" by (auto simp: grouped_by_fst_def) lemma union_grouped_out_arcs: "(\xs \ set (grouped_out_arcs G_list). set xs) = set (snd G_list)" by (cases G_list) (simp add: union_grouped_by_fst) lemma nil_not_in_grouped_out_arcs: "[] \ set (grouped_out_arcs G_list)" apply (cases G_list) apply (auto simp: grouped_by_fst_def ) by (metis (mono_tags) filter_empty_conv fst_conv) lemma set_grouped_out_arcs: assumes "pair_wf_digraph (list_digraph G_list)" shows "set ` set (grouped_out_arcs G_list) = {out_arcs (list_digraph G_list) v | v. v \ pverts (list_digraph G_list) \ out_arcs (list_digraph G_list) v \ {} }" (is "?L = ?R") proof - interpret pair_wf_digraph "list_digraph G_list" by fact define vs where "vs = remdups (map fst (snd G_list))" have "set vs = {v. out_arcs (list_digraph G_list) v \ {}}" by (auto simp: out_arcs_def list_digraph_ext_def vs_def intro: rev_image_eqI ) then have vs: "set vs = {v \ pverts (list_digraph G_list). out_arcs (list_digraph G_list) v \ {}}" by (auto dest: in_arcsD1) have goa: "grouped_out_arcs G_list = map (\u. filter (\x. fst x = u) (snd G_list)) vs" by (cases G_list) (auto simp: grouped_by_fst_def vs_def) have filter: "set \ (\u. filter (\x. fst x = u) (snd G_list)) = out_arcs (list_digraph G_list)" by (rule ext) (auto simp: list_digraph_ext_def) have "set (map set (grouped_out_arcs G_list)) = ?R" by (auto simp add: goa filter vs) then show ?thesis by simp qed lemma distincts_grouped_by_fst: assumes "distinct xs" shows "distincts (grouped_by_fst xs)" proof - have list_eq_setD: "\xs ys. xs = ys \ set xs = set ys" by auto have inj: "inj_on (\u. filter (\x. fst x = u) xs) (fst ` set xs)" by (rule inj_onI) (drule list_eq_setD, auto) with assms show ?thesis by (auto simp: grouped_by_fst_def distincts_def distinct_map filter_empty_conv) qed lemma distincts_grouped_arcs: assumes "distinct (snd G_list)" shows "distincts (grouped_out_arcs G_list)" using assms by (cases G_list) (simp add: distincts_grouped_by_fst) lemma distincts_in_all_maps_list: "distinct (snd X) \ xss \ set (all_maps_list X) \ distincts xss" by (simp add: all_maps_list_def distincts_grouped_arcs in_set_cyc_permutationss) definition to_map :: "('a \ 'a) set \ ('a \ 'a \ 'a \ 'a) \ ('a \ 'a) pre_map" where "to_map A f = \ edge_rev = swap_in A, edge_succ = f \" abbreviation "to_map' as xss \ to_map (set as) (lists_succ xss)" definition all_maps :: "'a pair_pre_digraph \ ('a \ 'a) pre_map set" where "all_maps G \ to_map (arcs G) ` {f. f permutes arcs G \ (\v \ verts G. out_arcs G v \ {} \ cyclic_on f (out_arcs G v))}" definition maps_all_maps_list :: "('a list \ ('a \ 'a) list) \ ('a \ 'a) pre_map list" where "maps_all_maps_list G_list = map (to_map (set (snd G_list)) o lists_succ) (all_maps_list G_list)" lemma (in pair_graph) all_maps_correct: shows "all_maps G = {M. digraph_map G M}" proof (intro set_eqI iffI) fix M assume A:"M \ all_maps G" then have [simp]: "edge_rev M = swap_in (arcs G)" "edge_succ M permutes parcs G" by (auto simp: all_maps_def to_map_def) have "digraph_map G M" proof (rule digraph_mapI) fix a assume "a \ parcs G" then show "edge_rev M a = a" by (auto simp: swap_in_def) next fix a assume "a \ parcs G" then show "edge_rev M (edge_rev M a) = a" "fst (edge_rev M a) = snd a" "edge_rev M a \ a" by (case_tac [!] "a") (auto intro: arcs_symmetric simp: swap_in_def dest: no_loops) next show "edge_succ M permutes parcs G" by simp next fix v assume "v \ pverts G" "out_arcs (with_proj G) v \ {}" then show "cyclic_on (edge_succ M) (out_arcs (with_proj G) v)" using A unfolding all_maps_def by (auto simp: to_map_def) qed then show "M \ {M. digraph_map G M}" by simp next fix M assume A: "M \ {M. digraph_map G M}" then interpret M: digraph_map G M by simp from A have "\x. fst (edge_rev M x) = fst (swap_in (arcs G) x)" "\x. snd (edge_rev M x) = snd (swap_in (arcs G) x)" using M.tail_arev M.head_arev by (auto simp: fun_eq_iff swap_in_def M.arev_eq) then have "edge_rev M = swap_in (arcs G)" by (metis prod.collapse fun_eq_iff) then show "M \ all_maps G" using M.edge_succ_permutes M.edge_succ_cyclic unfolding all_maps_def by (auto simp: to_map_def intro!: image_eqI[where x="edge_succ M"]) qed lemma set_maps_all_maps_list: assumes "pair_wf_digraph (list_digraph G_list)" "distinct (snd G_list)" shows "all_maps (list_digraph G_list) = set (maps_all_maps_list G_list)" proof - let ?G = "list_digraph G_list" { fix f have "(\x\set (grouped_out_arcs G_list). cyclic_on f (set x)) \ (\x\set ` set (grouped_out_arcs G_list). cyclic_on f x)" (is "?all1 = _") by simp also have "\ \ (\v\pverts ?G. out_arcs ?G v \ {} \ cyclic_on f (out_arcs ?G v))" (is "_ = ?all2") using assms by (auto simp: set_grouped_out_arcs) finally have "?all1 = ?all2" . } note all_eq = this have "lists_succ ` set (all_maps_list G_list) = {f. f permutes arcs ?G \ (\v \ pverts ?G. out_arcs ?G v \ {} \ cyclic_on f (out_arcs ?G v))}" unfolding all_maps_list_def using assms all_eq by (simp add: lists_succ_set_cyc_permutationss distincts_grouped_arcs union_grouped_out_arcs list_digraph_simps) then have *: "lists_succ ` set (all_maps_list G_list) = {f. f permutes set (snd G_list) \ (\v\set (fst G_list). out_arcs (with_proj \pverts = set (fst G_list), parcs = set (snd G_list)\) v \ {} \ cyclic_on f (out_arcs (with_proj \pverts = set (fst G_list), parcs = set (snd G_list)\) v))}" by (auto simp add: maps_all_maps_list_def all_maps_def list_digraph_simps list_digraph_ext_def) then have **: "\f. \ (f permutes set (snd G_list) \ (\a. a \ set (fst G_list) \ out_arcs (with_proj \pverts = set (fst G_list), parcs = set (snd G_list)\) a \ {} \ cyclic_on f (out_arcs (with_proj \pverts = set (fst G_list), parcs = set (snd G_list)\) a))) \ f \ lists_succ ` set (all_maps_list G_list)" by force from * show ?thesis by (auto simp add: maps_all_maps_list_def all_maps_def list_digraph_simps list_digraph_ext_def) (use ** in blast) qed section \Compute Face Cycles\ definition lists_fc_succ :: "('a \ 'a) list list \ ('a \ 'a) \ ('a \ 'a)" where "lists_fc_succ xss = (let sxss = \(sset xss) in (\x. lists_succ xss (swap_in sxss x)))" locale lists_digraph_map = fixes G_list :: "'b list \ ('b \ 'b) list" and xss :: "('b \ 'b) list list" assumes digraph_map: "digraph_map (list_digraph G_list) (to_map' (snd G_list) xss)" assumes no_loops: "\a. a \ parcs (list_digraph G_list) \ fst a \ snd a" assumes distincts_xss: "distincts xss" assumes parcs_xss: "parcs (list_digraph G_list) = \(sset xss)" begin abbreviation (input) "G \ list_digraph G_list" abbreviation (input) "M \ to_map' (snd G_list) xss" lemma edge_rev_simps: assumes "(u,v) \ parcs G" shows "edge_rev M (u,v) = (v,u)" using assms unfolding to_map_def list_digraph_ext_def by (auto simp: swap_in_def to_map_def) end sublocale lists_digraph_map \ digraph_map G M by (rule local.digraph_map) sublocale lists_digraph_map \ pair_graph G proof fix e assume "e \ parcs G" then have "e \ arcs G" by simp then have "head G e \ verts G" "tail G e \ verts G" by (blast dest: wellformed)+ then show "fst e \ pverts G" "snd e \ pverts G" by auto next fix e assume "e \ parcs G" then show "fst e \ snd e" using no_loops by simp next show "finite (pverts G)" "finite (parcs G)" unfolding list_digraph_ext_def by simp_all next { fix u v assume "(u,v) \ parcs G" then have "edge_rev M (u,v) \ parcs G" using edge_rev_in_arcs by simp then have "(v,u) \ parcs G" using \(u,v) \ _ \ by (simp add: edge_rev_simps) } then show "symmetric G" unfolding symmetric_def by (auto intro: symI) qed context lists_digraph_map begin definition "lists_fcs \ orbits_list (lists_fc_succ xss)" lemma M_simps: "edge_succ M = lists_succ xss" unfolding to_map_def by (cases G_list) auto lemma lists_fc_succ_permutes: "lists_fc_succ xss permutes (\(sset xss))" proof - have "\(u,v) \ \(sset xss). (v,u) \ \(sset xss)" using sym_arcs unfolding parcs_xss[symmetric] symmetric_def by (auto elim: symE) then have "swap_in (\(sset xss)) permutes \(sset xss)" using distincts_xss apply (auto simp: permutes_def split: if_splits) unfolding swap_in_def apply (simp_all split: if_splits prod.splits) apply metis+ done moreover have "lists_succ xss permutes (\(sset xss))" using lists_succ_permutes[OF distincts_xss] by simp moreover have "lists_fc_succ xss = lists_succ xss o swap_in (\(sset xss))" by (simp add: fun_eq_iff lists_fc_succ_def) ultimately show ?thesis by (metis permutes_compose) qed lemma permutation_lists_fc_succ[intro, simp]: "permutation (lists_fc_succ xss)" using lists_fc_succ_permutes by (auto simp: permutation_permutes) lemma face_cycle_succ_conv: "face_cycle_succ = lists_fc_succ xss" using parcs_xss unfolding face_cycle_succ_def by (simp add: fun_eq_iff to_map_def lists_fc_succ_def swap_in_def list_digraph_ext_def) lemma sset_lists_fcs: "sset (lists_fcs as) = {face_cycle_set a | a. a \ set as}" by (auto simp: lists_fcs_def sset_orbits_list face_cycle_set_def face_cycle_succ_conv) lemma distincts_lists_fcs: "distinct as \distincts (lists_fcs as)" by (simp add: lists_fcs_def distincts_orbits_list) lemma face_cycle_set_ss: "a \ parcs G \ face_cycle_set a \ parcs G" using in_face_cycle_setD with_proj_simps(2) by blast lemma face_cycle_succ_neq: assumes "a \ parcs G" shows "face_cycle_succ a \ a" using assms no_loops by (intro face_cycle_succ_neq) auto lemma card_face_cycle_sets_conv: shows "card (pre_digraph_map.face_cycle_sets G M) = length (lists_fcs (remdups (snd G_list)))" proof - interpret digraph_map G M by (rule digraph_map) have "face_cycle_sets = {face_cycle_set a | a. a \ parcs G}" by (auto simp: face_cycle_sets_def) also have "\ = sset (lists_fcs (remdups (snd G_list)))" unfolding sset_lists_fcs by (simp add: list_digraph_simps) also have "card \ = length (lists_fcs (remdups (snd G_list)))" by (simp add: card_image distincts_inj_on_set distinct_card distincts_distinct distincts_lists_fcs) finally show ?thesis . qed end definition "gen_succ \ \as xs. [b. (a,b) <- as, a \ set xs]" interpretation RTLI: set_access_gen set "\x xs. x \ set xs" "[]" "\xs ys. remdups (xs @ ys)" "gen_succ" by standard (auto simp: gen_succ_def) hide_const (open) gen_succ text \ It would suffice to check that @{term "set (RTLI.rtrancl_i A [u]) = set V"}. We don't do this here, since it makes the proof more complicated (and is not necessary for the graphs we care about \ definition sccs_verts_impl :: "'a list \ ('a \ 'a) list \ 'a set set" where "sccs_verts_impl G \ set ` (\x. RTLI.rtrancl_i (snd G) [x]) ` set (fst G)" definition isolated_verts_impl :: "'a list \ ('a \ 'a) list \ 'a list" where "isolated_verts_impl G = [v \ (fst G). \(\e \ set (snd G). fst e = v)]" definition pair_graph_impl :: "'a list \ ('a \ 'a) list \ bool" where "pair_graph_impl G \ case G of (V,A) \ (\(u,v) \ set A. u \ v \ u \ set V \ v \ set V \ (v,u) \ set A)" definition genus_impl :: "'a list \ ('a \ 'a) list \ ('a \ 'a) list list \ int" where "genus_impl G M \ case G of (V,A) \ (int (2*card (sccs_verts_impl G)) - int (length (isolated_verts_impl G)) - (int (length V) - int (length A) div 2 + int (length (orbits_list_impl (lists_fc_succ M) A)))) div 2" definition comb_planar_impl :: "'a list \ ('a \ 'a) list \ bool" where "comb_planar_impl G \ case G of (V,A) \ let i = int (2*card (sccs_verts_impl G)) - int (length (isolated_verts_impl G)) - int (length V) + int (length A) div 2 in (\M\set (all_maps_list G). (i - int (length (orbits_list_impl (lists_fc_succ M) A))) div 2 = 0)" (* definition comb_planar_impl :: "'a list \ ('a \ 'a) list \ bool" where "comb_planar_impl G \ \M\set (all_maps_list G). genus_impl G M = 0" *) lemma sccs_verts_impl_correct: assumes "pair_pseudo_graph (list_digraph G)" shows "pre_digraph.sccs_verts (list_digraph G) = sccs_verts_impl G" proof - interpret pair_pseudo_graph "list_digraph G" by fact { fix u assume "u \ set (fst G)" then have "\x. (u, x) \ (set (snd G))\<^sup>* \ x \ set (fst G)" by (metis in_arcsD2 list_digraph_simps rtrancl.cases) then have "set (RTLI.rtrancl_i (snd G) [u]) = {v. u \\<^sup>*\<^bsub>list_digraph G\<^esub> v }" unfolding RTLI.rtrancl_impl reachable_conv by (auto simp: list_digraph_simps \u \ _\) also have "\ = scc_of u" unfolding scc_of_def by (auto intro: symmetric_reachable') finally have "scc_of u = set (RTLI.rtrancl_i (snd G) [u])" by simp } then have "pre_digraph.sccs_verts (list_digraph G) = set ` (\x. RTLI.rtrancl_i (snd G) [x]) ` set (fst G)" unfolding sccs_verts_conv_scc_of list_digraph_simps by (force intro: rev_image_eqI) then show ?thesis unfolding sccs_verts_impl_def by simp qed lemma isolated_verts_impl_correct: "pre_digraph.isolated_verts (list_digraph G) = set (isolated_verts_impl G)" by (auto simp: pre_digraph.isolated_verts_def isolated_verts_impl_def list_digraph_simps out_arcs_def) lemma pair_graph_impl_correct[code]: "pair_graph (list_digraph G) = pair_graph_impl G" (is "?L = ?R") unfolding pair_graph_def pair_digraph_def pair_fin_digraph_def pair_wf_digraph_def pair_fin_digraph_axioms_def pair_loopfree_digraph_def pair_loopfree_digraph_axioms_def pair_sym_digraph_def pair_sym_digraph_axioms_def pair_pseudo_graph_def pair_graph_impl_def by (auto simp: pair_graph_impl_def list_digraph_simps symmetric_def intro: symI dest: symD split: prod.splits) lemma genus_impl_correct: assumes dist_V: "distinct (fst G)" and dist_A: "distinct (snd G)" assumes "lists_digraph_map G M" shows "pre_digraph_map.euler_genus (list_digraph G) (to_map' (snd G) M) = genus_impl G M" proof - interpret lists_digraph_map G M by fact obtain V A where G_eq: "G = (V,A)" by (cases G) moreover have "distinct (isolated_verts_impl G)" using dist_V by (auto simp: isolated_verts_impl_def ) moreover have faces: "card face_cycle_sets = length (orbits_list_impl (lists_fc_succ M) (snd G))" using dist_A by (simp add: card_face_cycle_sets_conv lists_fcs_def orbits_list_conv_impl distinct_remdups_id) ultimately show ?thesis using pair_pseudo_graph dist_V dist_A unfolding euler_genus_def euler_char_def genus_impl_def card_sccs_verts[symmetric] by (simp add: sccs_verts_impl_correct isolated_verts_impl_correct distinct_card list_digraph_simps zdiv_int) qed lemma elems_all_maps_list: assumes "M \ set (all_maps_list G)" "distinct (snd G)" shows "\(sset M) = set (snd G)" using assms by (simp add: all_maps_list_def in_set_cyc_permutationss distincts_grouped_arcs union_grouped_out_arcs[symmetric]) (metis set_map) lemma comb_planar_impl_altdef: "comb_planar_impl G = (\M\set (all_maps_list G). genus_impl G M = 0)" - unfolding comb_planar_impl_def Let_def genus_impl_def by (cases G) (simp add: sign_simps) + unfolding comb_planar_impl_def Let_def genus_impl_def by (cases G) (simp add: algebra_simps) lemma comb_planar_impl_correct: assumes "pair_graph (list_digraph G)" assumes dist_V: "distinct (fst G)" and dist_A: "distinct (snd G)" shows "comb_planar (list_digraph G) = comb_planar_impl G" (is "?L = ?R") proof - interpret G: pair_graph "list_digraph G" by fact let ?G = "list_digraph G" have *: "all_maps (list_digraph G) = set (maps_all_maps_list G)" by (rule set_maps_all_maps_list) (unfold_locales, simp add: dist_A) obtain V A where "G = (V,A)" by (cases G) { fix M assume "M \ set (all_maps_list G)" have "digraph_map (list_digraph G) (to_map' (snd G) M)" using \M \ _\ G.all_maps_correct by (auto simp: * maps_all_maps_list_def) then interpret G: digraph_map "list_digraph G" "to_map' (snd G) M" . have "distincts M" using \M \ _\ using dist_A distincts_in_all_maps_list by blast have "lists_digraph_map G M" using elems_all_maps_list[OF \M \ _\ \distinct (snd G)\] apply unfold_locales by (auto intro: \distincts M\ dest: G.adj_not_same) (auto simp: list_digraph_simps) } note ldm = this have "comb_planar ?G = (\M \ {M. digraph_map ?G M}. pre_digraph_map.euler_genus ?G M = 0)" unfolding comb_planar_def by simp also have "\ = (\M\set (all_maps_list G). pre_digraph_map.euler_genus (list_digraph G) (to_map (set (snd G)) (lists_succ M)) = 0)" unfolding comb_planar_def comb_planar_impl_def Let_def G.all_maps_correct[symmetric] set_maps_all_maps_list[OF G.pair_wf_digraph dist_A] maps_all_maps_list_def by simp also have "\ = (\M\set (all_maps_list G). genus_impl G M = 0)" using ldm assms by (simp add: genus_impl_correct) also have "\ = comb_planar_impl G" - unfolding comb_planar_impl_def genus_impl_def Let_def by (simp add: \G = (V,A)\ sign_simps) + unfolding comb_planar_impl_def genus_impl_def Let_def by (simp add: \G = (V,A)\ algebra_simps) finally show ?thesis . qed end diff --git a/thys/Polynomial_Interpolation/Missing_Polynomial.thy b/thys/Polynomial_Interpolation/Missing_Polynomial.thy --- a/thys/Polynomial_Interpolation/Missing_Polynomial.thy +++ b/thys/Polynomial_Interpolation/Missing_Polynomial.thy @@ -1,1401 +1,1401 @@ (* Author: René Thiemann Akihisa Yamada Jose Divason License: BSD *) section \Missing Polynomial\ text \The theory contains some basic results on polynomials which have not been detected in the distribution, especially on linear factors and degrees.\ theory Missing_Polynomial imports "HOL-Computational_Algebra.Polynomial_Factorial" Missing_Unsorted begin subsection \Basic Properties\ lemma degree_0_id: assumes "degree p = 0" shows "[: coeff p 0 :] = p" proof - have "\ x. 0 \ Suc x" by auto thus ?thesis using assms by (metis coeff_pCons_0 degree_pCons_eq_if pCons_cases) qed lemma degree0_coeffs: "degree p = 0 \ \ a. p = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) lemma degree1_coeffs: "degree p = 1 \ \ a b. p = [: b, a :] \ a \ 0" by (metis One_nat_def degree_pCons_eq_if nat.inject old.nat.distinct(2) pCons_0_0 pCons_cases) lemma degree2_coeffs: "degree p = 2 \ \ a b c. p = [: c, b, a :] \ a \ 0" by (metis Suc_1 Suc_neq_Zero degree1_coeffs degree_pCons_eq_if nat.inject pCons_cases) lemma poly_zero: fixes p :: "'a :: comm_ring_1 poly" assumes x: "poly p x = 0" shows "p = 0 \ degree p = 0" proof assume degp: "degree p = 0" hence "poly p x = coeff p (degree p)" by(subst degree_0_id[OF degp,symmetric], simp) hence "coeff p (degree p) = 0" using x by auto thus "p = 0" by auto qed auto lemma coeff_monom_Suc: "coeff (monom a (Suc d) * p) (Suc i) = coeff (monom a d * p) i" by (simp add: monom_Suc) lemma coeff_sum_monom: assumes n: "n \ d" shows "coeff (\i\d. monom (f i) i) n = f n" (is "?l = _") proof - have "?l = (\i\d. coeff (monom (f i) i) n)" (is "_ = sum ?cmf _") using coeff_sum. also have "{..d} = insert n ({..d}-{n})" using n by auto hence "sum ?cmf {..d} = sum ?cmf ..." by auto also have "... = sum ?cmf ({..d}-{n}) + ?cmf n" by (subst sum.insert,auto) also have "sum ?cmf ({..d}-{n}) = 0" by (subst sum.neutral, auto) finally show ?thesis by simp qed lemma linear_poly_root: "(a :: 'a :: comm_ring_1) \ set as \ poly (\ a \ as. [: - a, 1:]) a = 0" proof (induct as) case (Cons b as) show ?case proof (cases "a = b") case False with Cons have "a \ set as" by auto from Cons(1)[OF this] show ?thesis by simp qed simp qed simp lemma degree_lcoeff_sum: assumes deg: "degree (f q) = n" and fin: "finite S" and q: "q \ S" and degle: "\ p . p \ S - {q} \ degree (f p) < n" and cong: "coeff (f q) n = c" shows "degree (sum f S) = n \ coeff (sum f S) n = c" proof (cases "S = {q}") case True thus ?thesis using deg cong by simp next case False with q obtain p where "p \ S - {q}" by auto from degle[OF this] have n: "n > 0" by auto have "degree (sum f S) = degree (f q + sum f (S - {q}))" unfolding sum.remove[OF fin q] .. also have "\ = degree (f q)" proof (rule degree_add_eq_left) have "degree (sum f (S - {q})) \ n - 1" proof (rule degree_sum_le) fix p show "p \ S - {q} \ degree (f p) \ n - 1" using degle[of p] by auto qed (insert fin, auto) also have "\ < n" using n by simp finally show "degree (sum f (S - {q})) < degree (f q)" unfolding deg . qed finally show ?thesis unfolding deg[symmetric] cong[symmetric] proof (rule conjI) have id: "(\x\S - {q}. coeff (f x) (degree (f q))) = 0" by (rule sum.neutral, rule ballI, rule coeff_eq_0[OF degle[folded deg]]) show "coeff (sum f S) (degree (f q)) = coeff (f q) (degree (f q))" unfolding coeff_sum by (subst sum.remove[OF _ q], unfold id, insert fin, auto) qed qed lemma degree_sum_list_le: "(\ p . p \ set ps \ degree p \ n) \ degree (sum_list ps) \ n" proof (induct ps) case (Cons p ps) hence "degree (sum_list ps) \ n" "degree p \ n" by auto thus ?case unfolding sum_list.Cons by (metis degree_add_le) qed simp lemma degree_prod_list_le: "degree (prod_list ps) \ sum_list (map degree ps)" proof (induct ps) case (Cons p ps) show ?case unfolding prod_list.Cons by (rule order.trans[OF degree_mult_le], insert Cons, auto) qed simp lemma smult_sum: "smult (\i \ S. f i) p = (\i \ S. smult (f i) p)" by (induct S rule: infinite_finite_induct, auto simp: smult_add_left) lemma range_coeff: "range (coeff p) = insert 0 (set (coeffs p))" by (metis nth_default_coeffs_eq range_nth_default) lemma smult_power: "(smult a p) ^ n = smult (a ^ n) (p ^ n)" by (induct n, auto simp: field_simps) lemma poly_sum_list: "poly (sum_list ps) x = sum_list (map (\ p. poly p x) ps)" by (induct ps, auto) lemma poly_prod_list: "poly (prod_list ps) x = prod_list (map (\ p. poly p x) ps)" by (induct ps, auto) lemma sum_list_neutral: "(\ x. x \ set xs \ x = 0) \ sum_list xs = 0" by (induct xs, auto) lemma prod_list_neutral: "(\ x. x \ set xs \ x = 1) \ prod_list xs = 1" by (induct xs, auto) lemma (in comm_monoid_mult) prod_list_map_remove1: "x \ set xs \ prod_list (map f xs) = f x * prod_list (map f (remove1 x xs))" by (induct xs) (auto simp add: ac_simps) lemma poly_as_sum: fixes p :: "'a::comm_semiring_1 poly" shows "poly p x = (\i\degree p. x ^ i * coeff p i)" unfolding poly_altdef by (simp add: ac_simps) lemma poly_prod_0: "finite ps \ poly (prod f ps) x = (0 :: 'a :: field) \ (\ p \ ps. poly (f p) x = 0)" by (induct ps rule: finite_induct, auto) lemma coeff_monom_mult: shows "coeff (monom a d * p) i = (if d \ i then a * coeff p (i-d) else 0)" (is "?l = ?r") proof (cases "d \ i") case False thus ?thesis unfolding coeff_mult by simp next case True let ?f = "\j. coeff (monom a d) j * coeff p (i - j)" have "\j. j \ {0..i} - {d} \ ?f j = 0" by auto hence "0 = (\j \ {0..i} - {d}. ?f j)" by auto also have "... + ?f d = (\j \ insert d ({0..i} - {d}). ?f j)" by(subst sum.insert, auto) also have "... = (\j \ {0..i}. ?f j)" by (subst insert_Diff, insert True, auto) also have "... = (\j\i. ?f j)" by (rule sum.cong, auto) also have "... = ?l" unfolding coeff_mult .. finally show ?thesis using True by auto qed lemma poly_eqI2: assumes "degree p = degree q" and "\i. i \ degree p \ coeff p i = coeff q i" shows "p = q" apply(rule poly_eqI) by (metis assms le_degree) text \A nice extension rule for polynomials.\ lemma poly_ext[intro]: fixes p q :: "'a :: {ring_char_0, idom} poly" assumes "\x. poly p x = poly q x" shows "p = q" unfolding poly_eq_poly_eq_iff[symmetric] using assms by (rule ext) text \Copied from non-negative variants.\ lemma coeff_linear_power_neg[simp]: fixes a :: "'a::comm_ring_1" shows "coeff ([:a, -1:] ^ n) n = (-1)^n" apply (induct n, simp_all) apply (subst coeff_eq_0) apply (auto intro: le_less_trans degree_power_le) done lemma degree_linear_power_neg[simp]: fixes a :: "'a::{idom,comm_ring_1}" shows "degree ([:a, -1:] ^ n) = n" apply (rule order_antisym) apply (rule ord_le_eq_trans [OF degree_power_le], simp) apply (rule le_degree) unfolding coeff_linear_power_neg apply (auto) done subsection \Polynomial Composition\ lemmas [simp] = pcompose_pCons lemma pcompose_eq_0: fixes q :: "'a :: idom poly" assumes q: "degree q \ 0" shows "p \\<^sub>p q = 0 \ p = 0" proof (induct p) case 0 show ?case by auto next case (pCons a p) have id: "(pCons a p) \\<^sub>p q = [:a:] + q * (p \\<^sub>p q)" by simp show ?case proof (cases "p = 0") case True show ?thesis unfolding id unfolding True by simp next case False with pCons(2) have "p \\<^sub>p q \ 0" by auto from degree_mult_eq[OF _ this, of q] q have "degree (q * (p \\<^sub>p q)) \ 0" by force hence deg: "degree ([:a:] + q * (p \\<^sub>p q)) \ 0" by (subst degree_add_eq_right, auto) show ?thesis unfolding id using False deg by auto qed qed declare degree_pcompose[simp] subsection \Monic Polynomials\ abbreviation monic where "monic p \ coeff p (degree p) = 1" lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {field,normalization_semidom}) = x" by (cases "is_unit x") (auto simp: is_unit_unit_factor dvd_field_iff) lemma poly_gcd_monic: fixes p :: "'a :: {field,factorial_ring_gcd} poly" assumes "p \ 0 \ q \ 0" shows "monic (gcd p q)" proof - from assms have "1 = unit_factor (gcd p q)" by (auto simp: unit_factor_gcd) also have "\ = [:lead_coeff (gcd p q):]" unfolding unit_factor_poly_def by (simp add: monom_0) finally show ?thesis by (metis coeff_pCons_0 degree_1 lead_coeff_1) qed lemma normalize_monic: "monic p \ normalize p = p" by (simp add: normalize_poly_eq_map_poly is_unit_unit_factor) lemma lcoeff_monic_mult: assumes monic: "monic (p :: 'a :: comm_semiring_1 poly)" shows "coeff (p * q) (degree p + degree q) = coeff q (degree q)" proof - let ?pqi = "\ i. coeff p i * coeff q (degree p + degree q - i)" have "coeff (p * q) (degree p + degree q) = (\i\degree p + degree q. ?pqi i)" unfolding coeff_mult by simp also have "\ = ?pqi (degree p) + (sum ?pqi ({.. degree p + degree q} - {degree p}))" by (subst sum.remove[of _ "degree p"], auto) also have "?pqi (degree p) = coeff q (degree q)" unfolding monic by simp also have "(sum ?pqi ({.. degree p + degree q} - {degree p})) = 0" proof (rule sum.neutral, intro ballI) fix d assume d: "d \ {.. degree p + degree q} - {degree p}" show "?pqi d = 0" proof (cases "d < degree p") case True hence "degree p + degree q - d > degree q" by auto hence "coeff q (degree p + degree q - d) = 0" by (rule coeff_eq_0) thus ?thesis by simp next case False with d have "d > degree p" by auto hence "coeff p d = 0" by (rule coeff_eq_0) thus ?thesis by simp qed qed finally show ?thesis by simp qed lemma degree_monic_mult: assumes monic: "monic (p :: 'a :: comm_semiring_1 poly)" and q: "q \ 0" shows "degree (p * q) = degree p + degree q" proof - have "degree p + degree q \ degree (p * q)" by (rule degree_mult_le) also have "degree p + degree q \ degree (p * q)" proof - from q have cq: "coeff q (degree q) \ 0" by auto hence "coeff (p * q) (degree p + degree q) \ 0" unfolding lcoeff_monic_mult[OF monic] . thus "degree (p * q) \ degree p + degree q" by (rule le_degree) qed finally show ?thesis . qed lemma degree_prod_sum_monic: assumes S: "finite S" and nzd: "0 \ (degree o f) ` S" and monic: "(\ a . a \ S \ monic (f a))" shows "degree (prod f S) = (sum (degree o f) S) \ coeff (prod f S) (sum (degree o f) S) = 1" proof - from S nzd monic have "degree (prod f S) = sum (degree \ f) S \ (S \ {} \ degree (prod f S) \ 0 \ prod f S \ 0) \ coeff (prod f S) (sum (degree o f) S) = 1" proof (induct S rule: finite_induct) case (insert a S) have IH1: "degree (prod f S) = sum (degree o f) S" using insert by auto have IH2: "coeff (prod f S) (degree (prod f S)) = 1" using insert by auto have id: "degree (prod f (insert a S)) = sum (degree \ f) (insert a S) \ coeff (prod f (insert a S)) (sum (degree o f) (insert a S)) = 1" proof (cases "S = {}") case False with insert have nz: "prod f S \ 0" by auto from insert have monic: "coeff (f a) (degree (f a)) = 1" by auto have id: "(degree \ f) a = degree (f a)" by simp show ?thesis unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] id unfolding degree_monic_mult[OF monic nz] unfolding IH1[symmetric] unfolding lcoeff_monic_mult[OF monic] IH2 by simp qed (insert insert, auto) show ?case using id unfolding sum.insert[OF insert(1-2)] using insert by auto qed simp thus ?thesis by auto qed lemma degree_prod_monic: assumes "\ i. i < n \ degree (f i :: 'a :: comm_semiring_1 poly) = 1" and "\ i. i < n \ coeff (f i) 1 = 1" shows "degree (prod f {0 ..< n}) = n \ coeff (prod f {0 ..< n}) n = 1" proof - from degree_prod_sum_monic[of "{0 ..< n}" f] show ?thesis using assms by force qed lemma degree_prod_sum_lt_n: assumes "\ i. i < n \ degree (f i :: 'a :: comm_semiring_1 poly) \ 1" and i: "i < n" and fi: "degree (f i) = 0" shows "degree (prod f {0 ..< n}) < n" proof - have "degree (prod f {0 ..< n}) \ sum (degree o f) {0 ..< n}" by (rule degree_prod_sum_le, auto) also have "sum (degree o f) {0 ..< n} = (degree o f) i + sum (degree o f) ({0 ..< n} - {i})" by (rule sum.remove, insert i, auto) also have "(degree o f) i = 0" using fi by simp also have "sum (degree o f) ({0 ..< n} - {i}) \ sum (\ _. 1) ({0 ..< n} - {i})" by (rule sum_mono, insert assms, auto) also have "\ = n - 1" using i by simp also have "\ < n" using i by simp finally show ?thesis by simp qed lemma degree_linear_factors: "degree (\ a \ as. [: f a, 1:]) = length as" proof (induct as) case (Cons b as) note IH = this have id: "(\a\b # as. [:f a, 1:]) = [:f b,1 :] * (\a\as. [:f a, 1:])" by simp show ?case unfolding id by (subst degree_monic_mult, insert IH, auto) qed simp lemma monic_mult: fixes p q :: "'a :: idom poly" assumes "monic p" "monic q" shows "monic (p * q)" proof - from assms have nz: "p \ 0" "q \ 0" by auto show ?thesis unfolding degree_mult_eq[OF nz] coeff_mult_degree_sum using assms by simp qed lemma monic_factor: fixes p q :: "'a :: idom poly" assumes "monic (p * q)" "monic p" shows "monic q" proof - from assms have nz: "p \ 0" "q \ 0" by auto from assms[unfolded degree_mult_eq[OF nz] coeff_mult_degree_sum \monic p\] show ?thesis by simp qed lemma monic_prod: fixes f :: "'a \ 'b :: idom poly" assumes "\ a. a \ as \ monic (f a)" shows "monic (prod f as)" using assms proof (induct as rule: infinite_finite_induct) case (insert a as) hence id: "prod f (insert a as) = f a * prod f as" and *: "monic (f a)" "monic (prod f as)" by auto show ?case unfolding id by (rule monic_mult[OF *]) qed auto lemma monic_prod_list: fixes as :: "'a :: idom poly list" assumes "\ a. a \ set as \ monic a" shows "monic (prod_list as)" using assms by (induct as, auto intro: monic_mult) lemma monic_power: assumes "monic (p :: 'a :: idom poly)" shows "monic (p ^ n)" by (induct n, insert assms, auto intro: monic_mult) lemma monic_prod_list_pow: "monic (\(x::'a::idom, i)\xis. [:- x, 1:] ^ Suc i)" proof (rule monic_prod_list, goal_cases) case (1 a) then obtain x i where a: "a = [:-x, 1:]^Suc i" by force show "monic a" unfolding a by (rule monic_power, auto) qed lemma monic_degree_0: "monic p \ (degree p = 0) = (p = 1)" using le_degree poly_eq_iff by force subsection \Roots\ text \The following proof structure is completely similar to the one of @{thm poly_roots_finite}.\ lemma poly_roots_degree: fixes p :: "'a::idom poly" shows "p \ 0 \ card {x. poly p x = 0} \ degree p" proof (induct n \ "degree p" arbitrary: p) case (0 p) then obtain a where "a \ 0" and "p = [:a:]" by (cases p, simp split: if_splits) then show ?case by simp next case (Suc n p) show ?case proof (cases "\x. poly p x = 0") case True then obtain a where a: "poly p a = 0" .. then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) then obtain k where k: "p = [:-a, 1:] * k" .. with \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from Suc.hyps(1)[OF this \k \ 0\] have le: "card {x. poly k x = 0} \ degree k" . have "card {x. poly p x = 0} = card {x. poly ([:-a, 1:] * k) x = 0}" unfolding k .. also have "{x. poly ([:-a, 1:] * k) x = 0} = insert a {x. poly k x = 0}" by auto also have "card \ \ Suc (card {x. poly k x = 0})" unfolding card_insert_if[OF poly_roots_finite[OF \k \ 0\]] by simp also have "\ \ Suc (degree k)" using le by auto finally show ?thesis using \degree p = Suc (degree k)\ by simp qed simp qed lemma poly_root_factor: "(poly ([: r, 1:] * q) (k :: 'a :: idom) = 0) = (k = -r \ poly q k = 0)" (is ?one) "(poly (q * [: r, 1:]) k = 0) = (k = -r \ poly q k = 0)" (is ?two) "(poly [: r, 1 :] k = 0) = (k = -r)" (is ?three) proof - have [simp]: "r + k = 0 \ k = - r" by (simp add: minus_unique) show ?one unfolding poly_mult by auto show ?two unfolding poly_mult by auto show ?three by auto qed lemma poly_root_constant: "c \ 0 \ (poly (p * [:c:]) (k :: 'a :: idom) = 0) = (poly p k = 0)" unfolding poly_mult by auto lemma poly_linear_exp_linear_factors_rev: "([:b,1:])^(length (filter ((=) b) as)) dvd (\ (a :: 'a :: comm_ring_1) \ as. [: a, 1:])" proof (induct as) case (Cons a as) let ?ls = "length (filter ((=) b) (a # as))" let ?l = "length (filter ((=) b) as)" have prod: "(\ a \ Cons a as. [: a, 1:]) = [: a, 1 :] * (\ a \ as. [: a, 1:])" by simp show ?case proof (cases "a = b") case False hence len: "?ls = ?l" by simp show ?thesis unfolding prod len using Cons by (rule dvd_mult) next case True hence len: "[: b, 1 :] ^ ?ls = [: a, 1 :] * [: b, 1 :] ^ ?l" by simp show ?thesis unfolding prod len using Cons using dvd_refl mult_dvd_mono by blast qed qed simp lemma order_max: assumes dvd: "[: -a, 1 :] ^ k dvd p" and p: "p \ 0" shows "k \ order a p" proof (rule ccontr) assume "\ ?thesis" hence "\ j. k = Suc (order a p + j)" by arith then obtain j where k: "k = Suc (order a p + j)" by auto have "[: -a, 1 :] ^ Suc (order a p) dvd p" by (rule power_le_dvd[OF dvd[unfolded k]], simp) with order_2[OF p, of a] show False by blast qed subsection \Divisibility\ context assumes "SORT_CONSTRAINT('a :: idom)" begin lemma poly_linear_linear_factor: assumes dvd: "[:b,1:] dvd (\ (a :: 'a) \ as. [: a, 1:])" shows "b \ set as" proof - let ?p = "\ as. (\ a \ as. [: a, 1:])" let ?b = "[:b,1:]" from assms[unfolded dvd_def] obtain p where id: "?p as = ?b * p" .. from arg_cong[OF id, of "\ p. poly p (-b)"] have "poly (?p as) (-b) = 0" by simp thus ?thesis proof (induct as) case (Cons a as) have "?p (a # as) = [:a,1:] * ?p as" by simp from Cons(2)[unfolded this] have "poly (?p as) (-b) = 0 \ (a - b) = 0" by simp with Cons(1) show ?case by auto qed simp qed lemma poly_linear_exp_linear_factors: assumes dvd: "([:b,1:])^n dvd (\ (a :: 'a) \ as. [: a, 1:])" shows "length (filter ((=) b) as) \ n" proof - let ?p = "\ as. (\ a \ as. [: a, 1:])" let ?b = "[:b,1:]" from dvd show ?thesis proof (induct n arbitrary: as) case (Suc n as) have bs: "?b ^ Suc n = ?b * ?b ^ n" by simp from poly_linear_linear_factor[OF dvd_mult_left[OF Suc(2)[unfolded bs]], unfolded in_set_conv_decomp] obtain as1 as2 where as: "as = as1 @ b # as2" by auto have "?p as = [:b,1:] * ?p (as1 @ as2)" unfolding as proof (induct as1) case (Cons a as1) have "?p (a # as1 @ b # as2) = [:a,1:] * ?p (as1 @ b # as2)" by simp also have "?p (as1 @ b # as2) = [:b,1:] * ?p (as1 @ as2)" unfolding Cons by simp also have "[:a,1:] * \ = [:b,1:] * ([:a,1:] * ?p (as1 @ as2))" by (metis (no_types, lifting) mult.left_commute) finally show ?case by simp qed simp from Suc(2)[unfolded bs this dvd_mult_cancel_left] have "?b ^ n dvd ?p (as1 @ as2)" by simp from Suc(1)[OF this] show ?case unfolding as by simp qed simp qed end lemma const_poly_dvd: "([:a:] dvd [:b:]) = (a dvd b)" proof assume "a dvd b" then obtain c where "b = a * c" unfolding dvd_def by auto hence "[:b:] = [:a:] * [: c:]" by (auto simp: ac_simps) thus "[:a:] dvd [:b:]" unfolding dvd_def by blast next assume "[:a:] dvd [:b:]" then obtain pc where "[:b:] = [:a:] * pc" unfolding dvd_def by blast from arg_cong[OF this, of "\ p. coeff p 0", unfolded coeff_mult] have "b = a * coeff pc 0" by auto thus "a dvd b" unfolding dvd_def by blast qed lemma const_poly_dvd_1 [simp]: "[:a:] dvd 1 \ a dvd 1" by (metis const_poly_dvd one_poly_eq_simps(2)) lemma poly_dvd_1: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" shows "p dvd 1 \ degree p = 0 \ coeff p 0 dvd 1" proof (cases "degree p = 0") case False with divides_degree[of p 1] show ?thesis by auto next case True from degree0_coeffs[OF this] obtain a where p: "p = [:a:]" by auto show ?thesis unfolding p by auto qed text \Degree based version of irreducibility.\ definition irreducible\<^sub>d :: "'a :: comm_semiring_1 poly \ bool" where "irreducible\<^sub>d p = (degree p > 0 \ (\ q r. degree q < degree p \ degree r < degree p \ p \ q * r))" lemma irreducible\<^sub>dI [intro]: assumes 1: "degree p > 0" and 2: "\q r. degree q > 0 \ degree q < degree p \ degree r > 0 \ degree r < degree p \ p = q * r \ False" shows "irreducible\<^sub>d p" proof (unfold irreducible\<^sub>d_def, intro conjI allI impI notI 1) fix q r assume "degree q < degree p" and "degree r < degree p" and "p = q * r" with degree_mult_le[of q r] show False by (intro 2, auto) qed lemma irreducible\<^sub>dI2: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes deg: "degree p > 0" and ndvd: "\ q. degree q > 0 \ degree q \ degree p div 2 \ \ q dvd p" shows "irreducible\<^sub>d p" proof (rule ccontr) assume "\ ?thesis" from this[unfolded irreducible\<^sub>d_def] deg obtain q r where dq: "degree q < degree p" and dr: "degree r < degree p" and p: "p = q * r" by auto from deg have p0: "p \ 0" by auto with p have "q \ 0" "r \ 0" by auto from degree_mult_eq[OF this] p have dp: "degree p = degree q + degree r" by simp show False proof (cases "degree q \ degree p div 2") case True from ndvd[OF _ True] dq dr dp p show False by auto next case False with dp have dr: "degree r \ degree p div 2" by auto from p have dvd: "r dvd p" by auto from ndvd[OF _ dr] dvd dp dq show False by auto qed qed lemma reducible\<^sub>dI: assumes "degree p > 0 \ \q r. degree q < degree p \ degree r < degree p \ p = q * r" shows "\ irreducible\<^sub>d p" using assms by (auto simp: irreducible\<^sub>d_def) lemma irreducible\<^sub>dE [elim]: assumes "irreducible\<^sub>d p" and "degree p > 0 \ (\q r. degree q < degree p \ degree r < degree p \ p \ q * r) \ thesis" shows thesis using assms by (auto simp: irreducible\<^sub>d_def) lemma reducible\<^sub>dE [elim]: assumes red: "\ irreducible\<^sub>d p" and 1: "degree p = 0 \ thesis" and 2: "\q r. degree q > 0 \ degree q < degree p \ degree r > 0 \ degree r < degree p \ p = q * r \ thesis" shows thesis using red[unfolded irreducible\<^sub>d_def de_Morgan_conj not_not not_all not_imp] proof (elim disjE exE conjE) show "\degree p > 0 \ thesis" using 1 by auto next fix q r assume "degree q < degree p" and "degree r < degree p" and "p = q * r" with degree_mult_le[of q r] show thesis by (intro 2, auto) qed lemma irreducible\<^sub>dD: assumes "irreducible\<^sub>d p" shows "degree p > 0" "\q r. degree q < degree p \ degree r < degree p \ p \ q * r" using assms unfolding irreducible\<^sub>d_def by auto theorem irreducible\<^sub>d_factorization_exists: assumes "degree p > 0" shows "\fs. fs \ [] \ (\f \ set fs. irreducible\<^sub>d f \ degree f \ degree p) \ p = prod_list fs" and "\irreducible\<^sub>d p \ \fs. length fs > 1 \ (\f \ set fs. irreducible\<^sub>d f \ degree f < degree p) \ p = prod_list fs" proof (atomize(full), insert assms, induct "degree p" arbitrary:p rule: less_induct) case less then have deg_f: "degree p > 0" by auto show ?case proof (cases "irreducible\<^sub>d p") case True then have "set [p] \ Collect irreducible\<^sub>d" "p = prod_list [p]" by auto with True show ?thesis by (auto intro: exI[of _ "[p]"]) next case False with deg_f obtain g h where deg_g: "degree g < degree p" "degree g > 0" and deg_h: "degree h < degree p" "degree h > 0" and f_gh: "p = g * h" by auto from less.hyps[OF deg_g] less.hyps[OF deg_h] obtain gs hs where emp: "length gs > 0" "length hs > 0" and "\f \ set gs. irreducible\<^sub>d f \ degree f \ degree g" "g = prod_list gs" and "\f \ set hs. irreducible\<^sub>d f \ degree f \ degree h" "h = prod_list hs" by auto with f_gh deg_g deg_h have len: "length (gs@hs) > 1" and mem: "\f \ set (gs@hs). irreducible\<^sub>d f \ degree f < degree p" and p: "p = prod_list (gs@hs)" by (auto simp del: length_greater_0_conv) with False show ?thesis by (auto intro!: exI[of _ "gs@hs"] simp: less_imp_le) qed qed lemma irreducible\<^sub>d_factor: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree p > 0" shows "\ q r. irreducible\<^sub>d q \ p = q * r \ degree r < degree p" using assms proof (induct "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "irreducible\<^sub>d p") case False with less(2) obtain q r where q: "degree q < degree p" "degree q > 0" and r: "degree r < degree p" "degree r > 0" and p: "p = q * r" by auto from less(1)[OF q] obtain s t where IH: "irreducible\<^sub>d s" "q = s * t" by auto from p have p: "p = s * (t * r)" unfolding IH by (simp add: ac_simps) from less(2) have "p \ 0" by auto hence "degree p = degree s + (degree (t * r))" unfolding p by (subst degree_mult_eq, insert p, auto) with irreducible\<^sub>dD[OF IH(1)] have "degree p > degree (t * r)" by auto with p IH show ?thesis by auto next case True show ?thesis by (rule exI[of _ p], rule exI[of _ 1], insert True less(2), auto) qed qed context mult_zero begin (* least class with times and zero *) definition zero_divisor where "zero_divisor a \ \b. b \ 0 \ a * b = 0" lemma zero_divisorI[intro]: assumes "b \ 0" and "a * b = 0" shows "zero_divisor a" using assms by (auto simp: zero_divisor_def) lemma zero_divisorE[elim]: assumes "zero_divisor a" and "\b. b \ 0 \ a * b = 0 \ thesis" shows thesis using assms by (auto simp: zero_divisor_def) end lemma zero_divisor_0[simp]: "zero_divisor (0::'a::{mult_zero,zero_neq_one})" (* No need for one! *) by (auto intro!: zero_divisorI[of 1]) lemma not_zero_divisor_1: "\ zero_divisor (1 :: 'a :: {monoid_mult,mult_zero})" (* No need for associativity! *) by auto lemma zero_divisor_iff_eq_0[simp]: fixes a :: "'a :: {semiring_no_zero_divisors, zero_neq_one}" shows "zero_divisor a \ a = 0" by auto lemma mult_eq_0_not_zero_divisor_left[simp]: fixes a b :: "'a :: mult_zero" assumes "\ zero_divisor a" shows "a * b = 0 \ b = 0" using assms unfolding zero_divisor_def by force lemma mult_eq_0_not_zero_divisor_right[simp]: fixes a b :: "'a :: {ab_semigroup_mult,mult_zero}" (* No need for associativity! *) assumes "\ zero_divisor b" shows "a * b = 0 \ a = 0" using assms unfolding zero_divisor_def by (force simp: ac_simps) lemma degree_smult_not_zero_divisor_left[simp]: assumes "\ zero_divisor c" shows "degree (smult c p) = degree p" proof(cases "p = 0") case False then have "coeff (smult c p) (degree p) \ 0" using assms by auto from le_degree[OF this] degree_smult_le[of c p] show ?thesis by auto qed auto lemma degree_smult_not_zero_divisor_right[simp]: assumes "\ zero_divisor (lead_coeff p)" shows "degree (smult c p) = (if c = 0 then 0 else degree p)" proof(cases "c = 0") case False then have "coeff (smult c p) (degree p) \ 0" using assms by auto from le_degree[OF this] degree_smult_le[of c p] show ?thesis by auto qed auto lemma irreducible\<^sub>d_smult_not_zero_divisor_left: assumes c0: "\ zero_divisor c" assumes L: "irreducible\<^sub>d (smult c p)" shows "irreducible\<^sub>d p" proof (intro irreducible\<^sub>dI) from L have "degree (smult c p) > 0" by auto also note degree_smult_le finally show "degree p > 0" by auto fix q r assume deg_q: "degree q < degree p" and deg_r: "degree r < degree p" and p_qr: "p = q * r" then have 1: "smult c p = smult c q * r" by auto note degree_smult_le[of c q] also note deg_q finally have 2: "degree (smult c q) < degree (smult c p)" using c0 by auto from deg_r have 3: "degree r < \" using c0 by auto from irreducible\<^sub>dD(2)[OF L 2 3] 1 show False by auto qed lemmas irreducible\<^sub>d_smultI = irreducible\<^sub>d_smult_not_zero_divisor_left [where 'a = "'a :: {comm_semiring_1,semiring_no_zero_divisors}", simplified] lemma irreducible\<^sub>d_smult_not_zero_divisor_right: assumes p0: "\ zero_divisor (lead_coeff p)" and L: "irreducible\<^sub>d (smult c p)" shows "irreducible\<^sub>d p" proof- from L have "c \ 0" by auto with p0 have [simp]: "degree (smult c p) = degree p" by simp show "irreducible\<^sub>d p" proof (intro iffI irreducible\<^sub>dI conjI) from L show "degree p > 0" by auto fix q r assume deg_q: "degree q < degree p" and deg_r: "degree r < degree p" and p_qr: "p = q * r" then have 1: "smult c p = smult c q * r" by auto note degree_smult_le[of c q] also note deg_q finally have 2: "degree (smult c q) < degree (smult c p)" by simp from deg_r have 3: "degree r < \" by simp from irreducible\<^sub>dD(2)[OF L 2 3] 1 show False by auto qed qed lemma zero_divisor_mult_left: fixes a b :: "'a :: {ab_semigroup_mult, mult_zero}" assumes "zero_divisor a" shows "zero_divisor (a * b)" proof- from assms obtain c where c0: "c \ 0" and [simp]: "a * c = 0" by auto have "a * b * c = a * c * b" by (simp only: ac_simps) with c0 show ?thesis by auto qed lemma zero_divisor_mult_right: fixes a b :: "'a :: {semigroup_mult, mult_zero}" assumes "zero_divisor b" shows "zero_divisor (a * b)" proof- from assms obtain c where c0: "c \ 0" and [simp]: "b * c = 0" by auto have "a * b * c = a * (b * c)" by (simp only: ac_simps) with c0 show ?thesis by auto qed lemma not_zero_divisor_mult: fixes a b :: "'a :: {ab_semigroup_mult, mult_zero}" assumes "\ zero_divisor (a * b)" shows "\ zero_divisor a" and "\ zero_divisor b" using assms by (auto dest: zero_divisor_mult_right zero_divisor_mult_left) lemma zero_divisor_smult_left: assumes "zero_divisor a" shows "zero_divisor (smult a f)" proof- from assms obtain b where b0: "b \ 0" and "a * b = 0" by auto then have "smult a f * [:b:] = 0" by (simp add: ac_simps) with b0 show ?thesis by (auto intro!: zero_divisorI[of "[:b:]"]) qed lemma unit_not_zero_divisor: fixes a :: "'a :: {comm_monoid_mult, mult_zero}" assumes "a dvd 1" shows "\zero_divisor a" proof from assms obtain b where ab: "1 = a * b" by (elim dvdE) assume "zero_divisor a" then have "zero_divisor (1::'a)" by (unfold ab, intro zero_divisor_mult_left) then show False by auto qed lemma linear_irreducible\<^sub>d: assumes "degree p = 1" shows "irreducible\<^sub>d p" by (rule irreducible\<^sub>dI, insert assms, auto) lemma irreducible\<^sub>d_dvd_smult: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree p > 0" "irreducible\<^sub>d q" "p dvd q" shows "\ c. c \ 0 \ q = smult c p" proof - from assms obtain r where q: "q = p * r" by (elim dvdE, auto) from degree_mult_eq[of p r] assms(1) q have "\ degree p < degree q" and nz: "p \ 0" "q \ 0" apply (metis assms(2) degree_mult_eq_0 gr_implies_not_zero irreducible\<^sub>dD(2) less_add_same_cancel2) using assms by auto hence deg: "degree p \ degree q" by auto from \p dvd q\ obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with nz have "k \ 0" by auto from deg[unfolded q degree_mult_eq[OF \k \ 0\ \p \ 0\ ]] have "degree k = 0" unfolding q by auto then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed subsection \Map over Polynomial Coefficients\ lemma map_poly_simps: shows "map_poly f (pCons c p) = (if c = 0 \ p = 0 then 0 else pCons (f c) (map_poly f p))" proof (cases "c = 0") case True note c0 = this show ?thesis proof (cases "p = 0") case True thus ?thesis using c0 unfolding map_poly_def by simp next case False thus ?thesis unfolding map_poly_def by auto qed next case False thus ?thesis unfolding map_poly_def by auto qed lemma map_poly_pCons[simp]: assumes "c \ 0 \ p \ 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" unfolding map_poly_simps using assms by auto lemma map_poly_map_poly: assumes f0: "f 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" proof (induct p) case (pCons a p) show ?case proof(cases "g a \ 0 \ map_poly g p \ 0") case True show ?thesis unfolding map_poly_pCons[OF pCons(1)] unfolding map_poly_pCons[OF True] unfolding pCons(2) by simp next case False then show ?thesis unfolding map_poly_pCons[OF pCons(1)] unfolding pCons(2)[symmetric] by (simp add: f0) qed qed simp lemma map_poly_zero: assumes f: "\c. f c = 0 \ c = 0" shows [simp]: "map_poly f p = 0 \ p = 0" by (induct p; auto simp: map_poly_simps f) lemma map_poly_add: assumes h0: "h 0 = 0" and h_add: "\p q. h (p + q) = h p + h q" shows "map_poly h (p + q) = map_poly h p + map_poly h q" proof (induct p arbitrary: q) case (pCons a p) note pIH = this show ?case proof(induct "q") case (pCons b q) note qIH = this show ?case unfolding map_poly_pCons[OF qIH(1)] unfolding map_poly_pCons[OF pIH(1)] unfolding add_pCons unfolding pIH(2)[symmetric] unfolding h_add[rule_format,symmetric] unfolding map_poly_simps using h0 by auto qed auto qed auto subsection \Morphismic properties of @{term "pCons 0"}\ lemma monom_pCons_0_monom: "monom (pCons 0 (monom a n)) d = map_poly (pCons 0) (monom (monom a n) d)" apply (induct d) unfolding monom_0 unfolding map_poly_simps apply simp unfolding monom_Suc map_poly_simps by auto lemma pCons_0_add: "pCons 0 (p + q) = pCons 0 p + pCons 0 q" by auto lemma sum_pCons_0_commute: "sum (\i. pCons 0 (f i)) S = pCons 0 (sum f S)" by(induct S rule: infinite_finite_induct;simp) lemma pCons_0_as_mult: fixes p:: "'a :: comm_semiring_1 poly" shows "pCons 0 p = [:0,1:] * p" by auto subsection \Misc\ fun expand_powers :: "(nat \ 'a)list \ 'a list" where "expand_powers [] = []" | "expand_powers ((Suc n, a) # ps) = a # expand_powers ((n,a) # ps)" | "expand_powers ((0,a) # ps) = expand_powers ps" lemma expand_powers: fixes f :: "'a \ 'b :: comm_ring_1" shows "(\ (n,a) \ n_as. f a ^ n) = (\ a \ expand_powers n_as. f a)" by (rule sym, induct n_as rule: expand_powers.induct, auto) lemma poly_smult_zero_iff: fixes x :: "'a :: idom" shows "(poly (smult a p) x = 0) = (a = 0 \ poly p x = 0)" by simp lemma poly_prod_list_zero_iff: fixes x :: "'a :: idom" shows "(poly (prod_list ps) x = 0) = (\ p \ set ps. poly p x = 0)" by (induct ps, auto) lemma poly_mult_zero_iff: fixes x :: "'a :: idom" shows "(poly (p * q) x = 0) = (poly p x = 0 \ poly q x = 0)" by simp lemma poly_power_zero_iff: fixes x :: "'a :: idom" shows "(poly (p^n) x = 0) = (n \ 0 \ poly p x = 0)" by (cases n, auto) lemma sum_monom_0_iff: assumes fin: "finite S" and g: "\ i j. g i = g j \ i = j" shows "sum (\ i. monom (f i) (g i)) S = 0 \ (\ i \ S. f i = 0)" (is "?l = ?r") proof - { assume "\ ?r" then obtain i where i: "i \ S" and fi: "f i \ 0" by auto let ?g = "\ i. monom (f i) (g i)" have "coeff (sum ?g S) (g i) = f i + sum (\ j. coeff (?g j) (g i)) (S - {i})" by (unfold sum.remove[OF fin i], simp add: coeff_sum) also have "sum (\ j. coeff (?g j) (g i)) (S - {i}) = 0" by (rule sum.neutral, insert g, auto) finally have "coeff (sum ?g S) (g i) \ 0" using fi by auto hence "\ ?l" by auto } thus ?thesis by auto qed lemma degree_prod_list_eq: assumes "\ p. p \ set ps \ (p :: 'a :: idom poly) \ 0" shows "degree (prod_list ps) = sum_list (map degree ps)" using assms proof (induct ps) case (Cons p ps) show ?case unfolding prod_list.Cons by (subst degree_mult_eq, insert Cons, auto simp: prod_list_zero_iff) qed simp lemma degree_power_eq: assumes p: "p \ 0" shows "degree (p ^ n) = degree (p :: 'a :: idom poly) * n" proof (induct n) case (Suc n) from p have pn: "p ^ n \ 0" by auto show ?case using degree_mult_eq[OF p pn] Suc by auto qed simp lemma coeff_Poly: "coeff (Poly xs) i = (nth_default 0 xs i)" unfolding nth_default_coeffs_eq[of "Poly xs", symmetric] coeffs_Poly by simp lemma rsquarefree_def': "rsquarefree p = (p \ 0 \ (\a. order a p \ 1))" proof - have "\ a. order a p \ 1 \ order a p = 0 \ order a p = 1" by linarith thus ?thesis unfolding rsquarefree_def by auto qed lemma order_prod_list: "(\ p. p \ set ps \ p \ 0) \ order x (prod_list ps) = sum_list (map (order x) ps)" by (induct ps, auto, subst order_mult, auto simp: prod_list_zero_iff) lemma irreducible\<^sub>d_dvd_eq: fixes a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "irreducible\<^sub>d a" and "irreducible\<^sub>d b" and "a dvd b" and "monic a" and "monic b" shows "a = b" using assms by (metis (no_types, lifting) coeff_smult degree_smult_eq irreducible\<^sub>dD(1) irreducible\<^sub>d_dvd_smult mult.right_neutral smult_1_left) lemma monic_gcd_dvd: assumes fg: "f dvd g" and mon: "monic f" and gcd: "gcd g h \ {1, g}" shows "gcd f h \ {1, f}" proof (cases "coprime g h") case True with dvd_refl have "coprime f h" using fg by (blast intro: coprime_divisors) then show ?thesis by simp next case False with gcd have gcd: "gcd g h = g" by (simp add: coprime_iff_gcd_eq_1) with fg have "f dvd gcd g h" by simp then have "f dvd h" by simp then have "gcd f h = normalize f" by (simp add: gcd_proj1_iff) also have "normalize f = f" using mon by (rule normalize_monic) finally show ?thesis by simp qed lemma monom_power: "(monom a b)^n = monom (a^n) (b*n)" by (induct n, auto simp add: mult_monom) lemma poly_const_pow: "[:a:]^b = [:a^b:]" by (metis Groups.mult_ac(2) monom_0 monom_power mult_zero_right) lemma degree_pderiv_le: "degree (pderiv f) \ degree f - 1" proof (rule ccontr) assume "\ ?thesis" hence ge: "degree (pderiv f) \ Suc (degree f - 1)" by auto hence "pderiv f \ 0" by auto hence "coeff (pderiv f) (degree (pderiv f)) \ 0" by auto from this[unfolded coeff_pderiv] have "coeff f (Suc (degree (pderiv f))) \ 0" by auto moreover have "Suc (degree (pderiv f)) > degree f" using ge by auto ultimately show False by (simp add: coeff_eq_0) qed lemma map_div_is_smult_inverse: "map_poly (\x. x / (a :: 'a :: field)) p = smult (inverse a) p" unfolding smult_conv_map_poly by (simp add: divide_inverse_commute) lemma normalize_poly_old_def: "normalize (f :: 'a :: {normalization_semidom,field} poly) = smult (inverse (unit_factor (lead_coeff f))) f" by (simp add: normalize_poly_eq_map_poly map_div_is_smult_inverse) (* was in Euclidean_Algorithm in Number_Theory before, but has been removed *) lemma poly_dvd_antisym: fixes p q :: "'b::idom poly" assumes coeff: "coeff p (degree p) = coeff q (degree q)" assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" proof (cases "p = 0") case True with coeff show "p = q" by simp next case False with coeff have "q \ 0" by auto have degree: "degree p = degree q" using \p dvd q\ \q dvd p\ \p \ 0\ \q \ 0\ by (intro order_antisym dvd_imp_degree_le) from \p dvd q\ obtain a where a: "q = p * a" .. with \q \ 0\ have "a \ 0" by auto with degree a \p \ 0\ have "degree a = 0" by (simp add: degree_mult_eq) with coeff a show "p = q" by (cases a, auto split: if_splits) qed lemma coeff_f_0_code[code_unfold]: "coeff f 0 = (case coeffs f of [] \ 0 | x # _ \ x)" by (cases f, auto simp: cCons_def) lemma poly_compare_0_code[code_unfold]: "(f = 0) = (case coeffs f of [] \ True | _ \ False)" using coeffs_eq_Nil list.disc_eq_case(1) by blast text \Getting more efficient code for abbreviation @{term lead_coeff}"\ definition leading_coeff where [code_abbrev, simp]: "leading_coeff = lead_coeff" lemma leading_coeff_code [code]: "leading_coeff f = (let xs = coeffs f in if xs = [] then 0 else last xs)" by (simp add: last_coeffs_eq_coeff_degree) lemma nth_coeffs_coeff: "i < length (coeffs f) \ coeffs f ! i = coeff f i" by (metis nth_default_coeffs_eq nth_default_def) lemma degree_prod_eq_sum_degree: fixes A :: "'a set" and f :: "'a \ 'b::field poly" assumes f0: "\i\A. f i \ 0" shows "degree (\i\A. (f i)) = (\i\A. degree (f i))" using f0 proof (induct A rule: infinite_finite_induct) case (insert x A) have "(\i\insert x A. degree (f i)) = degree (f x) + (\i\A. degree (f i))" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... = degree (f x) + degree (\i\A. (f i))" by (simp add: insert.hyps insert.prems) also have "... = degree (f x * (\i\A. (f i)))" proof (rule degree_mult_eq[symmetric]) show "f x \ 0" using insert.prems by auto show "prod f A \ 0" by (simp add: insert.hyps(1) insert.prems) qed also have "... = degree (\i\insert x A. (f i))" by (simp add: insert.hyps) finally show ?case .. qed auto definition monom_mult :: "nat \ 'a :: comm_semiring_1 poly \ 'a poly" where "monom_mult n f = monom 1 n * f" lemma monom_mult_unfold [code_unfold]: "monom 1 n * f = monom_mult n f" "f * monom 1 n = monom_mult n f" by (auto simp: monom_mult_def ac_simps) lemma monom_mult_code [code abstract]: "coeffs (monom_mult n f) = (let xs = coeffs f in if xs = [] then xs else replicate n 0 @ xs)" by (rule coeffs_eqI) (auto simp add: Let_def monom_mult_def coeff_monom_mult nth_default_append nth_default_coeffs_eq) lemma coeff_pcompose_monom: fixes f :: "'a :: comm_ring_1 poly" assumes n: "j < n" shows "coeff (f \\<^sub>p monom 1 n) (n * i + j) = (if j = 0 then coeff f i else 0)" proof (induct f arbitrary: i) case (pCons a f i) note d = pcompose_pCons coeff_add coeff_monom_mult coeff_pCons show ?case proof (cases i) case 0 show ?thesis unfolding d 0 using n by (cases j, auto) next case (Suc ii) have id: "n * Suc ii + j - n = n * ii + j" using n by (simp add: diff_mult_distrib2) have id1: "(n \ n * Suc ii + j) = True" by auto have id2: "(case n * Suc ii + j of 0 \ a | Suc x \ coeff 0 x) = 0" using n by (cases "n * Suc ii + j", auto) show ?thesis unfolding d Suc id id1 id2 pCons(2) if_True by auto qed qed auto lemma coeff_pcompose_x_pow_n: fixes f :: "'a :: comm_ring_1 poly" assumes n: "n \ 0" shows "coeff (f \\<^sub>p monom 1 n) (n * i) = coeff f i" using coeff_pcompose_monom[of 0 n f i] n by auto lemma dvd_dvd_smult: "a dvd b \ f dvd g \ smult a f dvd smult b g" unfolding dvd_def by (metis mult_smult_left mult_smult_right smult_smult) definition sdiv_poly :: "'a :: idom_divide poly \ 'a \ 'a poly" where "sdiv_poly p a = (map_poly (\ c. c div a) p)" lemma smult_map_poly: "smult a = map_poly ((*) a)" by (rule ext, rule poly_eqI, subst coeff_map_poly, auto) lemma smult_exact_sdiv_poly: assumes "\ c. c \ set (coeffs p) \ a dvd c" shows "smult a (sdiv_poly p a) = p" unfolding smult_map_poly sdiv_poly_def by (subst map_poly_map_poly,simp,rule map_poly_idI, insert assms, auto) lemma coeff_sdiv_poly: "coeff (sdiv_poly f a) n = coeff f n div a" unfolding sdiv_poly_def by (rule coeff_map_poly, auto) lemma poly_pinfty_ge: fixes p :: "real poly" assumes "lead_coeff p > 0" "degree p \ 0" shows "\n. \ x \ n. poly p x \ b" proof - let ?p = "p - [:b - lead_coeff p :]" have id: "lead_coeff ?p = lead_coeff p" using assms(2) by (cases p, auto) with assms(1) have "lead_coeff ?p > 0" by auto from poly_pinfty_gt_lc[OF this, unfolded id] obtain n where "\ x. x \ n \ 0 \ poly p x - b" by auto thus ?thesis by auto qed lemma pderiv_sum: "pderiv (sum f I) = sum (\ i. (pderiv (f i))) I" by (induct I rule: infinite_finite_induct, auto simp: pderiv_add) lemma smult_sum2: "smult m (\i \ S. f i) = (\i \ S. smult m (f i))" by (induct S rule: infinite_finite_induct, auto simp add: smult_add_right) lemma degree_mult_not_eq: "degree (f * g) \ degree f + degree g \ lead_coeff f * lead_coeff g = 0" by (rule ccontr, auto simp: coeff_mult_degree_sum degree_mult_le le_antisym le_degree) lemma irreducible\<^sub>d_multD: fixes a b :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes l: "irreducible\<^sub>d (a*b)" shows "degree a = 0 \ a \ 0 \ irreducible\<^sub>d b \ degree b = 0 \ b \ 0 \ irreducible\<^sub>d a" proof- from l have a0: "a \ 0" and b0: "b \ 0" by auto note [simp] = degree_mult_eq[OF this] from l have "degree a = 0 \ degree b = 0" apply (unfold irreducible\<^sub>d_def) by force then show ?thesis proof(elim disjE) assume a: "degree a = 0" with l a0 have "irreducible\<^sub>d b" - unfolding irreducible\<^sub>d_def - by (smt add.left_neutral degree_mult_not_eq leading_coeff_0_iff sign_simps(4) mult_eq_0_iff) + by (simp add: irreducible\<^sub>d_def) + (metis degree_mult_eq degree_mult_eq_0 mult.left_commute plus_nat.add_0) with a a0 show ?thesis by auto next assume b: "degree b = 0" with l b0 have "irreducible\<^sub>d a" unfolding irreducible\<^sub>d_def by (smt add_cancel_left_right degree_mult_eq degree_mult_eq_0 neq0_conv semiring_normalization_rules(16)) with b b0 show ?thesis by auto qed qed lemma irreducible_connect_field[simp]: fixes f :: "'a :: field poly" shows "irreducible\<^sub>d f = irreducible f" (is "?l = ?r") proof show "?r \ ?l" apply (intro irreducible\<^sub>dI, force simp:is_unit_iff_degree) by (auto dest!: irreducible_multD simp: poly_dvd_1) next assume l: ?l show ?r proof (rule irreducibleI) from l show "f \ 0" "\ is_unit f" by (auto simp: poly_dvd_1) fix a b assume "f = a * b" from l[unfolded this] show "a dvd 1 \ b dvd 1" by (auto dest!: irreducible\<^sub>d_multD simp:is_unit_iff_degree) qed qed lemma is_unit_field_poly[simp]: fixes p :: "'a::field poly" shows "is_unit p \ p \ 0 \ degree p = 0" by (cases "p=0", auto simp: is_unit_iff_degree) lemma irreducible_smult_field[simp]: fixes c :: "'a :: field" shows "irreducible (smult c p) \ c \ 0 \ irreducible p" (is "?L \ ?R") proof (intro iffI conjI irreducible\<^sub>d_smult_not_zero_divisor_left[of c p, simplified]) assume "irreducible (smult c p)" then show "c \ 0" by auto next assume ?R then have c0: "c \ 0" and irr: "irreducible p" by auto show ?L proof (fold irreducible_connect_field, intro irreducible\<^sub>dI, unfold degree_smult_eq if_not_P[OF c0]) show "degree p > 0" using irr by auto fix q r from c0 have "p = smult (1/c) (smult c p)" by simp also assume "smult c p = q * r" finally have [simp]: "p = smult (1/c) \". assume main: "degree q < degree p" "degree r < degree p" have "\irreducible\<^sub>d p" by (rule reducible\<^sub>dI, rule exI[of _ "smult (1/c) q"], rule exI[of _ r], insert irr c0 main, simp) with irr show False by auto qed qed auto lemma irreducible_monic_factor: fixes p :: "'a :: field poly" assumes "degree p > 0" shows "\ q r. irreducible q \ p = q * r \ monic q" proof - from irreducible\<^sub>d_factorization_exists[OF assms] obtain fs where "fs \ []" and "set fs \ Collect irreducible" and "p = prod_list fs" by auto then have q: "irreducible (hd fs)" and p: "p = hd fs * prod_list (tl fs)" by (atomize(full), cases fs, auto) define c where "c = coeff (hd fs) (degree (hd fs))" from q have c: "c \ 0" unfolding c_def irreducible\<^sub>d_def by auto show ?thesis by (rule exI[of _ "smult (1/c) (hd fs)"], rule exI[of _ "smult c (prod_list (tl fs))"], unfold p, insert q c, auto simp: c_def) qed lemma monic_irreducible_factorization: fixes p :: "'a :: field poly" shows "monic p \ \ as f. finite as \ p = prod (\ a. a ^ Suc (f a)) as \ as \ {q. irreducible q \ monic q}" proof (induct "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "degree p > 0") case False with less(2) have "p = 1" by (simp add: coeff_eq_0 poly_eq_iff) thus ?thesis by (intro exI[of _ "{}"], auto) next case True from irreducible\<^sub>d_factor[OF this] obtain q r where p: "p = q * r" and q: "irreducible q" and deg: "degree r < degree p" by auto hence q0: "q \ 0" by auto define c where "c = coeff q (degree q)" let ?q = "smult (1/c) q" let ?r = "smult c r" from q0 have c: "c \ 0" "1 / c \ 0" unfolding c_def by auto hence p: "p = ?q * ?r" unfolding p by auto have deg: "degree ?r < degree p" using c deg by auto let ?Q = "{q. irreducible q \ monic (q :: 'a poly)}" have mon: "monic ?q" unfolding c_def using q0 by auto from monic_factor[OF \monic p\[unfolded p] this] have "monic ?r" . from less(1)[OF deg this] obtain f as where as: "finite as" "?r = (\ a \as. a ^ Suc (f a))" "as \ ?Q" by blast from q c have irred: "irreducible ?q" by simp show ?thesis proof (cases "?q \ as") case False let ?as = "insert ?q as" let ?f = "\ a. if a = ?q then 0 else f a" have "p = ?q * (\ a \as. a ^ Suc (f a))" unfolding p as by simp also have "(\ a \as. a ^ Suc (f a)) = (\ a \as. a ^ Suc (?f a))" by (rule prod.cong, insert False, auto) also have "?q * \ = (\ a \ ?as. a ^ Suc (?f a))" by (subst prod.insert, insert as False, auto) finally have p: "p = (\ a \ ?as. a ^ Suc (?f a))" . from as(1) have fin: "finite ?as" by auto from as mon irred have Q: "?as \ ?Q" by auto from fin p Q show ?thesis by(intro exI[of _ ?as] exI[of _ ?f], auto) next case True let ?f = "\ a. if a = ?q then Suc (f a) else f a" have "p = ?q * (\ a \as. a ^ Suc (f a))" unfolding p as by simp also have "(\ a \as. a ^ Suc (f a)) = ?q ^ Suc (f ?q) * (\ a \(as - {?q}). a ^ Suc (f a))" by (subst prod.remove[OF _ True], insert as, auto) also have "(\ a \(as - {?q}). a ^ Suc (f a)) = (\ a \(as - {?q}). a ^ Suc (?f a))" by (rule prod.cong, auto) also have "?q * (?q ^ Suc (f ?q) * \ ) = ?q ^ Suc (?f ?q) * \" by (simp add: ac_simps) also have "\ = (\ a \ as. a ^ Suc (?f a))" by (subst prod.remove[OF _ True], insert as, auto) finally have "p = (\ a \ as. a ^ Suc (?f a))" . with as show ?thesis by (intro exI[of _ as] exI[of _ ?f], auto) qed qed qed lemma monic_irreducible_gcd: "monic (f::'a::{field,euclidean_ring_gcd} poly) \ irreducible f \ gcd f u \ {1,f}" by (metis gcd_dvd1 irreducible_altdef insertCI is_unit_gcd_iff poly_dvd_antisym poly_gcd_monic) end diff --git a/thys/Polynomial_Interpolation/Missing_Unsorted.thy b/thys/Polynomial_Interpolation/Missing_Unsorted.thy --- a/thys/Polynomial_Interpolation/Missing_Unsorted.thy +++ b/thys/Polynomial_Interpolation/Missing_Unsorted.thy @@ -1,651 +1,652 @@ (* Author: René Thiemann Akihisa Yamada Jose Divason License: BSD *) section \Missing Unsorted\ text \This theory contains several lemmas which might be of interest to the Isabelle distribution. For instance, we prove that $b^n \cdot n^k$ is bounded by a constant whenever $0 < b < 1$.\ theory Missing_Unsorted imports HOL.Complex "HOL-Computational_Algebra.Factorial_Ring" begin lemma bernoulli_inequality: assumes x: "-1 \ (x :: 'a :: linordered_field)" shows "1 + of_nat n * x \ (1 + x) ^ n" proof (induct n) case (Suc n) have "1 + of_nat (Suc n) * x = 1 + x + of_nat n * x" by (simp add: field_simps) also have "\ \ \ + of_nat n * x ^ 2" by simp also have "\ = (1 + of_nat n * x) * (1 + x)" by (simp add: field_simps power2_eq_square) also have "\ \ (1 + x) ^ n * (1 + x)" by (rule mult_right_mono[OF Suc], insert x, auto) also have "\ = (1 + x) ^ (Suc n)" by simp finally show ?case . qed simp context fixes b :: "'a :: archimedean_field" assumes b: "0 < b" "b < 1" begin private lemma pow_one: "b ^ x \ 1" using power_Suc_less_one[OF b, of "x - 1"] by (cases x, auto) private lemma pow_zero: "0 < b ^ x" using b(1) by simp lemma exp_tends_to_zero: assumes c: "c > 0" shows "\ x. b ^ x \ c" proof (rule ccontr) assume not: "\ ?thesis" define bb where "bb = inverse b" define cc where "cc = inverse c" from b have bb: "bb > 1" unfolding bb_def by (rule one_less_inverse) from c have cc: "cc > 0" unfolding cc_def by simp define bbb where "bbb = bb - 1" have id: "bb = 1 + bbb" and bbb: "bbb > 0" and bm1: "bbb \ -1" unfolding bbb_def using bb by auto have "\ n. cc / bbb < of_nat n" by (rule reals_Archimedean2) then obtain n where lt: "cc / bbb < of_nat n" by auto from not have "\ b ^ n \ c" by auto hence bnc: "b ^ n > c" by simp have "bb ^ n = inverse (b ^ n)" unfolding bb_def by (rule power_inverse) also have "\ < cc" unfolding cc_def by (rule less_imp_inverse_less[OF bnc c]) also have "\ < bbb * of_nat n" using lt bbb by (metis mult.commute pos_divide_less_eq) also have "\ \ bb ^ n" using bernoulli_inequality[OF bm1, folded id, of n] by (simp add: ac_simps) finally show False by simp qed lemma linear_exp_bound: "\ p. \ x. b ^ x * of_nat x \ p" proof - from b have "1 - b > 0" by simp from exp_tends_to_zero[OF this] obtain x0 where x0: "b ^ x0 \ 1 - b" .. { fix x assume "x \ x0" hence "\ y. x = x0 + y" by arith then obtain y where x: "x = x0 + y" by auto have "b ^ x = b ^ x0 * b ^ y" unfolding x by (simp add: power_add) also have "\ \ b ^ x0" using pow_one[of y] pow_zero[of x0] by auto also have "\ \ 1 - b" by (rule x0) finally have "b ^ x \ 1 - b" . } note x0 = this define bs where "bs = insert 1 { b ^ Suc x * of_nat (Suc x) | x . x \ x0}" have bs: "finite bs" unfolding bs_def by auto define p where "p = Max bs" have bs: "\ b. b \ bs \ b \ p" unfolding p_def using bs by simp hence p1: "p \ 1" unfolding bs_def by auto show ?thesis proof (rule exI[of _ p], intro allI) fix x show "b ^ x * of_nat x \ p" proof (induct x) case (Suc x) show ?case proof (cases "x \ x0") case True show ?thesis by (rule bs, unfold bs_def, insert True, auto) next case False let ?x = "of_nat x :: 'a" have "b ^ (Suc x) * of_nat (Suc x) = b * (b ^ x * ?x) + b ^ Suc x" by (simp add: field_simps) also have "\ \ b * p + b ^ Suc x" by (rule add_right_mono[OF mult_left_mono[OF Suc]], insert b, auto) also have "\ = p - ((1 - b) * p - b ^ (Suc x))" by (simp add: field_simps) also have "\ \ p - 0" proof - have "b ^ Suc x \ 1 - b" using x0[of "Suc x"] False by auto also have "\ \ (1 - b) * p" using b p1 by auto finally show ?thesis by (intro diff_left_mono, simp) qed finally show ?thesis by simp qed qed (insert p1, auto) qed qed lemma poly_exp_bound: "\ p. \ x. b ^ x * of_nat x ^ deg \ p" proof - show ?thesis proof (induct deg) case 0 show ?case by (rule exI[of _ 1], intro allI, insert pow_one, auto) next case (Suc deg) then obtain q where IH: "\ x. b ^ x * (of_nat x) ^ deg \ q" by auto define p where "p = max 0 q" from IH have IH: "\ x. b ^ x * (of_nat x) ^ deg \ p" unfolding p_def using le_max_iff_disj by blast have p: "p \ 0" unfolding p_def by simp show ?case proof (cases "deg = 0") case True thus ?thesis using linear_exp_bound by simp next case False note deg = this define p' where "p' = p*p * 2 ^ Suc deg * inverse b" let ?f = "\ x. b ^ x * (of_nat x) ^ Suc deg" define f where "f = ?f" { fix x let ?x = "of_nat x :: 'a" have "f (2 * x) \ (2 ^ Suc deg) * (p * p)" proof (cases "x = 0") case False hence x1: "?x \ 1" by (cases x, auto) from x1 have x: "?x ^ (deg - 1) \ 1" by simp from x1 have xx: "?x ^ Suc deg \ 1" by (rule one_le_power) define c where "c = b ^ x * b ^ x * (2 ^ Suc deg)" have c: "c > 0" unfolding c_def using b by auto have "f (2 * x) = ?f (2 * x)" unfolding f_def by simp also have "b ^ (2 * x) = (b ^ x) * (b ^ x)" by (simp add: power2_eq_square power_even_eq) also have "of_nat (2 * x) = 2 * ?x" by simp also have "(2 * ?x) ^ Suc deg = 2 ^ Suc deg * ?x ^ Suc deg" by simp finally have "f (2 * x) = c * ?x ^ Suc deg" unfolding c_def by (simp add: ac_simps) also have "\ \ c * ?x ^ Suc deg * ?x ^ (deg - 1)" proof - have "c * ?x ^ Suc deg > 0" using c xx by simp thus ?thesis unfolding mult_le_cancel_left1 using x by simp qed also have "\ = c * ?x ^ (Suc deg + (deg - 1))" by (simp add: power_add) also have "Suc deg + (deg - 1) = deg + deg" using deg by simp also have "?x ^ (deg + deg) = (?x ^ deg) * (?x ^ deg)" by (simp add: power_add) also have "c * \ = (2 ^ Suc deg) * ((b ^ x * ?x ^ deg) * (b ^ x * ?x ^ deg))" unfolding c_def by (simp add: ac_simps) also have "\ \ (2 ^ Suc deg) * (p * p)" by (rule mult_left_mono[OF mult_mono[OF IH IH p]], insert pow_zero[of x], auto) finally show "f (2 * x) \ (2 ^ Suc deg) * (p * p)" . qed (auto simp: f_def) hence "?f (2 * x) \ (2 ^ Suc deg) * (p * p)" unfolding f_def . } note even = this show ?thesis proof (rule exI[of _ p'], intro allI) fix y show "?f y \ p'" proof (cases "even y") case True define x where "x = y div 2" have "y = 2 * x" unfolding x_def using True by simp from even[of x, folded this] have "?f y \ 2 ^ Suc deg * (p * p)" . also have "\ \ \ * inverse b" - unfolding mult_le_cancel_left1 using b p by (simp add:sign_simps one_le_inverse) + unfolding mult_le_cancel_left1 using b p + by (simp add: sign_simps one_le_inverse) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . next case False define x where "x = y div 2" have "y = 2 * x + 1" unfolding x_def using False by simp hence "?f y = ?f (2 * x + 1)" by simp also have "\ \ b ^ (2 * x + 1) * of_nat (2 * x + 2) ^ Suc deg" by (rule mult_left_mono[OF power_mono], insert b, auto) also have "b ^ (2 * x + 1) = b ^ (2 * x + 2) * inverse b" using b by auto also have "b ^ (2 * x + 2) * inverse b * of_nat (2 * x + 2) ^ Suc deg = inverse b * ?f (2 * (x + 1))" by (simp add: ac_simps) also have "\ \ inverse b * ((2 ^ Suc deg) * (p * p))" by (rule mult_left_mono[OF even], insert b, auto) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . qed qed qed qed qed end lemma prod_list_replicate[simp]: "prod_list (replicate n a) = a ^ n" by (induct n, auto) lemma prod_list_power: fixes xs :: "'a :: comm_monoid_mult list" shows "prod_list xs ^ n = (\x\xs. x ^ n)" by (induct xs, auto simp: power_mult_distrib) lemma set_upt_Suc: "{0 ..< Suc i} = insert i {0 ..< i}" by (fact atLeast0_lessThan_Suc) lemma prod_pow[simp]: "(\i = 0..a\ * y dvd x \ a * y dvd x" for x y a :: int using abs_dvd_iff [of "a * y"] abs_dvd_iff [of "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_right_int [simp]: "gcd x (\a\ * y) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * y"] gcd_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma lcm_abs_mult_right_int [simp]: "lcm x (\a\ * y) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * y"] lcm_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_left_int [simp]: "gcd x (a * \y\) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * \y\"] gcd_abs2_int [of _ "a * y"] by (simp add: abs_mult) lemma lcm_abs_mult_left_int [simp]: "lcm x (a * \y\) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * \y\"] lcm_abs2_int [of _ "a * y"] by (simp add: abs_mult) abbreviation (input) list_gcd :: "'a :: semiring_gcd list \ 'a" where "list_gcd \ gcd_list" abbreviation (input) list_lcm :: "'a :: semiring_gcd list \ 'a" where "list_lcm \ lcm_list" lemma list_gcd_simps: "list_gcd [] = 0" "list_gcd (x # xs) = gcd x (list_gcd xs)" by simp_all lemma list_gcd: "x \ set xs \ list_gcd xs dvd x" by (fact Gcd_fin_dvd) lemma list_gcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd (list_gcd xs)" by (fact gcd_list_greatest) lemma list_gcd_mult_int [simp]: fixes xs :: "int list" shows "list_gcd (map (times a) xs) = \a\ * list_gcd xs" by (simp add: Gcd_mult) lemma list_lcm_simps: "list_lcm [] = 1" "list_lcm (x # xs) = lcm x (list_lcm xs)" by simp_all lemma list_lcm: "x \ set xs \ x dvd list_lcm xs" by (fact dvd_Lcm_fin) lemma list_lcm_least: "(\ x. x \ set xs \ x dvd y) \ list_lcm xs dvd y" by (fact lcm_list_least) lemma lcm_mult_distrib_nat: "(k :: nat) * lcm m n = lcm (k * m) (k * n)" by (simp add: lcm_mult_left) lemma lcm_mult_distrib_int: "abs (k::int) * lcm m n = lcm (k * m) (k * n)" by (simp add: lcm_mult_left) lemma list_lcm_mult_int [simp]: fixes xs :: "int list" shows "list_lcm (map (times a) xs) = (if xs = [] then 1 else \a\ * list_lcm xs)" by (simp add: Lcm_mult) lemma list_lcm_pos: "list_lcm xs \ (0 :: int)" "0 \ set xs \ list_lcm xs \ 0" "0 \ set xs \ list_lcm xs > 0" proof - have "0 \ \Lcm (set xs)\" by (simp only: abs_ge_zero) then have "0 \ Lcm (set xs)" by simp then show "list_lcm xs \ 0" by simp assume "0 \ set xs" then show "list_lcm xs \ 0" by (simp add: Lcm_0_iff) with \list_lcm xs \ 0\ show "list_lcm xs > 0" by (simp add: le_less) qed lemma quotient_of_nonzero: "snd (quotient_of r) > 0" "snd (quotient_of r) \ 0" using quotient_of_denom_pos' [of r] by simp_all lemma quotient_of_int_div: assumes q: "quotient_of (of_int x / of_int y) = (a, b)" and y: "y \ 0" shows "\ z. z \ 0 \ x = a * z \ y = b * z" proof - let ?r = "rat_of_int" define z where "z = gcd x y" define x' where "x' = x div z" define y' where "y' = y div z" have id: "x = z * x'" "y = z * y'" unfolding x'_def y'_def z_def by auto from y have y': "y' \ 0" unfolding id by auto have z: "z \ 0" unfolding z_def using y by auto have cop: "coprime x' y'" unfolding x'_def y'_def z_def using div_gcd_coprime y by blast have "?r x / ?r y = ?r x' / ?r y'" unfolding id using z y y' by (auto simp: field_simps) from assms[unfolded this] have quot: "quotient_of (?r x' / ?r y') = (a, b)" by auto from quotient_of_coprime[OF quot] have cop': "coprime a b" . hence cop: "coprime b a" by (simp add: ac_simps) from quotient_of_denom_pos[OF quot] have b: "b > 0" "b \ 0" by auto from quotient_of_div[OF quot] quotient_of_denom_pos[OF quot] y' have "?r x' * ?r b = ?r a * ?r y'" by (auto simp: field_simps) hence id': "x' * b = a * y'" unfolding of_int_mult[symmetric] by linarith from id'[symmetric] have "b dvd y' * a" unfolding mult.commute[of y'] by auto with cop y' have "b dvd y'" by (simp add: coprime_dvd_mult_left_iff) then obtain z' where ybz: "y' = b * z'" unfolding dvd_def by auto from id[unfolded y' this] have y: "y = b * (z * z')" by auto with \y \ 0\ have zz: "z * z' \ 0" by auto from quotient_of_div[OF q] \y \ 0\ \b \ 0\ have "?r x * ?r b = ?r y * ?r a" by (auto simp: field_simps) hence id': "x * b = y * a" unfolding of_int_mult[symmetric] by linarith from this[unfolded y] b have x: "x = a * (z * z')" by auto show ?thesis unfolding x y using zz by blast qed fun max_list_non_empty :: "('a :: linorder) list \ 'a" where "max_list_non_empty [x] = x" | "max_list_non_empty (x # xs) = max x (max_list_non_empty xs)" lemma max_list_non_empty: "x \ set xs \ x \ max_list_non_empty xs" proof (induct xs) case (Cons y ys) note oCons = this show ?case proof (cases ys) case (Cons z zs) hence id: "max_list_non_empty (y # ys) = max y (max_list_non_empty ys)" by simp from oCons show ?thesis unfolding id by (auto simp: max.coboundedI2) qed (insert oCons, auto) qed simp lemma cnj_reals[simp]: "(cnj c \ \) = (c \ \)" using Reals_cnj_iff by fastforce lemma sgn_real_mono: "x \ y \ sgn x \ sgn (y :: real)" unfolding sgn_real_def by (auto split: if_splits) lemma sgn_minus_rat: "sgn (- (x :: rat)) = - sgn x" by (fact Rings.sgn_minus) lemma real_of_rat_sgn: "sgn (of_rat x) = real_of_rat (sgn x)" unfolding sgn_real_def sgn_rat_def by auto lemma inverse_le_iff_sgn: assumes sgn: "sgn x = sgn y" shows "(inverse (x :: real) \ inverse y) = (y \ x)" proof (cases "x = 0") case True with sgn have "sgn y = 0" by simp hence "y = 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0"; auto) thus ?thesis using True by simp next case False note x = this show ?thesis proof (cases "x < 0") case True with x sgn have "sgn y = -1" by simp hence "y < 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le_neg[OF True \y < 0\]) next case False with x have x: "x > 0" by auto with sgn have "sgn y = 1" by auto hence "y > 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le[OF x \y > 0\]) qed qed lemma inverse_le_sgn: assumes sgn: "sgn x = sgn y" and xy: "x \ (y :: real)" shows "inverse y \ inverse x" using xy inverse_le_iff_sgn[OF sgn] by auto lemma set_list_update: "set (xs [i := k]) = (if i < length xs then insert k (set (take i xs) \ set (drop (Suc i) xs)) else set xs)" proof (induct xs arbitrary: i) case (Cons x xs i) thus ?case by (cases i, auto) qed simp lemma prod_list_dvd: assumes "(x :: 'a :: comm_monoid_mult) \ set xs" shows "x dvd prod_list xs" proof - from assms[unfolded in_set_conv_decomp] obtain ys zs where xs: "xs = ys @ x # zs" by auto show ?thesis unfolding xs dvd_def by (intro exI[of _ "prod_list (ys @ zs)"], simp add: ac_simps) qed lemma dvd_prod: fixes A::"'b set" assumes "\b\A. a dvd f b" "finite A" shows "a dvd prod f A" using assms(2,1) proof (induct A) case (insert x A) thus ?case using comm_monoid_mult_class.dvd_mult dvd_mult2 insert_iff prod.insert by auto qed auto context fixes xs :: "'a :: comm_monoid_mult list" begin lemma prod_list_filter: "prod_list (filter f xs) * prod_list (filter (\ x. \ f x) xs) = prod_list xs" by (induct xs, auto simp: ac_simps) lemma prod_list_partition: assumes "partition f xs = (ys, zs)" shows "prod_list xs = prod_list ys * prod_list zs" using assms by (subst prod_list_filter[symmetric, of f], auto simp: o_def) end lemma dvd_imp_mult_div_cancel_left[simp]: assumes "(a :: 'a :: semidom_divide) dvd b" shows "a * (b div a) = b" proof(cases "b = 0") case True then show ?thesis by auto next case False with dvdE[OF assms] obtain c where *: "b = a * c" by auto also with False have "a \ 0" by auto then have "a * c div a = c" by auto also note *[symmetric] finally show ?thesis. qed lemma (in semidom) prod_list_zero_iff[simp]: "prod_list xs = 0 \ 0 \ set xs" by (induction xs, auto) context comm_monoid_mult begin lemma unit_prod [intro]: shows "a dvd 1 \ b dvd 1 \ (a * b) dvd 1" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff[simp]: shows "(a * b) dvd 1 \ a dvd 1 \ b dvd 1" by (auto dest: dvd_mult_left dvd_mult_right) end context comm_semiring_1 begin lemma irreducibleE[elim]: assumes "irreducible p" and "p \ 0 \ \ p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ thesis" shows thesis using assms by (auto simp: irreducible_def) lemma not_irreducibleE: assumes "\ irreducible x" and "x = 0 \ thesis" and "x dvd 1 \ thesis" and "\a b. x = a * b \ \ a dvd 1 \ \ b dvd 1 \ thesis" shows thesis using assms unfolding irreducible_def by auto lemma prime_elem_dvd_prod_list: assumes p: "prime_elem p" and pA: "p dvd prod_list A" shows "\a \ set A. p dvd a" proof(insert pA, induct A) case Nil with p show ?case by (simp add: prime_elem_not_unit) next case (Cons a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma prime_elem_dvd_prod_mset: assumes p: "prime_elem p" and pA: "p dvd prod_mset A" shows "\a \# A. p dvd a" proof(insert pA, induct A) case empty with p show ?case by (simp add: prime_elem_not_unit) next case (add a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma mult_unit_dvd_iff[simp]: assumes "b dvd 1" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" using dvd_mult_left[of a b c] by simp next assume "a dvd c" with assms mult_dvd_mono show "a * b dvd c" by fastforce qed lemma mult_unit_dvd_iff'[simp]: "a dvd 1 \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ b dvd 1" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "b dvd 1 \ c dvd 1" . thus ?thesis by (auto simp: c) qed end context idom begin text \ Following lemmas are adapted and generalized so that they don't use "algebraic" classes. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (auto simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma irreducibleI': assumes "a \ 0" "\ a dvd 1" "\b. b dvd a \ a dvd b \ b dvd 1" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ b dvd 1" by (intro assms) simp_all thus "b dvd 1 \ c dvd 1" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis using dvd_times_left_cancel_iff by fastforce qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: shows "irreducible x \ x \ 0 \ \ x dvd 1 \ (\b. b dvd x \ x dvd b \ b dvd 1)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma dvd_mult_unit_iff: assumes b: "b dvd 1" shows "a dvd c * b \ a dvd c" proof- from b obtain b' where 1: "b * b' = 1" by (elim dvdE, auto) then have b0: "b \ 0" by auto from 1 have "a = (a * b') * b" by (simp add: ac_simps) also have "\ dvd c * b \ a * b' dvd c" using b0 by auto finally show ?thesis by (auto intro: dvd_mult_left) qed lemma dvd_mult_unit_iff': "b dvd 1 \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma irreducible_mult_unit_left: shows "a dvd 1 \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma irreducible_mult_unit_right: shows "a dvd 1 \ irreducible (p * a) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma unit_imp_dvd [dest]: "b dvd 1 \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_mult_left_cancel: "a dvd 1 \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "a dvd 1 \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) text \New parts from here\ lemma irreducible_multD: assumes l: "irreducible (a*b)" shows "a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" proof- from l have "a dvd 1 \ b dvd 1" using irreducibleD by auto then show ?thesis proof(elim disjE) assume a: "a dvd 1" with l have "irreducible b" unfolding irreducible_def by (metis is_unit_mult_iff mult.left_commute mult_not_zero) with a show ?thesis by auto next assume a: "b dvd 1" with l have "irreducible a" unfolding irreducible_def by (meson is_unit_mult_iff mult_not_zero semiring_normalization_rules(16)) with a show ?thesis by auto qed qed end lemma (in field) irreducible_field[simp]: "irreducible x \ False" by (auto simp: dvd_field_iff irreducible_def) lemma (in idom) irreducible_mult: shows "irreducible (a*b) \ a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" by (auto dest: irreducible_multD simp: irreducible_mult_unit_left irreducible_mult_unit_right) end diff --git a/thys/Source_Coding_Theorem/Source_Coding_Theorem.thy b/thys/Source_Coding_Theorem/Source_Coding_Theorem.thy --- a/thys/Source_Coding_Theorem/Source_Coding_Theorem.thy +++ b/thys/Source_Coding_Theorem/Source_Coding_Theorem.thy @@ -1,603 +1,603 @@ (* Title: One Part of Shannon's Source Coding Theorem Author: Quentin Hibon , Lawrence Paulson , 2014 Maintainer: Quentin Hibon *) theory Source_Coding_Theorem imports "HOL-Probability.Information" begin section\Basic types\ type_synonym bit = bool type_synonym bword = "bit list" type_synonym letter = nat type_synonym 'b word = "'b list" type_synonym 'b encoder = "'b word \ bword" type_synonym 'b decoder = "bword \ 'b word option" section\Locale for the source coding theorem\ locale source_code = information_space + fixes fi :: "'b \ real" fixes X :: "'a \ 'b" assumes distr_i: "simple_distributed M X fi" assumes b_val: "b = 2" fixes enc::"'b encoder" fixes dec::"'b decoder" assumes real_code: "dec (enc x) = Some x" "enc w = [] \ w = []" "x \ [] \ enc x = enc [hd x] @ enc (tl x)" section\Source coding theorem, direct: the entropy is a lower bound of the code rate\ context source_code begin subsection\The letter set\ definition L :: "'b set" where "L \ X ` space M" lemma fin_L: "finite L" using L_def distr_i by auto lemma emp_L: "L \ {}" using L_def subprob_not_empty by auto subsection\Codes and words\ abbreviation real_word :: "'b word \ bool" where "real_word w \ (set w \ L)" abbreviation k_words :: "nat \ ('b word) set" where "k_words k \ {w. length w = k \ real_word w}" lemma rw_tail: assumes "real_word w" shows "w = [] \ real_word (tl w)" by (meson assms list.set_sel(2) subset_code(1)) definition code_word_length :: "'e encoder \ 'e \ nat" where "code_word_length e l = length (e [l])" abbreviation cw_len :: "'b \ nat" where "cw_len l \ code_word_length enc l" definition code_rate :: "'e encoder \ ('a \ 'e) \ real" where "code_rate e Xo = expectation (\a. (code_word_length e ((Xo) a)))" lemma fi_pos: "i\ L \ 0 \ fi i" using simple_distributed_nonneg[OF distr_i] L_def by auto lemma (in prob_space) simp_exp_composed: assumes X: "simple_distributed M X Px" shows "expectation (\a. f (X a)) = (\x \ X`space M. f x * Px x)" using distributed_integral[OF simple_distributed[OF X], of f] simple_distributed_nonneg[OF X] lebesgue_integral_count_space_finite[OF simple_distributed_finite[OF X], of "\x. f x * Px x"] by (simp add: ac_simps) lemma cr_rw: "code_rate enc X = (\i \ X ` space M. fi i * cw_len i)" using simp_exp_composed[OF distr_i, of "cw_len"] by (simp add: mult.commute code_rate_def) abbreviation cw_len_concat :: "'b word \ nat" where "cw_len_concat w \ foldr (\x s. (cw_len x) + s) w 0" lemma cw_len_length: "cw_len_concat w = length (enc w)" proof (induction w) case Nil show ?case using real_code by simp case (Cons a w) have "cw_len_concat (a # w) = cw_len a + cw_len_concat w" by simp thus ?case using code_word_length_def real_code Cons by (metis length_append list.distinct(1) list.sel(1) list.sel(3)) qed lemma maj_fold: assumes "\l. l\L \ f l \ bound" assumes "real_word w" shows "foldr (\x s. f x + s) w 0 \ length w * bound" using assms by(induction w) (simp,fastforce) definition max_len :: "nat" where "max_len = Max ((\x. cw_len x) ` L)" lemma max_cw: "l \ L \ cw_len l \ max_len" by (simp add: max_len_def fin_L) subsection\Related to the Kraft theorem\ definition \ :: "real" where "\ = (\i\L. 1 / b ^ (cw_len i))" lemma pos_cw_len: "0 < 1 / b ^ cw_len i" using b_gt_1 by simp lemma \_pos: "0 < \" using emp_L fin_L pos_cw_len sum_pos \_def by metis lemma \_pow: "\ = (\i\L. 1 / b powr cw_len i)" using powr_realpow b_gt_1 by (simp add: \_def) lemma k_words_rel: "k_words (Suc k) = {w. (hd w \ L \ tl w \ k_words k \ w \ [])}" proof fix k show "k_words (Suc k) \ {w. (hd w \ L \ tl w \ k_words k \ w \ [] )}" (is "?l \ ?r") proof fix w assume w_kw: "w \ k_words (Suc k)" hence "real_word w" by simp hence "hd w \ L" by (metis (mono_tags) w_kw hd_in_set list.size(3) mem_Collect_eq nat.distinct(1) subset_code(1)) moreover have "length w = Suc k" using w_kw by simp moreover hence "w \ []" by auto moreover have "real_word (tl w)" using \real_word w\ calculation(3) rw_tail by auto ultimately show "w \ ?r" using w_kw by simp qed next fix k show "k_words (Suc k) \ {w. (hd w \ L \ tl w \ k_words k \ w \ [])}" proof fix w assume asm: "w \ {w. hd w \ L \ tl w \ {w. length w = k \ real_word w} \ w \ []}" hence "hd w \ L \ length (tl w) = k \ real_word (tl w)" by simp hence "real_word w" by (metis empty_iff insert_subset list.collapse list.set(1) set_simps(2) subsetI) moreover hence "length w = Suc k" using asm by auto ultimately show "w \ k_words (Suc k)" by simp qed qed lemma bij_k_words: shows "bij_betw (\wi. Cons (fst wi) (snd wi)) (L \ k_words k) (k_words (Suc k))" unfolding bij_betw_def proof fix k let ?f = "(\wi. Cons (fst wi) (snd wi))" let ?S = "L \ (k_words k)" let ?T = "k_words (Suc k)" show "inj_on ?f ?S" by (simp add: inj_on_def) show "?f`?S = ?T" proof (rule ccontr) assume "?f ` ?S \ ?T" hence "\w. w\ ?T \ w \ ?f`?S" by auto then obtain w where asm: "w\ ?T \ w \ ?f`?S" by blast hence "w = ?f (hd w,tl w)" using k_words_rel by simp moreover have "(hd w,tl w) \ ?S" using k_words_rel asm by simp ultimately have "w \ ?f`?S" by blast thus "False" using asm by simp qed qed lemma finite_k_words: "finite (k_words k)" proof (induct k) case 0 show ?case by simp case (Suc n) thus ?case using bij_k_words bij_betw_finite fin_L by blast qed lemma cartesian_product: fixes f::"('c \ real)" fixes g::"('d \ real)" assumes "finite A" assumes "finite B" shows "(\b\B. g b) * (\a\A. f a) = (\ab\A\B. f (fst ab) * g (snd ab))" using bilinear_times bilinear_sum[where h="(\x y. x * y)" and f="f" and g="g"] assms by (metis (erased, lifting) sum.cong split_beta' Groups.ab_semigroup_mult_class.mult.commute) lemma \_power: shows "\^k = (\w \ (k_words k). 1 / b^(cw_len_concat w))" proof (induct k) case 0 have "k_words 0 = {[]}" by auto thus ?case by simp next case (Suc n) have " \ ^Suc n = \ ^n * \ " by simp also have "\ = (\w \ k_words n. 1 / b^cw_len_concat w) * (\i\L. 1 / b^cw_len i)" using Suc.hyps \_def by auto also have "\ = (\wi \ L \ k_words n. 1/b^cw_len (fst wi) * (1 / b^cw_len_concat (snd wi)))" using fin_L finite_k_words cartesian_product by blast also have "\ = (\wi \ L \ k_words n. 1 / b^(cw_len_concat (snd wi) + cw_len (fst wi)))" by (metis (no_types, lifting) power_add add.commute power_one_over) also have "\ = (\wi \ L \ k_words n. 1 / b^cw_len_concat (fst wi # snd wi))" by (metis (erased, lifting) add.commute comp_apply foldr.simps(2)) also have "\ = (\w \ (k_words (Suc n)). 1 / b^(cw_len_concat w))" using bij_k_words sum.reindex_bij_betw by fastforce finally show ?case by simp qed lemma bound_len_concat: shows "w \ k_words k \ cw_len_concat w \ k * max_len" using max_cw maj_fold by blast subsection\Inequality of the kraft sum (source coding theorem, direct)\ subsubsection\Sum manipulation lemmas and McMillan theorem\ lemma sum_vimage_proof: fixes g::"nat \ real" assumes "\w. f w < bd" shows "finite S \ (\w\S. g (f w)) = (\ m=0.. S) )* g m)" (is "_ \ _ = (\ m=0.. {0..h::(nat \ real). (\m=0..y\({0..m = 0..m\{0.. F) * g (f x) + g (f x)" by (simp add: semiring_normalization_rules(2), simp add: insert) ultimately have "(\m = 0..m\{0.. real" assumes bounded: "\w. w \ S \ f w < bd" and "0 < bd" assumes finite: "finite S" shows "(\w\S. g (f w)) = (\ m=0.. S) ) * g m)" (is "?s1 = ?s2") proof - let ?ff = "(\x. if x\S then f x else 0)" let ?ss1 = "(\w\S. g (?ff w))" let ?ss2 = "(\ m=0.. S) ) * g m)" have "?s1 =?ss1" by simp moreover have"\m. ?ff -`{m} \ S = f-`{m} \ S" by auto moreover hence "?s2 = ?ss2" by simp moreover have "\w . ?ff w < bd" using assms by simp moreover hence "?ss1 = ?ss2" using sum_vimage_proof[of "?ff"] finite by blast ultimately show "?s1 = ?s2" by metis qed lemma \_rw: "(\w \ (k_words k). 1 / b^(cw_len_concat w)) = (\m=0.. ((cw_len_concat) -` {m})) * (1 / b^m))" (is "?L = ?R") proof - have "\w. w \ k_words k \ cw_len_concat w < Suc ( k * max_len)" by (simp add: bound_len_concat le_imp_less_Suc) moreover have "?R = (\m = 0.. k_words k)) * (1 / b ^ m))" by (metis Int_commute) moreover have "0 < Suc (k*max_len)" by simp ultimately show ?thesis using finite_k_words sum_vimage[where f="cw_len_concat" and g = "\i. 1/ (b^i)"] by fastforce qed definition set_of_k_words_length_m :: "nat \ nat \ 'b word set" where "set_of_k_words_length_m k m = {xk. xk \ k_words k} \ (cw_len_concat)-`{m}" lemma am_inj_code: "inj_on enc ((cw_len_concat)-`{m})" (is "inj_on _ ?s") using inj_on_def[of enc "?s"] real_code by (metis option.inject) lemma img_inc: "enc`cw_len_concat-`{m} \ {bl. length bl = m}" using cw_len_length by auto lemma bool_lists_card: "card {bl::bool list. length bl = m} = b^m" using card_lists_length_eq[of "UNIV::bool set"] by (simp add: b_val) lemma bool_list_fin: "finite {bl::bool list. length bl = m}" using finite_lists_length_eq[of "UNIV::bool set"] by (simp add: b_val) lemma set_of_k_words_bound: shows "card (set_of_k_words_length_m k m) \ b^m" (is "?c \ ?b") proof - have card_w_len_m_bound: "card (cw_len_concat-`{m}) \ b^m" by (metis (no_types, lifting) am_inj_code bool_list_fin bool_lists_card card_image card_mono img_inc of_nat_le_iff) have "set_of_k_words_length_m k m \ (cw_len_concat)-`{m}" by (simp add: set_of_k_words_length_m_def) hence "card (set_of_k_words_length_m k m) \ card ((cw_len_concat)-`{m})" by (metis (no_types, lifting) am_inj_code bool_list_fin card.infinite card_0_eq card_image card_mono empty_iff finite_subset img_inc inf_img_fin_dom) thus ?thesis using card_w_len_m_bound by simp qed lemma empty_set_k_words: assumes "0 < k" shows "set_of_k_words_length_m k 0 = {}" proof(rule ccontr) assume "\ set_of_k_words_length_m k 0 = {}" hence "\x. x \ set_of_k_words_length_m k 0" by auto then obtain x where x_def: "x \ set_of_k_words_length_m k 0" by auto hence "x \ []" unfolding set_of_k_words_length_m_def using assms by auto moreover have "cw_len_concat (hd x#tl x) = cw_len_concat (tl x) + cw_len (hd x)" by (metis add.commute comp_apply foldr.simps(2)) moreover have "enc [(hd x)] \ []" using assms real_code by blast moreover hence "0 < cw_len (hd x)" unfolding code_word_length_def by simp ultimately have "x \ set_of_k_words_length_m k 0" by (simp add:set_of_k_words_length_m_def) thus "False" using x_def by simp qed lemma \_rw2: assumes "0 < k" shows "(\m=0.. (k * max_len)" proof - have "(\m=1.. (\m=1..m. (card (set_of_k_words_length_m k m))/b^m)" "\m. b^m /b^m"] by simp moreover have"(\m=1..m=1..m=1..m = 1.. k * max_len" by (metis One_nat_def card_atLeastLessThan card_eq_sum diff_Suc_Suc real_of_card) thus ?thesis using empty_set_k_words assms by (simp add: sum_shift_lb_Suc0_0_upt split: if_split_asm) qed lemma \_power_bound : assumes "0 < k" shows " \^k \ k * max_len" using assms \_power \_rw \_rw2 by (simp add: set_of_k_words_length_m_def) theorem McMillan : shows "\ \ 1" proof - have ineq: "\k. 0 < k \ \ \ root k k * root k max_len" using \_pos \_power_bound by (metis (no_types, hide_lams) not_less of_nat_0_le_iff of_nat_mult power_strict_mono real_root_mult real_root_pos_pos_le real_root_pos_unique real_root_power) hence "0 < max_len \ (\k. root k k * root k max_len) \ 1" by (auto intro!: tendsto_eq_intros LIMSEQ_root LIMSEQ_root_const) moreover have "\n\1. \ \ root n n * root n max_len" using ineq by simp moreover have "max_len = 0 \ \ \ 1" using ineq by fastforce ultimately show " \ \ 1" using LIMSEQ_le_const by blast qed lemma entropy_rw: "\(X) = -(\i \ L. fi i * log b (fi i))" using entropy_simple_distributed[OF distr_i] by (simp add: L_def) subsubsection\Technical lemmas about the logarithm\ lemma log_mult_ext3: "0 \ x \ 0 < y \ 0 < z \ x * log b (x*y*z) = x * log b (x*y) + x * log b z" by(cases "x=0")(simp add: log_mult_eq abs_of_pos distrib_left less_eq_real_def)+ lemma log_mult_ext2: "0 \ x \ 0 < y \ x * log b (x*y) = x * log b x + x * log b y" using log_mult_ext3[where y=1] by simp subsubsection \KL divergence and properties\ definition KL_div ::"'b set \ ('b \ real) \ ('b \ real) \ real" where "KL_div S a d = (\ i \ S. a i * log b (a i / d i))" lemma KL_div_mul: assumes "0 < d" "d \ 1" assumes "\i. i\S \ 0 \ a i" assumes "\i. i\S \ 0 < e i" shows "KL_div S a e \ KL_div S a (\i. e i / d)" unfolding KL_div_def proof - { fix i assume "i\S" hence "a i / (e i / d) \ a i / e i" using assms by (metis (no_types) div_by_1 frac_le less_imp_triv not_less) hence "log b (a i / (e i / d)) \ log b (a i / e i)" using assms(1) by (metis (full_types) b_gt_1 divide_divide_eq_left inverse_divide le_less_linear log_le log_neg_const order_refl times_divide_eq_right zero_less_mult_iff) } thus "(\i\S. a i * log b (a i / (e i / d))) \ (\i\S. a i * log b (a i / e i))" by (meson mult_left_mono assms sum_mono) qed lemma KL_div_pos: fixes a e::"'b \ real" assumes fin: "finite S" assumes nemp: "S \ {}" assumes non_null: "\i. i\S \ 0 < a i" "\i. i\ S \ 0 < e i" assumes sum_a_one: "(\ i \ S. a i) = 1" assumes sum_c_one: "(\ i \ S. e i) = 1" shows "0 \ KL_div S a e" unfolding KL_div_def proof - let ?f = "\i. e i / a i" have f_pos: "\i. i\S \ 0 < ?f i" using non_null by simp have a_pos: "\i. i\ S \ 0 \ a i" using non_null by (simp add: order.strict_implies_order) have "- log b (\i\S. a i * e i / a i) \ (\i\S. a i * - log b (e i / a i))" using convex_on_sum[OF fin nemp minus_log_convex[OF b_gt_1] convex_real_interval(3) sum_a_one a_pos, of "\i. e i / a i"] f_pos by simp also have "-log b (\i\S. a i * e i / a i) = -log b (\i\S. e i)" proof - from non_null(1) have "\i. i \ S \ a i * e i / a i = e i" by force thus ?thesis by simp qed finally have "0 \ (\i\S. a i * - log b (e i / a i))" by (simp add: sum_c_one) thus "0 \ (\i\S. a i * log b (a i / e i))" using b_gt_1 log_divide non_null by simp qed lemma KL_div_pos_emp: "0 \ KL_div {} a e" by (simp add: KL_div_def) lemma KL_div_pos_gen: fixes a d::"'b \ real" assumes fin: "finite S" assumes non_null: "\i. i\S \ 0 < a i" "\i. i\ S \ 0 < d i" assumes sum_a_one: "(\ i \ S. a i) = 1" assumes sum_d_one: "(\ i \ S. d i) = 1" shows "0 \ KL_div S a d" using KL_div_pos KL_div_pos_emp assms by metis theorem KL_div_pos2: fixes a d::"'b \ real" assumes fin: "finite S" assumes non_null: "\i. i\S \ 0 \ a i" "\i. i\ S \ 0 < d i" assumes sum_a_one: "(\ i \ S. a i) = 1" assumes sum_c_one: "(\ i \ S. d i) = 1" shows "0 \ KL_div S a d" proof - have "S = (S \ {i. 0 < a i}) \ (S \ {i. 0 = a i})" using non_null(1) by fastforce moreover have "(S \ {i. 0 < a i}) \ (S \ {i. 0 = a i}) = {}" by auto ultimately have eq: "KL_div S a d = KL_div (S \ {i. 0 < a i}) a d + KL_div (S \ {i. 0 = a i}) a d" unfolding KL_div_def by (metis (mono_tags, lifting) fin finite_Un sum.union_disjoint) have "KL_div (S \ {i. 0 = a i}) a d = 0" unfolding KL_div_def by simp hence "KL_div S a d = KL_div (S \ {i. 0 < a i}) a d" using eq by simp moreover have "0 \ KL_div (S \ {i. 0 < a i}) a d" proof(cases "(S \ {i. 0 < a i}) = {}") case True thus ?thesis unfolding KL_div_def by simp next case False let ?c = "\i. d i / (\j \(S \ {i. 0 < a i}). d j)" have 1: "(\i. i \ S \ {i. 0 < a i} \ 0 < a i)" by simp have 2: "(\i. i \ S \ {i. 0 < a i} \ 0 < ?c i)" by (metis False IntD1 divide_pos_pos fin finite_Int non_null(2) sum_pos) have 3: "(\i\ (S \ {i. 0 < a i}). a i) = 1" using sum.cong[of S, of S, of "(\x. if x \ {i. 0 < a i} then a x else 0)", of a] sum.inter_restrict[OF fin, of a] non_null(1) sum_a_one by fastforce have "(\i\S \ {j. 0 < a j}. ?c i) = (\i\S \ {j. 0 < a j}. d i) / (\i\S \ {j. 0 < a j}. d i)" by (metis sum_divide_distrib) hence 5: "(\i\S \ {j. 0 < a j}. ?c i) = 1" using 2 False by force hence "0 \ KL_div (S \ {j. 0 < a j}) a ?c" using KL_div_pos_gen[ OF finite_Int[OF disjI1, of S, of "{j. 0 < a j}"], of a, of ?c ] 1 2 3 by (metis fin) have fstdb: "0 < (\i\S \ {i. 0 < a i}. d i)" using non_null(2) False by (metis Int_Collect fin finite_Int sum_pos) have 6: "0 \ KL_div (S \ {i. 0 < a i}) a (\i. d i / (\i\(S \ {i. 0 < a i}). d i))" using 2 3 5 KL_div_pos_gen[ OF finite_Int[OF disjI1, OF fin], of "{i. 0 < a i}", of "a", of "?c" ] by simp hence "KL_div (S \ {j. 0 < a j}) a (\i. d i / (\i\(S \ {i. 0 < a i}). d i)) \ KL_div (S \ {j. 0 < a j}) a d" using non_null sum.inter_restrict[OF fin, of d, of "{i. 0 < a i}"] sum_mono[of S, of "(\x. if x \ {i. 0 < a i} then d x else 0)", of d] non_null(2) sum_c_one non_null(2) fstdb KL_div_mul by force moreover have "0 \ KL_div (S \ {j. 0 < a j}) a (\i. d i / (\i\(S \ {i. 0 < a i}). d i))" using KL_div_pos_gen[ OF finite_Int[OF disjI1, OF fin]] using 2 3 5 by fastforce ultimately show "0 \ KL_div (S \ {j. 0 < a j}) a d" by simp qed ultimately show ?thesis by simp qed lemma sum_div_1: fixes f::"'b \ 'c::field" assumes "(\i\A. f i) \ 0" shows "(\i\A. f i / (\j\A. f j)) = 1" by (metis (no_types) assms right_inverse_eq sum_divide_distrib) theorem rate_lower_bound: shows "\(X) \ code_rate enc X" proof - let ?cr = "code_rate enc X" let ?r = "(\i. 1 / ((b powr cw_len i) * \))" have pos_pi: "\i. i \ L \ 0 \ fi i" using fi_pos by simp { fix i assume "i \ L" hence "fi i * (log b (1 / (1 / b powr (cw_len i))) + log b (fi i)) = fi i * log b (fi i / (1 / b powr (cw_len i)))" - using log_mult_ext2[OF pos_pi] b_gt_1 - by (simp add: linordered_field_class.sign_simps(36)) + using log_mult_ext2 [OF pos_pi, of i] b_gt_1 + by simp (simp add: algebra_simps) } hence eqpi: "\i. i\ L \ fi i * (log b (1 / (1 / b powr (cw_len i))) + log b (fi i)) = fi i * log b (fi i / (1 / b powr (cw_len i)))" by simp have sum_one_L: "(\ i \ L. fi i) = 1" using simple_distributed_sum_space[OF distr_i] by (simp add: L_def) { fix i assume "i \ L" hence h1: "0 \ fi i" using pos_pi by blast have h2: "0 < \ / (1/b powr cw_len i)" using b_gt_1 \_pos by auto have h3: "0 < 1 / \" using \_pos by simp have "fi i * log b (fi i * \ / (1/b powr cw_len i) * (1/ \)) = fi i * log b (fi i * \ / (1/b powr cw_len i)) + fi i * log b (1/ \)" using log_mult_ext3[OF h1 h2 h3] by (metis times_divide_eq_right) } hence big_eq: "\i. i \ L \ fi i * log b (fi i * \ / (1/b powr cw_len i) * (1 / \)) = fi i * log b (fi i * \ / (1/b powr cw_len i)) + fi i * log b (1 / \)" by (simp add: inverse_eq_divide) have 1: "?cr - \(X) = (\i \ L. fi i * cw_len i) + (\i \ L. fi i * log b (fi i))" using \_def entropy_rw cr_rw L_def by simp also have 2: "(\i\L. fi i * cw_len i) = (\i \ L. fi i * (-log b (1/(b powr (cw_len i)))))" using b_gt_1 log_divide by simp also have "\ = -1 * (\i \ L. fi i * (log b (1/(b powr (cw_len i)))))" using sum_distrib_left[of "-1" "(\i. fi i * (- 1 * log b (1 / b powr (cw_len i))))" L] by simp finally have "?cr - \(X) = -(\i \ L. fi i * log b (1/b powr cw_len i)) + (\i \ L. fi i * log b (fi i))" by simp have "?cr - \(X) = (\i \ L. fi i * ((log b (1/ (1/(b powr (cw_len i))))) + log b (fi i)))" using b_gt_1 1 by (simp add: distrib_left sum.distrib) also have "\ = (\i \ L. fi i *((log b (fi i / (1/(b powr (cw_len i)))))))" using Finite_Cartesian_Product.sum_cong_aux[OF eqpi] by simp also from big_eq have "\ = (\i\L. fi i * (log b (fi i * \ / (1 / b powr (cw_len i))))) + (\i \ L. fi i) * log b (1/ \)" using \_pos by (simp add: sum_distrib_right sum.distrib) also have "\ = (\i\L. fi i * (log b (fi i * \ / (1 / b powr (cw_len i))))) - log b (\)" using \_pos by (simp add: log_inverse_eq divide_inverse sum_one_L) also have "\ = (\ i \ L. fi i * log b (fi i / ?r i)) - log b (\)" by (metis (mono_tags, hide_lams) divide_divide_eq_left divide_divide_eq_right) also have "\ = KL_div L fi ?r - log b ( \)" using b_gt_1 \_pos log_inverse KL_div_def by simp also have "\ = KL_div L fi ?r + log b (1 / \)" using log_inverse b_val \_pos by (simp add: inverse_eq_divide) finally have code_ent_kl_log: "?cr - \(X) = KL_div L fi ?r + log b (1 / \)" by simp have "(\i\L. ?r i) = 1" using sum_div_1[of "\i. 1 / (b powr (cw_len i))"] \_pos \_pow by simp moreover have "\i. 0 < ?r i" using b_gt_1 \_pos by simp moreover have "(\i\L. fi i) = 1" using sum_one_L by simp ultimately have "0 \ KL_div L fi ?r" using KL_div_pos2[OF fin_L fi_pos] by simp hence "log b (1 / \) \ ?cr - \(X)" using code_ent_kl_log by simp moreover from McMillan have "0 \ log b (1 / \)" using \_pos by (simp add: b_gt_1) ultimately show ?thesis by simp qed end end diff --git a/thys/Taylor_Models/Interval.thy b/thys/Taylor_Models/Interval.thy --- a/thys/Taylor_Models/Interval.thy +++ b/thys/Taylor_Models/Interval.thy @@ -1,840 +1,840 @@ (* Author: Christoph Traut, TU Muenchen Fabian Immler, TU Muenchen *) section \Interval Type\ theory Interval imports "HOL-Analysis.Analysis" "HOL-Library.Set_Algebras" "HOL-Library.Float" begin text \A type of non-empty, closed intervals.\ typedef (overloaded) 'a interval = "{(a::'a::preorder, b). a \ b}" morphisms bounds_of_interval Interval by auto setup_lifting type_definition_interval lift_definition lower::"('a::preorder) interval \ 'a" is fst . lift_definition upper::"('a::preorder) interval \ 'a" is snd . lemma interval_eq_iff: "a = b \ lower a = lower b \ upper a = upper b" by transfer auto lemma interval_eqI: "lower a = lower b \ upper a = upper b \ a = b" by (auto simp: interval_eq_iff) lemma lower_le_upper[simp]: "lower i \ upper i" by transfer auto lift_definition set_of :: "'a::preorder interval \ 'a set" is "\x. {fst x .. snd x}" . lemma set_of_eq: "set_of x = {lower x .. upper x}" by transfer simp context notes [[typedef_overloaded]] begin lift_definition(code_dt) Interval'::"'a::preorder \ 'a::preorder \ 'a interval option" is "\a b. if a \ b then Some (a, b) else None" by auto end instantiation "interval" :: ("{preorder,equal}") equal begin definition "equal_class.equal a b \ (lower a = lower b) \ (upper a = upper b)" instance proof qed (simp add: equal_interval_def interval_eq_iff) end instantiation interval :: ("preorder") ord begin definition less_eq_interval :: "'a interval \ 'a interval \ bool" where "less_eq_interval a b \ lower b \ lower a \ upper a \ upper b" definition less_interval :: "'a interval \ 'a interval \ bool" where "less_interval x y = (x \ y \ \ y \ x)" instance proof qed end instantiation interval :: ("lattice") semilattice_sup begin lift_definition sup_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a, b) (c, d). (inf a c, sup b d)" by (auto simp: le_infI1 le_supI1) lemma lower_sup[simp]: "lower (sup A B) = inf (lower A) (lower B)" by transfer auto lemma upper_sup[simp]: "upper (sup A B) = sup (upper A) (upper B)" by transfer auto instance proof qed (auto simp: less_eq_interval_def less_interval_def interval_eq_iff) end lemma set_of_interval_union: "set_of A \ set_of B \ set_of (sup A B)" for A::"'a::lattice interval" by (auto simp: set_of_eq) lemma interval_union_commute: "sup A B = sup B A" for A::"'a::lattice interval" by (auto simp add: interval_eq_iff inf.commute sup.commute) lemma interval_union_mono1: "set_of a \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lemma interval_union_mono2: "set_of A \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lift_definition interval_of :: "'a::preorder \ 'a interval" is "\x. (x, x)" by auto lemma lower_interval_of[simp]: "lower (interval_of a) = a" by transfer auto lemma upper_interval_of[simp]: "upper (interval_of a) = a" by transfer auto definition width :: "'a::{preorder,minus} interval \ 'a" where "width i = upper i - lower i" instantiation "interval" :: ("ordered_ab_semigroup_add") ab_semigroup_add begin lift_definition plus_interval::"'a interval \ 'a interval \ 'a interval" is "\(a, b). \(c, d). (a + c, b + d)" by (auto intro!: add_mono) lemma lower_plus[simp]: "lower (plus A B) = plus (lower A) (lower B)" by transfer auto lemma upper_plus[simp]: "upper (plus A B) = plus (upper A) (upper B)" by transfer auto instance proof qed (auto simp: interval_eq_iff less_eq_interval_def ac_simps) end instance "interval" :: ("{ordered_ab_semigroup_add, lattice}") ordered_ab_semigroup_add proof qed (auto simp: less_eq_interval_def intro!: add_mono) instantiation "interval" :: ("{preorder,zero}") zero begin lift_definition zero_interval::"'a interval" is "(0, 0)" by auto lemma lower_zero[simp]: "lower 0 = 0" by transfer auto lemma upper_zero[simp]: "upper 0 = 0" by transfer auto instance proof qed end instance "interval" :: ("{ordered_comm_monoid_add}") comm_monoid_add proof qed (auto simp: interval_eq_iff) instance "interval" :: ("{ordered_comm_monoid_add,lattice}") ordered_comm_monoid_add .. instantiation "interval" :: ("{ordered_ab_group_add}") uminus begin lift_definition uminus_interval::"'a interval \ 'a interval" is "\(a, b). (-b, -a)" by auto lemma lower_uminus[simp]: "lower (- A) = - upper A" by transfer auto lemma upper_uminus[simp]: "upper (- A) = - lower A" by transfer auto instance .. end instantiation "interval" :: ("{ordered_ab_group_add}") minus begin definition minus_interval::"'a interval \ 'a interval \ 'a interval" where "minus_interval a b = a + - b" lemma lower_minus[simp]: "lower (minus A B) = minus (lower A) (upper B)" by (auto simp: minus_interval_def) lemma upper_minus[simp]: "upper (minus A B) = minus (upper A) (lower B)" by (auto simp: minus_interval_def) instance .. end instantiation "interval" :: (linordered_semiring) times begin lift_definition times_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a1, a2). \(b1, b2). (let x1 = a1 * b1; x2 = a1 * b2; x3 = a2 * b1; x4 = a2 * b2 in (min x1 (min x2 (min x3 x4)), max x1 (max x2 (max x3 x4))))" by (auto simp: Let_def intro!: min.coboundedI1 max.coboundedI1) lemma lower_times: "lower (times A B) = Min {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) lemma upper_times: "upper (times A B) = Max {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) instance .. end lemma interval_eq_set_of_iff: "X = Y \ set_of X = set_of Y" for X Y::"'a::order interval" by (auto simp: set_of_eq interval_eq_iff) subsection \Membership\ abbreviation (in preorder) in_interval ("(_/ \\<^sub>i _)" [51, 51] 50) where "in_interval x X \ x \ set_of X" lemma in_interval_to_interval[intro!]: "a \\<^sub>i interval_of a" by (auto simp: set_of_eq) lemma plus_in_intervalI: fixes x y :: "'a :: ordered_ab_semigroup_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x + y \\<^sub>i X + Y" by (simp add: add_mono_thms_linordered_semiring(1) set_of_eq) lemma connected_set_of[intro, simp]: "connected (set_of X)" for X::"'a::linear_continuum_topology interval" by (auto simp: set_of_eq ) lemma ex_sum_in_interval_lemma: "\xa\{la .. ua}. \xb\{lb .. ub}. x = xa + xb" if "la \ ua" "lb \ ub" "la + lb \ x" "x \ ua + ub" "ua - la \ ub - lb" for la b c d::"'a::linordered_ab_group_add" proof - define wa where "wa = ua - la" define wb where "wb = ub - lb" define w where "w = wa + wb" define d where "d = x - la - lb" define da where "da = max 0 (min wa (d - wa))" define db where "db = d - da" from that have nonneg: "0 \ wa" "0 \ wb" "0 \ w" "0 \ d" "d \ w" by (auto simp add: wa_def wb_def w_def d_def add.commute le_diff_eq) have "0 \ db" by (auto simp: da_def nonneg db_def intro!: min.coboundedI2) have "x = (la + da) + (lb + db)" by (simp add: da_def db_def d_def) moreover have "x - la - ub \ da" using that unfolding da_def by (intro max.coboundedI2) (auto simp: wa_def d_def diff_le_eq diff_add_eq) then have "db \ wb" by (auto simp: db_def d_def wb_def algebra_simps) with \0 \ db\ that nonneg have "lb + db \ {lb..ub}" by (auto simp: wb_def algebra_simps) moreover have "da \ wa" by (auto simp: da_def nonneg) then have "la + da \ {la..ua}" by (auto simp: da_def wa_def algebra_simps) ultimately show ?thesis by force qed lemma ex_sum_in_interval: "\xa\la. xa \ ua \ (\xb\lb. xb \ ub \ x = xa + xb)" if a: "la \ ua" and b: "lb \ ub" and x: "la + lb \ x" "x \ ua + ub" for la b c d::"'a::linordered_ab_group_add" proof - from linear consider "ua - la \ ub - lb" | "ub - lb \ ua - la" by blast then show ?thesis proof cases case 1 from ex_sum_in_interval_lemma[OF that 1] show ?thesis by auto next case 2 from x have "lb + la \ x" "x \ ub + ua" by (simp_all add: ac_simps) from ex_sum_in_interval_lemma[OF b a this 2] show ?thesis by auto qed qed lemma Icc_plus_Icc: "{a .. b} + {c .. d} = {a + c .. b + d}" if "a \ b" "c \ d" for a b c d::"'a::linordered_ab_group_add" using ex_sum_in_interval[OF that] by (auto intro: add_mono simp: atLeastAtMost_iff Bex_def set_plus_def) lemma set_of_plus: fixes A :: "'a::linordered_ab_group_add interval" shows "set_of (A + B) = set_of A + set_of B" using Icc_plus_Icc[of "lower A" "upper A" "lower B" "upper B"] by (auto simp: set_of_eq) lemma plus_in_intervalE: fixes xy :: "'a :: linordered_ab_group_add" assumes "xy \\<^sub>i X + Y" obtains x y where "xy = x + y" "x \\<^sub>i X" "y \\<^sub>i Y" using assms unfolding set_of_plus set_plus_def by auto lemma set_of_uminus: "set_of (-X) = {- x | x. x \ set_of X}" for X :: "'a :: ordered_ab_group_add interval" by (auto simp: set_of_eq simp: le_minus_iff minus_le_iff intro!: exI[where x="-x" for x]) lemma uminus_in_intervalI: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ -x \\<^sub>i -X" by (auto simp: set_of_uminus) lemma uminus_in_intervalD: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i - X \ - x \\<^sub>i X" by (auto simp: set_of_uminus) lemma minus_in_intervalI: fixes x y :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x - y \\<^sub>i X - Y" by (metis diff_conv_add_uminus minus_interval_def plus_in_intervalI uminus_in_intervalI) lemma set_of_minus: "set_of (X - Y) = {x - y | x y . x \ set_of X \ y \ set_of Y}" for X Y :: "'a :: linordered_ab_group_add interval" unfolding minus_interval_def set_of_plus set_of_uminus set_plus_def by force lemma times_in_intervalI: fixes x y::"'a::linordered_ring" assumes "x \\<^sub>i X" "y \\<^sub>i Y" shows "x * y \\<^sub>i X * Y" proof - define X1 where "X1 \ lower X" define X2 where "X2 \ upper X" define Y1 where "Y1 \ lower Y" define Y2 where "Y2 \ upper Y" from assms have assms: "X1 \ x" "x \ X2" "Y1 \ y" "y \ Y2" by (auto simp: X1_def X2_def Y1_def Y2_def set_of_eq) have "(X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y) \ (X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y)" proof (cases x "0::'a" rule: linorder_cases) case x0: less show ?thesis proof (cases "y < 0") case y0: True from y0 x0 assms have "x * y \ X1 * y" by (intro mult_right_mono_neg, auto) also from x0 y0 assms have "X1 * y \ X1 * Y1" by (intro mult_left_mono_neg, auto) finally have 1: "x * y \ X1 * Y1". show ?thesis proof(cases "X2 \ 0") case True with assms have "X2 * Y2 \ X2 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X2 * Y1 \ X2 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y1 \ x * y". with 1 show ?thesis by auto qed next case False then have y0: "y \ 0" by auto from x0 y0 assms have "X1 * Y2 \ x * Y2" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ x * y" by (intro mult_left_mono_neg, auto) finally have 1: "X1 * Y2 \ x * y". show ?thesis proof(cases "X2 \ 0") case X2: True from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y1" by (auto intro: mult_left_mono_neg) finally have "x * y \ X2 * Y1". with 1 show ?thesis by auto next case X2: False from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y2" by (auto intro: mult_left_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed qed next case [simp]: equal with assms show ?thesis by (cases "Y2 \ 0", auto intro:mult_sign_intros) next case x0: greater show ?thesis proof (cases "y < 0") case y0: True from x0 y0 assms have "X2 * Y1 \ X2 * y" by (intro mult_left_mono, auto) also from y0 x0 assms have "X2 * y \ x * y" by (intro mult_right_mono_neg, auto) finally have 1: "X2 * Y1 \ x * y". show ?thesis proof(cases "Y2 \ 0") case Y2: True from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X1 * Y2" by (auto intro: mult_right_mono_neg) finally have "x * y \ X1 * Y2". with 1 show ?thesis by auto next case Y2: False from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X2 * Y2" by (auto intro: mult_right_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed next case y0: False from x0 y0 assms have "x * y \ X2 * y" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ X2 * Y2" by (intro mult_left_mono, auto) finally have 1: "x * y \ X2 * Y2". show ?thesis proof(cases "X1 \ 0") case True with assms have "X1 * Y2 \ X1 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X1 * Y1 \ X1 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y1 \ x * y". with 1 show ?thesis by auto qed qed qed hence min:"min (X1 * Y1) (min (X1 * Y2) (min (X2 * Y1) (X2 * Y2))) \ x * y" and max:"x * y \ max (X1 * Y1) (max (X1 * Y2) (max (X2 * Y1) (X2 * Y2)))" by (auto simp:min_le_iff_disj le_max_iff_disj) show ?thesis using min max by (auto simp: Let_def X1_def X2_def Y1_def Y2_def set_of_eq lower_times upper_times) qed lemma times_in_intervalE: fixes xy :: "'a :: {linordered_semiring, real_normed_algebra, linear_continuum_topology}" \ \TODO: linear continuum topology is pretty strong\ assumes "xy \\<^sub>i X * Y" obtains x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" proof - let ?mult = "\(x, y). x * y" let ?XY = "set_of X \ set_of Y" have cont: "continuous_on ?XY ?mult" by (auto intro!: tendsto_eq_intros simp: continuous_on_def split_beta') have conn: "connected (?mult ` ?XY)" by (rule connected_continuous_image[OF cont]) auto have "lower (X * Y) \ ?mult ` ?XY" "upper (X * Y) \ ?mult ` ?XY" by (auto simp: set_of_eq lower_times upper_times min_def max_def split: if_splits) from connectedD_interval[OF conn this, of xy] assms obtain x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" by (auto simp: set_of_eq) then show ?thesis .. qed lemma set_of_times: "set_of (X * Y) = {x * y | x y. x \ set_of X \ y \ set_of Y}" for X Y::"'a :: {linordered_ring, real_normed_algebra, linear_continuum_topology} interval" by (auto intro!: times_in_intervalI elim!: times_in_intervalE) instance "interval" :: (linordered_idom) cancel_semigroup_add proof qed (auto simp: interval_eq_iff) lemma interval_mul_commute: "A * B = B * A" for A B:: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_right[simp]: "A * 0 = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_left[simp]: "0 * A = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) instantiation "interval" :: ("{preorder,one}") one begin lift_definition one_interval::"'a interval" is "(1, 1)" by auto lemma lower_one[simp]: "lower 1 = 1" by transfer auto lemma upper_one[simp]: "upper 1 = 1" by transfer auto instance proof qed end instance interval :: ("{one, preorder, linordered_semiring}") power proof qed lemma set_of_one[simp]: "set_of (1::'a::{one, order} interval) = {1}" by (auto simp: set_of_eq) instance "interval" :: ("{linordered_idom,linordered_ring, real_normed_algebra, linear_continuum_topology}") monoid_mult apply standard unfolding interval_eq_set_of_iff set_of_times subgoal for a b c by (auto simp: interval_eq_set_of_iff set_of_times; metis mult.assoc) by auto lemma one_times_ivl_left[simp]: "1 * A = A" for A :: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps min_def max_def) lemma one_times_ivl_right[simp]: "A * 1 = A" for A :: "'a::linordered_idom interval" by (metis interval_mul_commute one_times_ivl_left) lemma set_of_power_mono: "a^n \ set_of (A^n)" if "a \ set_of A" for a :: "'a::linordered_idom" using that by (induction n) (auto intro!: times_in_intervalI) lemma set_of_add_cong: "set_of (A + B) = set_of (A' + B')" if "set_of A = set_of A'" "set_of B = set_of B'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus that .. lemma set_of_add_inc_left: "set_of (A + B) \ set_of (A' + B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus using that by (auto simp: set_plus_def) lemma set_of_add_inc_right: "set_of (A + B) \ set_of (A + B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that] by (simp add: add.commute) lemma set_of_add_inc: "set_of (A + B) \ set_of (A' + B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that(1)] set_of_add_inc_right[OF that(2)] by auto lemma set_of_neg_inc: "set_of (-A) \ set_of (-A')" if "set_of A \ set_of A'" for A :: "'a::ordered_ab_group_add interval" using that unfolding set_of_uminus by auto lemma set_of_sub_inc_left: "set_of (A - B) \ set_of (A' - B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc_right: "set_of (A - B) \ set_of (A - B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc: "set_of (A - B) \ set_of (A' - B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_idom interval" using set_of_sub_inc_left[OF that(1)] set_of_sub_inc_right[OF that(2)] by auto lemma set_of_mul_inc_right: "set_of (A * B) \ set_of (A * B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ring interval" using that apply transfer apply (auto simp: Let_def) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) done lemma set_of_distrib_left: "set_of (B * (A1 + A2)) \ set_of (B * A1 + B * A2)" for A1 :: "'a::linordered_ring interval" apply transfer apply (auto simp: Let_def add_mono distrib_left distrib_right) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.assoc min.cobounded2) apply (meson add_mono_thms_linordered_semiring(1) dual_order.trans max.cobounded1 max.cobounded2) apply (meson add_mono_thms_linordered_semiring(1) dual_order.trans max.cobounded1 max.cobounded2) apply (meson add_mono_thms_linordered_semiring(1) dual_order.trans max.cobounded1 max.cobounded2) done lemma set_of_distrib_right: "set_of ((A1 + A2) * B) \ set_of (A1 * B + A2 * B)" for A1 A2 B :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" unfolding set_of_times set_of_plus set_plus_def apply clarsimp subgoal for b a1 a2 apply (rule exI[where x="a1 * b"]) apply (rule conjI) subgoal by force subgoal apply (rule exI[where x="a2 * b"]) apply (rule conjI) subgoal by force subgoal by (simp add: algebra_simps) done done done lemma set_of_mul_inc_left: "set_of (A * B) \ set_of (A' * B)" if "set_of A \ set_of A'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_mul_inc: "set_of (A * B) \ set_of (A' * B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_pow_inc: "set_of (A^n) \ set_of (A'^n)" if "set_of A \ set_of A'" for A :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" using that by (induction n, simp_all add: set_of_mul_inc) lemma set_of_distrib_right_left: "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * B1 + A1 * B2 + A2 * B1 + A2 * B2)" for A1 :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" proof- have "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * (B1 + B2) + A2 * (B1 + B2))" by (rule set_of_distrib_right) also have "... \ set_of ((A1 * B1 + A1 * B2) + A2 * (B1 + B2))" by (rule set_of_add_inc_left[OF set_of_distrib_left]) also have "... \ set_of ((A1 * B1 + A1 * B2) + (A2 * B1 + A2 * B2))" by (rule set_of_add_inc_right[OF set_of_distrib_left]) finally show ?thesis by (simp add: add.assoc) qed lemma mult_bounds_enclose_zero1: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "la \ 0" "0 \ ua" for la lb ua ub:: "'a::linordered_idom" subgoal by (metis (no_types, hide_lams) that eq_iff min_le_iff_disj mult_zero_left mult_zero_right zero_le_mult_iff) subgoal by (metis that le_max_iff_disj mult_zero_right order_refl zero_le_mult_iff) done lemma mult_bounds_enclose_zero2: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "lb \ 0" "0 \ ub" for la lb ua ub:: "'a::linordered_idom" using mult_bounds_enclose_zero1[OF that, of la ua] by (simp_all add: ac_simps) lemma set_of_mul_contains_zero: "0 \ set_of (A * B)" if "0 \ set_of A \ 0 \ set_of B" for A :: "'a::linordered_idom interval" using that - by (auto simp: set_of_eq lower_times upper_times sign_simps mult_le_0_iff + by (auto simp: set_of_eq lower_times upper_times algebra_simps mult_le_0_iff mult_bounds_enclose_zero1 mult_bounds_enclose_zero2) instance "interval" :: (linordered_semiring) mult_zero apply standard subgoal by transfer auto subgoal by transfer auto done lift_definition min_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (min l1 l2, min u1 u2)" by (auto simp: min_def) lemma lower_min_interval[simp]: "lower (min_interval x y) = min (lower x) (lower y)" by transfer auto lemma upper_min_interval[simp]: "upper (min_interval x y) = min (upper x) (upper y)" by transfer auto lemma min_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ min a b \\<^sub>i min_interval A B" by (auto simp: set_of_eq min_def) lift_definition max_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (max l1 l2, max u1 u2)" by (auto simp: max_def) lemma lower_max_interval[simp]: "lower (max_interval x y) = max (lower x) (lower y)" by transfer auto lemma upper_max_interval[simp]: "upper (max_interval x y) = max (upper x) (upper y)" by transfer auto lemma max_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ max a b \\<^sub>i max_interval A B" by (auto simp: set_of_eq max_def) lift_definition abs_interval::"'a::linordered_idom interval \ 'a interval" is "(\(l,u). (if l < 0 \ 0 < u then 0 else min \l\ \u\, max \l\ \u\))" by auto lemma lower_abs_interval[simp]: "lower (abs_interval x) = (if lower x < 0 \ 0 < upper x then 0 else min \lower x\ \upper x\)" by transfer auto lemma upper_abs_interval[simp]: "upper (abs_interval x) = max \lower x\ \upper x\" by transfer auto lemma in_abs_intervalI1: "lx < 0 \ 0 < ux \ 0 \ xa \ xa \ max (- lx) (ux) \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (metis abs_minus_cancel abs_of_nonneg atLeastAtMost_iff image_eqI le_less le_max_iff_disj le_minus_iff neg_le_0_iff_le order_trans) lemma in_abs_intervalI2: "min (\lx\) \ux\ \ xa \ xa \ max \lx\ \ux\ \ lx \ ux \ 0 \ lx \ ux \ 0 \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (force intro: image_eqI[where x="-xa"] image_eqI[where x="xa"]) lemma set_of_abs_interval: "set_of (abs_interval x) = abs ` set_of x" by (auto simp: set_of_eq not_less intro: in_abs_intervalI1 in_abs_intervalI2 cong del: image_cong_simp) fun split_domain :: "('a::preorder interval \ 'a interval list) \ 'a interval list \ 'a interval list list" where "split_domain split [] = [[]]" | "split_domain split (I#Is) = ( let S = split I; D = split_domain split Is in concat (map (\d. map (\s. s # d) S) D) )" context notes [[typedef_overloaded]] begin lift_definition(code_dt) split_interval::"'a::linorder interval \ 'a \ ('a interval \ 'a interval)" is "\(l, u) x. ((min l x, max l x), (min u x, max u x))" by (auto simp: min_def) end lemma split_domain_nonempty: assumes "\I. split I \ []" shows "split_domain split I \ []" using last_in_set assms by (induction I, auto) lemma split_intervalD: "split_interval X x = (A, B) \ set_of X \ set_of A \ set_of B" unfolding set_of_eq by transfer (auto simp: min_def max_def split: if_splits) definition "split_float_interval x = split_interval x ((lower x + upper x) * Float 1 (-1))" lemma split_float_intervalD: "split_float_interval X = (A, B) \ set_of X \ set_of A \ set_of B" by (auto dest!: split_intervalD simp: split_float_interval_def) lemmas float_round_down_le[intro] = order_trans[OF float_round_down] and float_round_up_ge[intro] = order_trans[OF _ float_round_up] instantiation interval :: ("{topological_space, preorder}") topological_space begin definition open_interval_def[code del]: "open (X::'a interval set) = (\x\X. \A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ X)" instance proof show "open (UNIV :: ('a interval) set)" unfolding open_interval_def by auto next fix S T :: "('a interval) set" assume "open S" "open T" show "open (S \ T)" unfolding open_interval_def proof (safe) fix x assume "x \ S" "x \ T" from \x \ S\ \open S\ obtain Sl Su where S: "open Sl" "open Su" "lower x \ Sl" "upper x \ Su" "Interval ` (Sl \ Su) \ S" by (auto simp: open_interval_def) from \x \ T\ \open T\ obtain Tl Tu where T: "open Tl" "open Tu" "lower x \ Tl" "upper x \ Tu" "Interval ` (Tl \ Tu) \ T" by (auto simp: open_interval_def) let ?L = "Sl \ Tl" and ?U = "Su \ Tu" have "open ?L \ open ?U \ lower x \ ?L \ upper x \ ?U \ Interval ` (?L \ ?U) \ S \ T" using S T by (auto simp add: open_Int) then show "\A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ S \ T" by fast qed qed (unfold open_interval_def, fast) end definition mid :: "float interval \ float" where "mid i = (lower i + upper i) * Float 1 (-1)" lemma mid_in_interval: "mid i \\<^sub>i i" using lower_le_upper[of i] by (auto simp: mid_def set_of_eq powr_minus) definition centered :: "float interval \ float interval" where "centered i = i - interval_of (mid i)" subsection \Quickcheck\ lift_definition Ivl::"'a \ 'a::preorder \ 'a interval" is "\a b. (min a b, b)" by (auto simp: min_def) instantiation interval :: ("{exhaustive,preorder}") exhaustive begin definition exhaustive_interval::"('a interval \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "exhaustive_interval f d = Quickcheck_Exhaustive.exhaustive (\x. Quickcheck_Exhaustive.exhaustive (\y. f (Ivl x y)) d) d" instance .. end definition (in term_syntax) [code_unfold]: "valtermify_interval x y = Code_Evaluation.valtermify (Ivl::'a::{preorder,typerep}\_) {\} x {\} y" instantiation interval :: ("{full_exhaustive,preorder,typerep}") full_exhaustive begin definition full_exhaustive_interval:: "('a interval \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_interval f d = Quickcheck_Exhaustive.full_exhaustive (\x. Quickcheck_Exhaustive.full_exhaustive (\y. f (valtermify_interval x y)) d) d" instance .. end instantiation interval :: ("{random,preorder,typerep}") random begin definition random_interval :: "natural \ natural \ natural \ ('a interval \ (unit \ term)) \ natural \ natural" where "random_interval i = scomp (Quickcheck_Random.random i) (\man. scomp (Quickcheck_Random.random i) (\exp. Pair (valtermify_interval man exp)))" instance .. end end diff --git a/thys/Taylor_Models/Interval_Approximation.thy b/thys/Taylor_Models/Interval_Approximation.thy --- a/thys/Taylor_Models/Interval_Approximation.thy +++ b/thys/Taylor_Models/Interval_Approximation.thy @@ -1,774 +1,774 @@ section \Approximate Operations on Intervals of Floating Point Numbers\ theory Interval_Approximation imports "HOL-Decision_Procs.Approximation_Bounds" Interval begin lifting_update float.lifting \ \TODO: in Float!\ lifting_forget float.lifting text \TODO: many of the lemmas should move to theories Float or Approximation (the latter should be based on type @{type interval}.\ subsection "Intervals with Floating Point Bounds" lift_definition round_interval :: "nat \ float interval \ float interval" is "\p. \(l, u). (float_round_down p l, float_round_up p u)" by (auto simp: intro!: float_round_down_le float_round_up_le) lemma lower_round_ivl[simp]: "lower (round_interval p x) = float_round_down p (lower x)" by transfer auto lemma upper_round_ivl[simp]: "upper (round_interval p x) = float_round_up p (upper x)" by transfer auto lemma round_ivl_correct: "set_of A \ set_of (round_interval prec A)" by (auto simp: set_of_eq float_round_down_le float_round_up_le) lift_definition truncate_ivl :: "nat \ real interval \ real interval" is "\p. \(l, u). (truncate_down p l, truncate_up p u)" by (auto intro!: truncate_down_le truncate_up_le) lemma lower_truncate_ivl[simp]: "lower (truncate_ivl p x) = truncate_down p (lower x)" by transfer auto lemma upper_truncate_ivl[simp]: "upper (truncate_ivl p x) = truncate_up p (upper x)" by transfer auto lemma truncate_ivl_correct: "set_of A \ set_of (truncate_ivl prec A)" by (auto simp: set_of_eq intro!: truncate_down_le truncate_up_le) lift_definition real_interval::"float interval \ real interval" is "\(l, u). (real_of_float l, real_of_float u)" by auto lemma lower_real_interval[simp]: "lower (real_interval x) = lower x" by transfer auto lemma upper_real_interval[simp]: "upper (real_interval x) = upper x" by transfer auto definition "set_of' x = (case x of None \ UNIV | Some i \ set_of (real_interval i))" lemma real_interval_min_interval[simp]: "real_interval (min_interval a b) = min_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_min) lemma real_interval_max_interval[simp]: "real_interval (max_interval a b) = max_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max) subsection \Intervals for standard functions\ lift_definition power_float_interval :: "nat \ nat \ float interval \ float interval" is "\p n (l, u). float_power_bnds p n l u" using float_power_bnds by (auto simp: bnds_power dest!: float_power_bnds[OF sym]) lemma lower_power_float_interval[simp]: "lower (power_float_interval p n x) = fst (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma upper_power_float_interval[simp]: "upper (power_float_interval p n x) = snd (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma power_float_intervalI: "x \\<^sub>i real_interval X \ x ^ n \\<^sub>i real_interval (power_float_interval p n X)" using float_power_bnds[OF prod.collapse] by (auto simp: set_of_eq ) lift_definition mult_float_interval::"nat \ float interval \ float interval \ float interval" is "\prec. \(a1, a2). \(b1, b2). bnds_mult prec a1 a2 b1 b2" by (auto dest!: bnds_mult[OF sym]) lemma lower_mult_float_interval[simp]: "lower (mult_float_interval p x y) = fst (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma upper_mult_float_interval[simp]: "upper (mult_float_interval p x y) = snd (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma mult_float_interval: "set_of (real_interval A) * set_of (real_interval B) \ set_of (real_interval (mult_float_interval prec A B))" proof - let ?bm = "bnds_mult prec (lower A) (upper A) (lower B) (upper B)" show ?thesis using bnds_mult[of "fst ?bm" "snd ?bm", simplified, OF refl] by (auto simp: set_of_eq set_times_def) qed lemma mult_float_intervalI: "x * y \\<^sub>i (real_interval (mult_float_interval prec A B))" if "x \\<^sub>i real_interval A" "y \\<^sub>i real_interval B" using mult_float_interval[of A B] that by (auto simp: ) lift_definition sqrt_float_interval::"nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_sqrt prec lx, ub_sqrt prec ux)" using bnds_sqrt' by auto (meson order_trans real_sqrt_le_iff) lemma lower_float_interval[simp]: "lower (sqrt_float_interval prec X) = lb_sqrt prec (lower X)" by transfer auto lemma upper_float_interval[simp]: "upper (sqrt_float_interval prec X) = ub_sqrt prec (upper X)" by transfer auto lemma sqrt_float_interval: "sqrt ` set_of (real_interval X) \ set_of (real_interval (sqrt_float_interval prec X))" using bnds_sqrt by (auto simp: set_of_eq) lemma sqrt_float_intervalI: "sqrt x \\<^sub>i real_interval (sqrt_float_interval p X)" if "x \ set_of (real_interval X)" using sqrt_float_interval[of X p] that by auto lemmas [simp del] = lb_arctan.simps ub_arctan.simps lemma lb_arctan: "arctan (real_of_float x) \ y \ real_of_float (lb_arctan prec x) \ y" and ub_arctan: "y \ arctan x \ y \ ub_arctan prec x" for x::float and y::real using arctan_boundaries[of x prec] by auto lift_definition arctan_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_arctan prec lx, ub_arctan prec ux)" by (auto intro!: lb_arctan ub_arctan arctan_monotone') lemma lower_arctan_float_interval[simp]: "lower (arctan_float_interval p x) = lb_arctan p (lower x)" by transfer auto lemma upper_arctan_float_interval[simp]: "upper (arctan_float_interval p x) = ub_arctan p (upper x)" by transfer auto lemma arctan_float_interval: "arctan ` set_of (real_interval x) \ set_of (real_interval (arctan_float_interval p x))" by (auto simp: set_of_eq intro!: lb_arctan ub_arctan arctan_monotone') lemma arctan_float_intervalI: "arctan x \\<^sub>i real_interval (arctan_float_interval p X)" if "x \ set_of (real_interval X)" using arctan_float_interval[of X p] that by auto lemma bnds_cos_lower: "\x. real_of_float xl \ x \ x \ real_of_float xu \ cos x \ y \ real_of_float (fst (bnds_cos prec xl xu)) \ y" and bnds_cos_upper: "\x. real_of_float xl \ x \ x \ real_of_float xu \ y \ cos x \ y \ real_of_float (snd (bnds_cos prec xl xu))" for xl xu::float and y::real using bnds_cos[of "fst (bnds_cos prec xl xu)" "snd (bnds_cos prec xl xu)" prec] by force+ lift_definition cos_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). bnds_cos prec lx ux" using bnds_cos by auto (metis (full_types) order_refl order_trans) lemma lower_cos_float_interval[simp]: "lower (cos_float_interval p x) = fst (bnds_cos p (lower x) (upper x))" by transfer auto lemma upper_cos_float_interval[simp]: "upper (cos_float_interval p x) = snd (bnds_cos p (lower x) (upper x))" by transfer auto lemma cos_float_interval: "cos ` set_of (real_interval x) \ set_of (real_interval (cos_float_interval p x))" by (auto simp: set_of_eq bnds_cos_lower bnds_cos_upper) lemma cos_float_intervalI: "cos x \\<^sub>i real_interval (cos_float_interval p X)" if "x \ set_of (real_interval X)" using cos_float_interval[of X p] that by auto lemma lb_exp: "exp x \ y \ lb_exp prec x \ y" and ub_exp: "y \ exp x \ y \ ub_exp prec x" for x::float and y::real using exp_boundaries[of x prec] by auto lift_definition exp_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_exp prec lx, ub_exp prec ux)" by (auto simp: lb_exp ub_exp) lemma lower_exp_float_interval[simp]: "lower (exp_float_interval p x) = lb_exp p (lower x)" by transfer auto lemma upper_exp_float_interval[simp]: "upper (exp_float_interval p x) = ub_exp p (upper x)" by transfer auto lemma exp_float_interval: "exp ` set_of (real_interval x) \ set_of (real_interval (exp_float_interval p x))" using exp_boundaries apply (auto simp: set_of_eq) apply (smt exp_le_cancel_iff) apply (smt exp_le_cancel_iff) done lemma exp_float_intervalI: "exp x \\<^sub>i real_interval (exp_float_interval p X)" if "x \ set_of (real_interval X)" using exp_float_interval[of X p] that by auto lemmas [simp del] = lb_ln.simps ub_ln.simps lemma lb_lnD: "y \ ln x \ 0 < real_of_float x" if "lb_ln prec x = Some y" using lb_ln[OF that[symmetric]] by auto lemma ub_lnD: "ln x \ y\ 0 < real_of_float x" if "ub_ln prec x = Some y" using ub_ln[OF that[symmetric]] by auto lift_definition(code_dt) ln_float_interval :: "nat \ float interval \ float interval option" is "\prec. \(lx, ux). Option.bind (lb_ln prec lx) (\l. Option.bind (ub_ln prec ux) (\u. Some (l, u)))" by (auto simp: pred_option_def bind_eq_Some_conv ln_le_cancel_iff[symmetric] simp del: ln_le_cancel_iff dest!: lb_lnD ub_lnD) lemma ln_float_interval_eq_Some_conv[simp]: "ln_float_interval p x = Some y \ lb_ln p (lower x) = Some (lower y) \ ub_ln p (upper x) = Some (upper y)" by transfer (auto simp: bind_eq_Some_conv) lemma ln_float_interval: "ln ` set_of (real_interval x) \ set_of (real_interval y)" if "ln_float_interval p x = Some y" using that by (simp add: set_of_eq) (smt atLeastAtMost_iff bnds_ln image_subset_iff) lemma ln_float_intervalI: "ln x \ set_of' (ln_float_interval p X)" if "x \\<^sub>i (real_interval X)" using ln_float_interval[of p X] that by (auto simp: set_of'_def split: option.splits) lift_definition(code_dt) powr_float_interval :: "nat \ float interval \ float interval \ float interval option" is "\prec. \(l1, u1). \(l2, u2). bnds_powr prec l1 u1 l2 u2" by (auto simp: pred_option_def dest!: bnds_powr[OF sym]) lemma powr_float_interval: "{x powr y | x y. x \ set_of (real_interval X) \ y \ set_of (real_interval Y)} \ set_of (real_interval R)" if "powr_float_interval prec X Y = Some R" using that by transfer (auto dest!: bnds_powr[OF sym]) lemma powr_float_intervalI: "x powr y \ set_of' (powr_float_interval p X Y)" if "x \\<^sub>i real_interval X" "y \\<^sub>i real_interval Y" using powr_float_interval[of p X Y] that by (auto simp: set_of'_def split: option.splits) lift_definition(code_dt) inverse_float_interval::"nat \ float interval \ float interval option" is "\prec (l, u). if (0 < l \ u < 0) then Some (float_divl prec 1 u, float_divr prec 1 l) else None" by (auto intro!: order_trans[OF float_divl] order_trans[OF _ float_divr] simp: divide_simps) lemma inverse_float_interval_eq_Some_conv[simp]: defines "one \ (1::float)" shows "inverse_float_interval p X = Some R \ (lower X > 0 \ upper X < 0) \ lower R = float_divl p one (upper X) \ upper R = float_divr p one (lower X)" by clarsimp (transfer fixing: one, force simp: one_def split: if_splits) lemma inverse_float_interval: "inverse ` set_of (real_interval X) \ set_of (real_interval Y)" if "inverse_float_interval p X = Some Y" using that apply (clarsimp simp: set_of_eq) by (intro order_trans[OF float_divl] order_trans[OF _ float_divr] conjI) (auto simp: divide_simps) lemma inverse_float_intervalI: "x \ set_of (real_interval X) \ inverse x \ set_of' (inverse_float_interval p X)" using inverse_float_interval[of p X] by (auto simp: set_of'_def split: option.splits) lift_definition pi_float_interval::"nat \ float interval" is "\prec. (lb_pi prec, ub_pi prec)" using pi_boundaries by (auto intro: order_trans) lemma lower_pi_float_interval[simp]: "lower (pi_float_interval prec) = lb_pi prec" by transfer auto lemma upper_pi_float_interval[simp]: "upper (pi_float_interval prec) = ub_pi prec" by transfer auto lemma pi_float_interval: "pi \ set_of (real_interval (pi_float_interval prec))" using pi_boundaries by (auto simp: set_of_eq) lemma real_interval_abs_interval[simp]: "real_interval (abs_interval x) = abs_interval (real_interval x)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max real_of_float_min) lift_definition floor_float_interval::"float interval \ float interval" is "\(l, u). (floor_fl l, floor_fl u)" by (auto intro!: floor_mono simp: floor_fl.rep_eq) lemma lower_floor_float_interval[simp]: "lower (floor_float_interval x) = floor_fl (lower x)" by transfer auto lemma upper_floor_float_interval[simp]: "upper (floor_float_interval x) = floor_fl (upper x)" by transfer auto lemma floor_float_intervalI: "\x\ \\<^sub>i real_interval (floor_float_interval X)" if "x \\<^sub>i real_interval X" using that by (auto simp: set_of_eq floor_fl_def floor_mono) lemma in_intervalI: "x \\<^sub>i X" if "lower X \ x" "x \ upper X" using that by (auto simp: set_of_eq) abbreviation in_real_interval ("(_/ \\<^sub>r _)" [51, 51] 50) where "x \\<^sub>r X \ x \\<^sub>i real_interval X" lemma in_real_intervalI: "x \\<^sub>r X" if "lower X \ x" "x \ upper X" for x::real and X::"float interval" using that by (intro in_intervalI) auto lemma lower_Interval: "lower (Interval x) = fst x" and upper_Interval: "upper (Interval x) = snd x" if "fst x \ snd x" using that by (auto simp: lower_def upper_def Interval_inverse split_beta') definition all_in_i :: "'a::preorder list \ 'a interval list \ bool" (infix "(all'_in\<^sub>i)" 50) where "x all_in\<^sub>i I = (length x = length I \ (\i < length I. x ! i \\<^sub>i I ! i))" definition all_in :: "real list \ float interval list \ bool" (infix "(all'_in)" 50) where "x all_in I = (length x = length I \ (\i < length I. x ! i \\<^sub>r I ! i))" definition all_subset :: "'a::order interval list \ 'a interval list \ bool" (infix "(all'_subset)" 50) where "I all_subset J = (length I = length J \ (\i < length I. set_of (I!i) \ set_of (J!i)))" lemmas [simp] = all_in_def all_subset_def lemma all_subsetD: assumes "I all_subset J" assumes "x all_in I" shows "x all_in J" using assms by (auto simp: set_of_eq; fastforce) lemma plus_down_mono: "plus_down p a b \ plus_down p c d" if "a + b \ c + d" by (auto simp: plus_down_def intro!: truncate_down_mono that) lemma plus_up_mono: "plus_up p a b \ plus_up p c d" if "a + b \ c + d" by (auto simp: plus_up_def intro!: truncate_up_mono that) lemma round_interval_mono: "set_of (round_interval prec X) \ set_of (round_interval prec Y)" if "set_of X \ set_of Y" using that by transfer (auto simp: float_round_down.rep_eq float_round_up.rep_eq truncate_down_mono truncate_up_mono) lemma mult_mono_nonpos_nonneg: "a * b \ c * d" if "a \ c" "a \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" apply (rule order_trans[OF mult_left_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_right_mono; fact) lemma mult_mono_nonneg_nonpos: "b * a \ d * c" if "a \ c" "c \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" apply (rule order_trans[OF mult_right_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_left_mono; fact) lemma mult_mono_nonpos_nonpos: "a * b \ c * d" if "a \ c" "a \ 0" "b \ d" "d \ 0" for a b c d::real apply (rule order_trans[OF mult_left_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_right_mono_neg; fact) lemma mult_float_mono1: notes mono_rules = plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_down prec (nprt aa * pprt bc) (plus_down prec (nprt ba * nprt bc) (plus_down prec (pprt aa * pprt ac) (pprt ba * nprt ac))) \ plus_down prec (nprt a * pprt bb) (plus_down prec (nprt b * nprt bb) (plus_down prec (pprt a * pprt ab) (pprt b * nprt ab)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ lemma mult_float_mono2: notes mono_rules = plus_up_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_up prec (pprt b * pprt bb) (plus_up prec (pprt a * nprt bb) (plus_up prec (nprt b * pprt ab) (nprt a * nprt ab))) \ plus_up prec (pprt ba * pprt bc) (plus_up prec (pprt aa * nprt bc) (plus_up prec (nprt ba * pprt ac) (nprt aa * nprt ac)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ lemma mult_float_interval_mono: "set_of (mult_float_interval prec A B) \ set_of (mult_float_interval prec X Y)" if "set_of A \ set_of X" "set_of B \ set_of Y" using that apply transfer unfolding bnds_mult_def atLeastatMost_subset_iff float_plus_down.rep_eq float_plus_up.rep_eq by (auto simp: float_plus_down.rep_eq float_plus_up.rep_eq mult_float_mono1 mult_float_mono2) lemma Ivl_simps[simp]: "lower (Ivl a b) = min a b" "upper (Ivl a b) = b" subgoal by transfer simp subgoal by transfer simp done lemmas [simp del] = power_down.simps(2) power_up.simps(2) lemmas power_down_simp = power_down.simps(2) lemmas power_up_simp = power_up.simps(2) lemma power_down_even_nonneg: "even n \ 0 \ power_down p x n" by (induct p x n rule: power_down.induct) (auto simp: power_down_simp simp del: odd_Suc_div_two intro!: truncate_down_nonneg ) lemma truncate_down_less_zero_iff[simp]: "truncate_down p x < 0 \ x < 0" by (metis le_less_trans not_less_iff_gr_or_eq truncate_down truncate_down_pos truncate_down_zero) lemma truncate_down_nonneg_iff[simp]: "truncate_down p x \ 0 \ x \ 0" using truncate_down_less_zero_iff[of p x] truncate_down_nonneg[of x p] by linarith lemma truncate_down_eq_zero_iff[simp]: "truncate_down prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_down_less_zero_iff truncate_down_pos truncate_down_zero) lemma power_down_eq_zero_iff[simp]: "power_down prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_down_simp[of _ _ "x - 1"] - by (cases x) (auto simp: sign_simps zero_le_mult_iff div2_less_self) + by (cases x) (auto simp add: div2_less_self) qed lemma power_down_nonneg_iff[simp]: "power_down prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_down_simp[of _ _ "x - 1"] by (cases x) (auto simp: sign_simps zero_le_mult_iff) qed lemma power_down_neg_iff[simp]: "power_down prec b n < 0 \ b < 0 \ odd n" using power_down_nonneg_iff[of prec b n] by (auto simp del: power_down_nonneg_iff) lemma power_down_nonpos_iff[simp]: notes [simp del] = power_down_neg_iff power_down_eq_zero_iff shows "power_down prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_down_neg_iff[of prec b n] power_down_eq_zero_iff[of prec b n] by auto lemma power_down_mono: "power_down prec a n \ power_down prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_down.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_down prec a j \ b * power_down prec b j" by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_down_even_nonneg) then have "truncate_down (Suc prec) (a * power_down prec a j) \ truncate_down (Suc prec) (b * power_down prec b j)" by (auto intro!: truncate_down_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_down_simp) next assume [simp]: "odd j" have "power_down prec 0 (Suc (j div 2)) \ - power_down prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_down (Suc prec) ((power_down prec a (Suc (j div 2)))\<^sup>2) \ truncate_down (Suc prec) ((power_down prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_down_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_down_simp) qed qed simp qed lemma truncate_up_nonneg: "0 \ truncate_up p x" if "0 \ x" by (simp add: that truncate_up_le) lemma truncate_up_pos: "0 < truncate_up p x" if "0 < x" by (meson less_le_trans that truncate_up) lemma truncate_up_less_zero_iff[simp]: "truncate_up p x < 0 \ x < 0" proof - have f1: "\n r. truncate_up n r + truncate_down n (- 1 * r) = 0" by (simp add: truncate_down_uminus_eq) have f2: "(\v0 v1. truncate_up v0 v1 + truncate_down v0 (- 1 * v1) = 0) = (\v0 v1. truncate_up v0 v1 = - 1 * truncate_down v0 (- 1 * v1))" by (auto simp: truncate_up_eq_truncate_down) have f3: "\x1. ((0::real) < x1) = (\ x1 \ 0)" by fastforce have "(- 1 * x \ 0) = (0 \ x)" by force then have "0 \ x \ \ truncate_down p (- 1 * x) \ 0" using f3 by (meson truncate_down_pos) then have "(0 \ truncate_up p x) \ (\ 0 \ x)" using f2 f1 truncate_up_nonneg by force then show ?thesis by linarith qed lemma truncate_up_nonneg_iff[simp]: "truncate_up p x \ 0 \ x \ 0" using truncate_up_less_zero_iff[of p x] truncate_up_nonneg[of x] by linarith lemma power_up_even_nonneg: "even n \ 0 \ power_up p x n" by (induct p x n rule: power_up.induct) (auto simp: power_up.simps simp del: odd_Suc_div_two intro!: ) lemma truncate_up_eq_zero_iff[simp]: "truncate_up prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_up_less_zero_iff truncate_up_pos truncate_up_zero) lemma power_up_eq_zero_iff[simp]: "power_up prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_up_simp[of _ _ "x - 1"] by (cases x) (auto simp: sign_simps zero_le_mult_iff div2_less_self) qed lemma power_up_nonneg_iff[simp]: "power_up prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_up_simp[of _ _ "x - 1"] by (cases x) (auto simp: sign_simps zero_le_mult_iff) qed lemma power_up_neg_iff[simp]: "power_up prec b n < 0 \ b < 0 \ odd n" using power_up_nonneg_iff[of prec b n] by (auto simp del: power_up_nonneg_iff) lemma power_up_nonpos_iff[simp]: notes [simp del] = power_up_neg_iff power_up_eq_zero_iff shows "power_up prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_up_neg_iff[of prec b n] power_up_eq_zero_iff[of prec b n] by auto lemma power_up_mono: "power_up prec a n \ power_up prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_up.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_up prec a j \ b * power_up prec b j" by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_up_even_nonneg) then have "truncate_up prec (a * power_up prec a j) \ truncate_up prec (b * power_up prec b j)" by (auto intro!: truncate_up_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_up_simp) next assume [simp]: "odd j" have "power_up prec 0 (Suc (j div 2)) \ - power_up prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_up prec ((power_up prec a (Suc (j div 2)))\<^sup>2) \ truncate_up prec ((power_up prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_up_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_up_simp) qed qed simp qed lemma set_of_subset_iff: "set_of X \ set_of Y \ lower Y \ lower X \ upper X \ upper Y" for X Y::"'a::linorder interval" by (auto simp: set_of_eq subset_iff) lemma power_float_interval_mono: "set_of (power_float_interval prec n A) \ set_of (power_float_interval prec n B)" if "set_of A \ set_of B" proof - define la where "la = real_of_float (lower A)" define ua where "ua = real_of_float (upper A)" define lb where "lb = real_of_float (lower B)" define ub where "ub = real_of_float (upper B)" have ineqs: "lb \ la" "la \ ua" "ua \ ub" "lb \ ub" using that lower_le_upper[of A] lower_le_upper[of B] by (auto simp: la_def ua_def lb_def ub_def set_of_eq) show ?thesis using ineqs by (simp add: set_of_subset_iff float_power_bnds_def max_def power_down_fl.rep_eq power_up_fl.rep_eq la_def[symmetric] ua_def[symmetric] lb_def[symmetric] ub_def[symmetric]) (auto intro!: power_down_mono power_up_mono intro: order_trans[where y=0]) qed lemma bounds_of_interval_eq_lower_upper: "bounds_of_interval ivl = (lower ivl, upper ivl)" if "lower ivl \ upper ivl" using that by (auto simp: lower.rep_eq upper.rep_eq) lemma real_interval_Ivl: "real_interval (Ivl a b) = Ivl a b" by transfer (auto simp: min_def) lemma set_of_mul_contains_real_zero: "0 \\<^sub>r (A * B)" if "0 \\<^sub>r A \ 0 \\<^sub>r B" using that set_of_mul_contains_zero[of A B] by (auto simp: set_of_eq) fun subdivide_interval :: "nat \ float interval \ float interval list" where "subdivide_interval 0 I = [I]" | "subdivide_interval (Suc n) I = ( let m = mid I in (subdivide_interval n (Ivl (lower I) m)) @ (subdivide_interval n (Ivl m (upper I))) )" lemma subdivide_interval_length: shows "length (subdivide_interval n I) = 2^n" by(induction n arbitrary: I, simp_all add: Let_def) lemma lower_le_mid: "lower x \ mid x" "real_of_float (lower x) \ mid x" and mid_le_upper: "mid x \ upper x" "real_of_float (mid x) \ upper x" unfolding mid_def subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto done lemma subdivide_interval_correct: "list_ex (\i. x \\<^sub>r i) (subdivide_interval n I)" if "x \\<^sub>r I" for x::real using that proof(induction n arbitrary: x I) case 0 then show ?case by simp next case (Suc n) from \x \\<^sub>r I\ consider "x \\<^sub>r Ivl (lower I) (mid I)" | "x \\<^sub>r Ivl (mid I) (upper I)" by (cases "x \ real_of_float (mid I)") (auto simp: set_of_eq min_def lower_le_mid mid_le_upper) from this[case_names lower upper] show ?case by cases (use Suc.IH in \auto simp: Let_def\) qed fun interval_list_union :: "'a::lattice interval list \ 'a interval" where "interval_list_union [] = undefined" | "interval_list_union [I] = I" | "interval_list_union (I#Is) = sup I (interval_list_union Is)" lemma interval_list_union_correct: assumes "S \ []" assumes "i < length S" shows "set_of (S!i) \ set_of (interval_list_union S)" using assms proof(induction S arbitrary: i) case (Cons a S i) thus ?case proof(cases S) fix b S' assume "S = b # S'" hence "S \ []" by simp show ?thesis proof(cases i) case 0 show ?thesis apply(cases S) using interval_union_mono1 by (auto simp add: 0) next case (Suc i_prev) hence "i_prev < length S" using Cons(3) by simp from Cons(1)[OF \S \ []\ this] Cons(1) have "set_of ((a # S) ! i) \ set_of (interval_list_union S)" by (simp add: \i = Suc i_prev\) also have "... \ set_of (interval_list_union (a # S))" using \S \ []\ apply(cases S) using interval_union_mono2 by auto finally show ?thesis . qed qed simp qed simp lemma split_domain_correct: fixes x :: "real list" assumes "x all_in I" assumes split_correct: "\x a I. x \\<^sub>r I \ list_ex (\i::float interval. x \\<^sub>r i) (split I)" shows "list_ex (\s. x all_in s) (split_domain split I)" using assms(1) proof(induction I arbitrary: x) case (Cons I Is x) have "x \ []" using Cons(2) by auto obtain x' xs where x_decomp: "x = x' # xs" using \x \ []\ list.exhaust by auto hence "x' \\<^sub>r I" "xs all_in Is" using Cons(2) by auto show ?case using Cons(1)[OF \xs all_in Is\] split_correct[OF \x' \\<^sub>r I\] - apply(simp add: list_ex_iff set_of_eq) + apply (auto simp add: list_ex_iff set_of_eq) by (smt length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc x_decomp) qed simp end \ No newline at end of file diff --git a/thys/Winding_Number_Eval/Missing_Topology.thy b/thys/Winding_Number_Eval/Missing_Topology.thy --- a/thys/Winding_Number_Eval/Missing_Topology.thy +++ b/thys/Winding_Number_Eval/Missing_Topology.thy @@ -1,739 +1,740 @@ (* Author: Wenda Li *) section \Some useful lemmas in topology\ theory Missing_Topology imports "HOL-Analysis.Analysis" begin subsection \Misc\ lemma open_times_image: fixes S::"'a::real_normed_field set" assumes "open S" "c\0" shows "open (((*) c) ` S)" proof - let ?f = "\x. x/c" and ?g="((*) c)" have "continuous_on UNIV ?f" using \c\0\ by (auto intro:continuous_intros) then have "open (?f -` S)" using \open S\ by (auto elim:open_vimage) moreover have "?g ` S = ?f -` S" using \c\0\ apply auto using image_iff by fastforce ultimately show ?thesis by auto qed lemma image_linear_greaterThan: fixes x::"'a::linordered_field" assumes "c\0" shows "((\x. c*x+b) ` {x<..}) = (if c>0 then {c*x+b <..} else {..< c*x+b})" using \c\0\ apply (auto simp add:image_iff field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) done lemma image_linear_lessThan: fixes x::"'a::linordered_field" assumes "c\0" shows "((\x. c*x+b) ` {..0 then {..c\0\ apply (auto simp add:image_iff field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) done lemma continuous_on_neq_split: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes "\x\s. f x\y" "continuous_on s f" "connected s" shows "(\x\s. f x>y) \ (\x\s. f x f ` s" using assms(1) by blast then have "(aa \ s \ y < f aa) \ aaa \ s \ f aaa < y" by (meson Topological_Spaces.connected_continuous_image assms(2) assms(3) connectedD_interval image_eqI linorder_not_le) } then show ?thesis by blast qed lemma fixes f::"'a::linorder_topology \ 'b::topological_space" assumes "continuous_on {a..b} f" "aclosure ({..a} - {b})" using \a by auto then show ?thesis using at_within_eq_bot_iff by auto qed then have "(f \ f b) (at b within {..a})" by auto moreover have "(f \ f b) (at b within {a..b})" using assms unfolding continuous_on by auto moreover have "{..a} \ {a..b} = {..b}" using \a by auto ultimately have "(f \ f b) (at b within {..b})" using Lim_Un[of f "f b" b "{..a}" "{a..b}"] by presburger then have "(f \ f b) (at b within {..closure ({b..} - {a})" using \a by auto then show ?thesis using at_within_eq_bot_iff by auto qed then have "(f \ f a) (at a within {b..})" by auto moreover have "(f \ f a) (at a within {a..b})" using assms unfolding continuous_on by auto moreover have "{b..} \ {a..b} = {a..}" using \a by auto ultimately have "(f \ f a) (at a within {a..})" using Lim_Un[of f "f a" a "{b..}" "{a..b}"] by presburger then have "(f \ f a) (at a within {a<..})" apply (elim tendsto_within_subset) by auto then show "continuous (at_right a) f" using continuous_within by auto qed subsection \More about @{term eventually}\ lemma eventually_comp_filtermap: "eventually (P o f) F \ eventually P (filtermap f F)" unfolding comp_def using eventually_filtermap by auto lemma eventually_uminus_at_top_at_bot: fixes P::"'a::{ordered_ab_group_add,linorder} \ bool" shows "eventually (P o uminus) at_bot \ eventually P at_top" "eventually (P o uminus) at_top \ eventually P at_bot" unfolding eventually_comp_filtermap by (fold at_top_mirror at_bot_mirror,auto) lemma eventually_at_infinityI: fixes P::"'a::real_normed_vector \ bool" assumes "\x. c \ norm x \ P x" shows "eventually P at_infinity" unfolding eventually_at_infinity using assms by auto lemma eventually_at_bot_linorderI: fixes c::"'a::linorder" assumes "\x. x \ c \ P x" shows "eventually P at_bot" using assms by (auto simp: eventually_at_bot_linorder) lemma eventually_times_inverse_1: fixes f::"'a \ 'b::{field,t2_space}" assumes "(f \ c) F" "c\0" shows "\\<^sub>F x in F. inverse (f x) * f x = 1" proof - have "\\<^sub>F x in F. f x\0" using assms tendsto_imp_eventually_ne by blast then show ?thesis apply (elim eventually_mono) by auto qed subsection \More about @{term filtermap}\ lemma filtermap_linear_at_within: assumes "bij f" and cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (at a within S) = at (f a) within f`S" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (at a within S))" then obtain T where "open T" "a \ T" and impP:"\x\T. x\a \ x\S\ P (f x)" by (auto simp: eventually_filtermap eventually_at_topological) then show "eventually P (at (f a) within f ` S)" unfolding eventually_at_topological apply (intro exI[of _ "f`T"]) using \bij f\ open_map by (metis bij_pointE imageE imageI) next fix P assume "eventually P (at (f a) within f ` S)" then obtain T1 where "open T1" "f a \ T1" and impP:"\x\T1. x\f a \ x\f`S\ P (x)" unfolding eventually_at_topological by auto then obtain T2 where "open T2" "a \ T2" "(\x'\T2. f x' \ T1)" using cont[unfolded continuous_at_open,rule_format,of T1] by blast then have "\x\T2. x\a \ x\S\ P (f x)" using impP by (metis assms(1) bij_pointE imageI) then show "eventually P (filtermap f (at a within S))" unfolding eventually_filtermap eventually_at_topological apply (intro exI[of _ T2]) using \open T2\ \a \ T2\ by auto qed lemma filtermap_at_bot_linear_eq: fixes c::"'a::linordered_field" assumes "c\0" shows "filtermap (\x. x * c + b) at_bot = (if c>0 then at_bot else at_top)" proof (cases "c>0") case True then have "filtermap (\x. x * c + b) at_bot = at_bot" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_bot_linorder filterlim_at_bot - by (meson linordered_field_class.sign_simps(42) pos_divide_le_eq) + by (auto simp add: field_simps) subgoal unfolding eventually_at_bot_linorder filterlim_at_bot - by (metis linordered_field_class.sign_simps(24) real_affinity_le) + by (metis mult.commute real_affinity_le) by auto then show ?thesis using \c>0\ by auto next case False then have "c<0" using \c\0\ by auto then have "filtermap (\x. x * c + b) at_bot = at_top" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_top_linorder filterlim_at_bot by (meson le_diff_eq neg_divide_le_eq) subgoal unfolding eventually_at_bot_linorder filterlim_at_top - using linordered_field_class.sign_simps(42) neg_le_divide_eq by blast + using \c < 0\ by (meson False diff_le_eq le_divide_eq) by auto then show ?thesis using \c<0\ by auto qed lemma filtermap_linear_at_left: fixes c::"'a::{linordered_field,linorder_topology,real_normed_field}" assumes "c\0" shows "filtermap (\x. c*x+b) (at_left x) = (if c>0 then at_left (c*x+b) else at_right (c*x+b))" proof - let ?f = "\x. c*x+b" have "filtermap (\x. c*x+b) (at_left x) = (at (?f x) within ?f ` {..c\0\ by (auto intro!: o_bij[of "\x. (x-b)/c"]) show "isCont ?f x" by auto show "\S. open S \ open (?f ` S)" using open_times_image[OF _ \c\0\,THEN open_translation,of _ b] by (simp add:image_image add.commute) show "at (?f x) within ?f ` {..0" using image_linear_lessThan[OF \c\0\,of b x] that by auto moreover have "?f ` {.. c>0" using image_linear_lessThan[OF \c\0\,of b x] that by auto ultimately show ?thesis by auto qed lemma filtermap_linear_at_right: fixes c::"'a::{linordered_field,linorder_topology,real_normed_field}" assumes "c\0" shows "filtermap (\x. c*x+b) (at_right x) = (if c>0 then at_right (c*x+b) else at_left (c*x+b))" proof - let ?f = "\x. c*x+b" have "filtermap ?f (at_right x) = (at (?f x) within ?f ` {x<..})" proof (subst filtermap_linear_at_within) show "bij ?f" using \c\0\ by (auto intro!: o_bij[of "\x. (x-b)/c"]) show "isCont ?f x" by auto show "\S. open S \ open (?f ` S)" using open_times_image[OF _ \c\0\,THEN open_translation,of _ b] by (simp add:image_image add.commute) show "at (?f x) within ?f ` {x<..} = at (?f x) within ?f ` {x<..}" by simp qed moreover have "?f ` {x<..} = {?f x<..}" when "c>0" using image_linear_greaterThan[OF \c\0\,of b x] that by auto moreover have "?f ` {x<..} = {.. c>0" using image_linear_greaterThan[OF \c\0\,of b x] that by auto ultimately show ?thesis by auto qed lemma filtermap_at_top_linear_eq: fixes c::"'a::linordered_field" assumes "c\0" shows "filtermap (\x. x * c + b) at_top = (if c>0 then at_top else at_bot)" proof (cases "c>0") case True then have "filtermap (\x. x * c + b) at_top = at_top" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_top_linorder filterlim_at_top by (meson le_diff_eq pos_le_divide_eq) subgoal unfolding eventually_at_top_linorder filterlim_at_top - by (metis linordered_field_class.sign_simps(24) real_le_affinity) + apply auto + by (metis mult.commute real_le_affinity) by auto then show ?thesis using \c>0\ by auto next case False then have "c<0" using \c\0\ by auto then have "filtermap (\x. x * c + b) at_top = at_bot" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) - subgoal unfolding eventually_at_bot_linorder filterlim_at_top - by (meson linordered_field_class.sign_simps(42) neg_le_divide_eq) + subgoal unfolding eventually_at_bot_linorder filterlim_at_top + by (auto simp add: field_simps) subgoal unfolding eventually_at_top_linorder filterlim_at_bot by (meson le_diff_eq neg_divide_le_eq) by auto then show ?thesis using \c<0\ by auto qed lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (nhds a))" then obtain S where "open S" "a \ S" "\x\S. P (f x)" by (auto simp: eventually_filtermap eventually_nhds) then show "eventually P (nhds (f a))" unfolding eventually_nhds apply (intro exI[of _ "f`S"]) by (auto intro!: open_map) qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) subsection \More about @{term filterlim}\ lemma filterlim_at_infinity_times: fixes f :: "'a \ 'b::real_normed_field" assumes "filterlim f at_infinity F" "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ 0 * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at 0) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma filterlim_at_top_at_bot[elim]: fixes f::"'a \ 'b::unbounded_dense_linorder" and F::"'a filter" assumes top:"filterlim f at_top F" and bot: "filterlim f at_bot F" and "F\bot" shows False proof - obtain c::'b where True by auto have "\\<^sub>F x in F. c < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c" using bot unfolding filterlim_at_bot_dense by auto ultimately have "\\<^sub>F x in F. c < f x \ f x < c" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_top_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_top F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'>c" using gt_ex by blast have "\\<^sub>F x in F. c' < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c'" using order_tendstoD[OF tendsto,of c'] \c'>c\ by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_bot_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_bot F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'\<^sub>F x in F. c' > f x" using top unfolding filterlim_at_bot_dense by auto moreover have "\\<^sub>F x in F. f x > c'" using order_tendstoD[OF tendsto,of c'] \c' by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_top_linear_iff: fixes f::"'a::linordered_field \ 'b" assumes "c\0" shows "(LIM x at_top. f (x * c + b) :> F2) \ (if c>0 then (LIM x at_top. f x :> F2) else (LIM x at_bot. f x :> F2))" unfolding filterlim_def apply (subst filtermap_filtermap[of f "\x. x * c + b",symmetric]) using assms by (auto simp add:filtermap_at_top_linear_eq) lemma filterlim_at_bot_linear_iff: fixes f::"'a::linordered_field \ 'b" assumes "c\0" shows "(LIM x at_bot. f (x * c + b) :> F2) \ (if c>0 then (LIM x at_bot. f x :> F2) else (LIM x at_top. f x :> F2)) " unfolding filterlim_def apply (subst filtermap_filtermap[of f "\x. x * c + b",symmetric]) using assms by (auto simp add:filtermap_at_bot_linear_eq) lemma filterlim_tendsto_add_at_top_iff: assumes f: "(f \ c) F" shows "(LIM x F. (f x + g x :: real) :> at_top) \ (LIM x F. g x :> at_top)" proof assume "LIM x F. f x + g x :> at_top" moreover have "((\x. - f x) \ - c) F" using f by (intro tendsto_intros,simp) ultimately show "filterlim g at_top F" using filterlim_tendsto_add_at_top by fastforce qed (auto simp add:filterlim_tendsto_add_at_top[OF f]) lemma filterlim_tendsto_add_at_bot_iff: fixes c::real assumes f: "(f \ c) F" shows "(LIM x F. f x + g x :> at_bot) \ (LIM x F. g x :> at_bot)" proof - have "(LIM x F. f x + g x :> at_bot) \ (LIM x F. - f x + (- g x) :> at_top)" apply (subst filterlim_uminus_at_top) by (rule filterlim_cong,auto) also have "... = (LIM x F. - g x :> at_top)" apply (subst filterlim_tendsto_add_at_top_iff[of _ "-c"]) by (auto intro:tendsto_intros simp add:f) also have "... = (LIM x F. g x :> at_bot)" apply (subst filterlim_uminus_at_top) by (rule filterlim_cong,auto) finally show ?thesis . qed lemma tendsto_inverse_0_at_infinity: "LIM x F. f x :> at_infinity \ ((\x. inverse (f x) :: real) \ 0) F" by (metis filterlim_at filterlim_inverse_at_iff) lemma filterlim_at_infinity_divide_iff: fixes f::"'a \ 'b::real_normed_field" assumes "(f \ c) F" "c\0" shows "(LIM x F. f x / g x :> at_infinity) \ (LIM x F. g x :> at 0)" proof assume asm:"LIM x F. f x / g x :> at_infinity" have "LIM x F. inverse (f x) * (f x / g x) :> at_infinity" apply (rule tendsto_mult_filterlim_at_infinity[of _ "inverse c", OF _ _ asm]) by (auto simp add: assms(1) assms(2) tendsto_inverse) then have "LIM x F. inverse (g x) :> at_infinity" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono simp add:field_simps) then show "filterlim g (at 0) F" using filterlim_inverse_at_iff[symmetric] by force next assume "filterlim g (at 0) F" then have "filterlim (\x. inverse (g x)) at_infinity F" using filterlim_compose filterlim_inverse_at_infinity by blast then have "LIM x F. f x * inverse (g x) :> at_infinity" using tendsto_mult_filterlim_at_infinity[OF assms, of "\x. inverse(g x)"] by simp then show "LIM x F. f x / g x :> at_infinity" by (simp add: divide_inverse) qed lemma filterlim_tendsto_pos_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "0 < c" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_top)" proof assume "filterlim g at_top F" then show "LIM x F. f x * g x :> at_top" using filterlim_tendsto_pos_mult_at_top[OF assms] by auto next assume asm:"LIM x F. f x * g x :> at_top" have "((\x. inverse (f x)) \ inverse c) F" using tendsto_inverse[OF assms(1)] \0 by auto moreover have "inverse c >0" using assms(2) by auto ultimately have "LIM x F. inverse (f x) * (f x * g x) :> at_top" using filterlim_tendsto_pos_mult_at_top[OF _ _ asm,of "\x. inverse (f x)" "inverse c"] by auto then show "LIM x F. g x :> at_top" apply (elim filterlim_mono_eventually) apply simp_all[2] using eventually_times_inverse_1[OF assms(1)] \c>0\ eventually_mono by fastforce qed lemma filterlim_tendsto_pos_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 < c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_bot F" using filterlim_tendsto_pos_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "c < 0" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_bot)" proof - have "(LIM x F. f x * g x :> at_top) = (LIM x F. - g x :> at_top)" apply (rule filterlim_tendsto_pos_mult_at_top_iff[of "\x. - f x" "-c" F "\x. - g x", simplified]) using assms by (auto intro: tendsto_intros ) also have "... = (LIM x F. g x :> at_bot)" using filterlim_uminus_at_bot[symmetric] by auto finally show ?thesis . qed lemma filterlim_tendsto_neg_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 > c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_top F" using filterlim_tendsto_neg_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_top by simp lemma Lim_add: fixes f g::"_ \ 'a::{t2_space,topological_monoid_add}" assumes "\y. (f \ y) F" and "\y. (g \ y) F" and "F\bot" shows "Lim F f + Lim F g = Lim F (\x. f x+g x)" apply (rule tendsto_Lim[OF \F\bot\, symmetric]) apply (auto intro!:tendsto_eq_intros) using assms tendsto_Lim by blast+ (* lemma filterlim_at_top_tendsto[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_top F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain cc where "cc>c" using gt_ex by blast have "\\<^sub>F x in F. cc < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < cc" using tendsto order_tendstoD(2)[OF _ \cc>c\] by auto ultimately have "\\<^sub>F x in F. cc < f x \ f x < cc" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_bot_tendsto[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_bot F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain cc where "cc\<^sub>F x in F. cc > f x" using top unfolding filterlim_at_bot_dense by auto moreover have "\\<^sub>F x in F. f x > cc" using tendsto order_tendstoD(1)[OF _ \cc] by auto ultimately have "\\<^sub>F x in F. cc < f x \ f x < cc" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed *) subsection \Isolate and discrete\ definition (in topological_space) isolate:: "'a \ 'a set \ bool" (infixr "isolate" 60) where "x isolate S \ (x\S \ (\T. open T \ T \ S = {x}))" definition (in topological_space) discrete:: "'a set \ bool" where "discrete S \ (\x\S. x isolate S)" definition (in metric_space) uniform_discrete :: "'a set \ bool" where "uniform_discrete S \ (\e>0. \x\S. \y\S. dist x y < e \ x = y)" lemma uniformI1: assumes "e>0" "\x y. \x\S;y\S;dist x y \ x =y " shows "uniform_discrete S" unfolding uniform_discrete_def using assms by auto lemma uniformI2: assumes "e>0" "\x y. \x\S;y\S;x\y\ \ dist x y\e " shows "uniform_discrete S" unfolding uniform_discrete_def using assms not_less by blast lemma isolate_islimpt_iff:"(x isolate S) \ (\ (x islimpt S) \ x\S)" unfolding isolate_def islimpt_def by auto lemma isolate_dist_Ex_iff: fixes x::"'a::metric_space" shows "x isolate S \ (x\S \ (\e>0. \y\S. dist x y < e \ y=x))" unfolding isolate_islimpt_iff islimpt_approachable by (metis dist_commute) lemma discrete_empty[simp]: "discrete {}" unfolding discrete_def by auto lemma uniform_discrete_empty[simp]: "uniform_discrete {}" unfolding uniform_discrete_def by (simp add: gt_ex) lemma isolate_insert: fixes x :: "'a::t1_space" shows "x isolate (insert a S) \ x isolate S \ (x=a \ \ (x islimpt S))" by (meson insert_iff islimpt_insert isolate_islimpt_iff) (* TODO. Other than uniform_discrete S \ discrete S uniform_discrete S \ closed S , we should be able to prove discrete S \ closed S \ uniform_discrete S but the proof (based on Tietze Extension Theorem) seems not very trivial to me. Informal proofs can be found in http://topology.auburn.edu/tp/reprints/v30/tp30120.pdf http://msp.org/pjm/1959/9-2/pjm-v9-n2-p19-s.pdf *) lemma uniform_discrete_imp_closed: "uniform_discrete S \ closed S" by (meson discrete_imp_closed uniform_discrete_def) lemma uniform_discrete_imp_discrete: "uniform_discrete S \ discrete S" by (metis discrete_def isolate_dist_Ex_iff uniform_discrete_def) lemma isolate_subset:"x isolate S \ T \ S \ x\T \ x isolate T" unfolding isolate_def by fastforce lemma discrete_subset[elim]: "discrete S \ T \ S \ discrete T" unfolding discrete_def using islimpt_subset isolate_islimpt_iff by blast lemma uniform_discrete_subset[elim]: "uniform_discrete S \ T \ S \ uniform_discrete T" by (meson subsetD uniform_discrete_def) lemma continuous_on_discrete: "discrete S \ continuous_on S f" unfolding continuous_on_topological by (metis discrete_def islimptI isolate_islimpt_iff) (* Is euclidean_space really necessary?*) lemma uniform_discrete_insert: fixes S :: "'a::euclidean_space set" shows "uniform_discrete (insert a S) \ uniform_discrete S" proof assume asm:"uniform_discrete S" let ?thesis = "uniform_discrete (insert a S)" have ?thesis when "a\S" using that asm by (simp add: insert_absorb) moreover have ?thesis when "S={}" using that asm by (simp add: uniform_discrete_def) moreover have ?thesis when "a\S" "S\{}" proof - obtain e1 where "e1>0" and e1_dist:"\x\S. \y\S. dist y x < e1 \ y = x" using asm unfolding uniform_discrete_def by auto define e2 where "e2 \ min (setdist {a} S) e1" have "closed S" using asm uniform_discrete_imp_closed by auto then have "e2>0" by (simp add: \0 < e1\ e2_def setdist_gt_0_compact_closed that(1) that(2)) moreover have "x = y" when "x\insert a S" "y\insert a S" "dist x y < e2 " for x y proof - have ?thesis when "x=a" "y=a" using that by auto moreover have ?thesis when "x=a" "y\S" using that setdist_le_dist[of x "{a}" y S] \dist x y < e2\ unfolding e2_def by fastforce moreover have ?thesis when "y=a" "x\S" using that setdist_le_dist[of y "{a}" x S] \dist x y < e2\ unfolding e2_def by (simp add: dist_commute) moreover have ?thesis when "x\S" "y\S" using e1_dist[rule_format, OF that] \dist x y < e2\ unfolding e2_def by (simp add: dist_commute) ultimately show ?thesis using that by auto qed ultimately show ?thesis unfolding uniform_discrete_def by meson qed ultimately show ?thesis by auto qed (simp add: subset_insertI uniform_discrete_subset) lemma discrete_compact_finite_iff: fixes S :: "'a::t1_space set" shows "discrete S \ compact S \ finite S" proof assume "finite S" then have "compact S" using finite_imp_compact by auto moreover have "discrete S" unfolding discrete_def using isolate_islimpt_iff islimpt_finite[OF \finite S\] by auto ultimately show "discrete S \ compact S" by auto next assume "discrete S \ compact S" then show "finite S" by (meson discrete_def Heine_Borel_imp_Bolzano_Weierstrass isolate_islimpt_iff order_refl) qed lemma uniform_discrete_finite_iff: fixes S :: "'a::heine_borel set" shows "uniform_discrete S \ bounded S \ finite S" proof assume "uniform_discrete S \ bounded S" then have "discrete S" "compact S" using uniform_discrete_imp_discrete uniform_discrete_imp_closed compact_eq_bounded_closed by auto then show "finite S" using discrete_compact_finite_iff by auto next assume asm:"finite S" let ?thesis = "uniform_discrete S \ bounded S" have ?thesis when "S={}" using that by auto moreover have ?thesis when "S\{}" proof - have "\x. \d>0. \y\S. y \ x \ d \ dist x y" using finite_set_avoid[OF \finite S\] by auto then obtain f where f_pos:"f x>0" and f_dist: "\y\S. y \ x \ f x \ dist x y" if "x\S" for x by metis define f_min where "f_min \ Min (f ` S)" have "f_min > 0" unfolding f_min_def by (simp add: asm f_pos that) moreover have "\x\S. \y\S. f_min > dist x y \ x=y" using f_dist unfolding f_min_def by (metis Min_gr_iff all_not_in_conv asm dual_order.irrefl eq_iff finite_imageI imageI less_eq_real_def) ultimately have "uniform_discrete S" unfolding uniform_discrete_def by auto moreover have "bounded S" using \finite S\ by auto ultimately show ?thesis by auto qed ultimately show ?thesis by blast qed lemma uniform_discrete_image_scale: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "uniform_discrete S" and dist:"\x\S. \y\S. dist x y = c * dist (f x) (f y)" shows "uniform_discrete (f ` S)" proof - have ?thesis when "S={}" using that by auto moreover have ?thesis when "S\{}" "c\0" proof - obtain x1 where "x1\S" using \S\{}\ by auto have ?thesis when "S-{x1} = {}" proof - have "S={x1}" using that \S\{}\ by auto then show ?thesis using uniform_discrete_insert[of "f x1"] by auto qed moreover have ?thesis when "S-{x1} \ {}" proof - obtain x2 where "x2\S-{x1}" using \S-{x1} \ {}\ by auto then have "x2\S" "x1\x2" by auto then have "dist x1 x2 > 0" by auto moreover have "dist x1 x2 = c * dist (f x1) (f x2)" using dist[rule_format, OF \x1\S\ \x2\S\] . moreover have "dist (f x2) (f x2) \ 0" by auto ultimately have False using \c\0\ by (simp add: zero_less_mult_iff) then show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "S\{}" "c>0" proof - obtain e1 where "e1>0" and e1_dist:"\x\S. \y\S. dist y x < e1 \ y = x" using \uniform_discrete S\ unfolding uniform_discrete_def by auto define e where "e= e1/c" have "x1 = x2" when "x1\ f ` S" "x2\ f ` S" "dist x1 x2 < e " for x1 x2 proof - obtain y1 where y1:"y1\S" "x1=f y1" using \x1\ f ` S\ by auto obtain y2 where y2:"y2\S" "x2=f y2" using \x2\ f ` S\ by auto have "dist y1 y2 < e1" using dist[rule_format, OF y1(1) y2(1)] \c>0\ \dist x1 x2 < e\ unfolding e_def apply (fold y1(2) y2(2)) by (auto simp add:divide_simps mult.commute) then have "y1=y2" using e1_dist[rule_format, OF y2(1) y1(1)] by simp then show "x1=x2" using y1(2) y2(2) by auto qed moreover have "e>0" using \e1>0\ \c>0\ unfolding e_def by auto ultimately show ?thesis unfolding uniform_discrete_def by meson qed ultimately show ?thesis by fastforce qed end diff --git a/thys/Winding_Number_Eval/Winding_Number_Eval_Examples.thy b/thys/Winding_Number_Eval/Winding_Number_Eval_Examples.thy --- a/thys/Winding_Number_Eval/Winding_Number_Eval_Examples.thy +++ b/thys/Winding_Number_Eval/Winding_Number_Eval_Examples.thy @@ -1,97 +1,97 @@ (* Author: Wenda Li *) section \Some examples of applying the method winding\_eval\ theory Winding_Number_Eval_Examples imports Winding_Number_Eval begin lemma example1: assumes "R>1" shows "winding_number (part_circlepath 0 R 0 pi +++ linepath (-R) R) \ = 1" proof (eval_winding,simp_all) define CR where "CR \part_circlepath 0 R 0 pi" define L where "L\ linepath (- (complex_of_real R)) R" show "\ \ path_image CR" unfolding CR_def using \R>1\ by (intro not_on_circlepathI,auto) show *:"\ \ closed_segment (- (of_real R)) R" using \R>1\ complex_eq_iff by (intro not_on_closed_segmentI,auto) from cindex_pathE_linepath[OF this] have "cindex_pathE L \ = -1" unfolding L_def using \R>1\ by auto moreover have "cindex_pathE CR \ = -1" unfolding CR_def using \R>1\ apply (subst cindex_pathE_part_circlepath) by (simp_all add:jumpF_pathstart_part_circlepath jumpF_pathfinish_part_circlepath) ultimately show "- complex_of_real (cindex_pathE CR \) - cindex_pathE L \ = 2" unfolding L_def CR_def by auto qed lemma example2: assumes "R>1" shows "winding_number (part_circlepath 0 R 0 pi +++ linepath (-R) R) (-\) = 0" proof (eval_winding,simp_all) define CR where "CR \part_circlepath 0 R 0 pi" define L where "L\ linepath (- (complex_of_real R)) R" show "-\ \ path_image CR" unfolding CR_def using \R>1\ by (intro not_on_circlepathI,auto) show *:"-\ \ closed_segment (- (of_real R)) R" using \R>1\ complex_eq_iff by (intro not_on_closed_segmentI,auto) from cindex_pathE_linepath[OF this] have "cindex_pathE L (-\) = 1" unfolding L_def using \R>1\ by auto moreover have "cindex_pathE CR (-\) = -1" unfolding CR_def using \R>1\ apply (subst cindex_pathE_part_circlepath) by (simp_all add:jumpF_pathstart_part_circlepath jumpF_pathfinish_part_circlepath) ultimately show "-cindex_pathE CR (-\) = cindex_pathE L (-\)" unfolding L_def CR_def by auto qed lemma example3: fixes lb ub z :: complex defines "rec \ linepath lb (Complex (Re ub) (Im lb)) +++ linepath (Complex (Re ub) (Im lb)) ub +++ linepath ub (Complex (Re lb) (Im ub)) +++ linepath (Complex (Re lb) (Im ub)) lb" assumes order_asms:"Re lb < Re z" "Re z < Re ub" "Im lb < Im z" "Im z < Im ub" shows "winding_number rec z = 1" unfolding rec_def proof (eval_winding) let ?l1 = "linepath lb (Complex (Re ub) (Im lb))" and ?l2 = "linepath (Complex (Re ub) (Im lb)) ub" and ?l3 = "linepath ub (Complex (Re lb) (Im ub))" and ?l4 = "linepath (Complex (Re lb) (Im ub)) lb" show l1: "z \ path_image ?l1" apply (auto intro!: not_on_closed_segmentI_complex) - using order_asms by (metis diff_gt_0_iff_gt mult_pos_pos order.asym zero_less_mult_pos2) + using order_asms by (simp add: algebra_simps crossproduct_eq) show l2:"z \ path_image ?l2" apply (auto intro!: not_on_closed_segmentI_complex) - using order_asms by (metis diff_gt_0_iff_gt linordered_field_class.sign_simps(44) order.asym) + using order_asms by (simp add: algebra_simps crossproduct_eq) show l3:"z \ path_image ?l3" apply (auto intro!: not_on_closed_segmentI_complex) - using order_asms by (metis diff_less_0_iff_less linordered_field_class.sign_simps(44) order.asym) + using order_asms by (simp add: algebra_simps crossproduct_eq) show l4:"z \ path_image ?l4" apply (auto intro!: not_on_closed_segmentI_complex) - using order_asms by (metis diff_less_0_iff_less linordered_field_class.sign_simps(44) order.asym) + using order_asms by (simp add: algebra_simps crossproduct_eq) show "- complex_of_real (cindex_pathE ?l1 z + (cindex_pathE ?l2 z + (cindex_pathE ?l3 z + cindex_pathE ?l4 z))) = 2 * 1" proof - have "(Im z - Im ub) * (Re ub - Re lb) < 0" using mult_less_0_iff order_asms(1) order_asms(2) order_asms(4) by fastforce then have "cindex_pathE ?l3 z = -1" apply (subst cindex_pathE_linepath) using l3 order_asms by (auto simp add:algebra_simps) moreover have "(Im lb - Im z) * (Re ub - Re lb) <0" using mult_less_0_iff order_asms(1) order_asms(2) order_asms(3) by fastforce then have "cindex_pathE ?l1 z = -1" apply (subst cindex_pathE_linepath) using l1 order_asms by (auto simp add:algebra_simps) moreover have "cindex_pathE ?l2 z = 0" apply (subst cindex_pathE_linepath) using l2 order_asms by (auto simp add:algebra_simps) moreover have "cindex_pathE ?l4 z = 0" apply (subst cindex_pathE_linepath) using l4 order_asms by (auto simp add:algebra_simps) ultimately show ?thesis by auto qed qed end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,5729 +1,5728 @@ (* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Complex_Main Word_Next Word_Enum "HOL-Library.Sublist" begin text \Set up quickcheck to support words\ quickcheck_generator word constructors: "zero_class.zero :: ('a::len) word", "numeral :: num \ ('a::len) word", "uminus :: ('a::len) word \ ('a::len) word" instantiation Enum.finite_1 :: len begin definition "len_of_finite_1 (x :: Enum.finite_1 itself) \ (1 :: nat)" instance by (standard, auto simp: len_of_finite_1_def) end instantiation Enum.finite_2 :: len begin definition "len_of_finite_2 (x :: Enum.finite_2 itself) \ (2 :: nat)" instance by (standard, auto simp: len_of_finite_2_def) end instantiation Enum.finite_3 :: len begin definition "len_of_finite_3 (x :: Enum.finite_3 itself) \ (4 :: nat)" instance by (standard, auto simp: len_of_finite_3_def) end (* Provide wf and less_induct for word. wf may be more useful in loop proofs, less_induct in recursion proofs. *) lemma word_less_wf: "wf {(a, b). a < (b :: ('a::len) word)}" apply (rule wf_subset) apply (rule wf_measure) apply safe apply (subst in_measure) apply (erule unat_mono) done lemma word_less_induct: "\ \x::('a::len) word. (\y. y < x \ P y) \ P x \ \ P a" using word_less_wf by induct blast instantiation word :: (len) wellorder begin instance by (intro_classes) (metis word_less_induct) end lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" assumes suc_m: "Suc m < LENGTH('a::len)" assumes 2: "unat (2::'a::len word) = 2" shows "2^n < (2::'a::len word)^m" proof - from suc_n have "(2::nat) * 2 ^ n mod 2 ^ LENGTH('a::len) = 2 * 2^n" apply (subst mod_less) apply (subst power_Suc[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done moreover from suc_m have "(2::nat) * 2 ^ m mod 2 ^ LENGTH('a::len) = 2 * 2^m" apply (subst mod_less) apply (subst power_Suc[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done ultimately have "2 * 2 ^ n < (2::nat) * 2 ^ m" using x apply (unfold word_less_nat_alt) apply simp apply (subst (asm) unat_word_ariths(2))+ apply (subst (asm) 2)+ apply (subst (asm) word_unat_power, subst (asm) unat_of_nat)+ apply (simp add: mod_mult_right_eq) done with suc_n suc_m show ?thesis unfolding word_less_nat_alt apply (subst word_unat_power, subst unat_of_nat)+ apply simp done qed lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" assumes 2: "unat (2::'a::len word) = 2" shows "x < y" using x apply (induct x arbitrary: y) apply (case_tac y; simp) apply (case_tac y; clarsimp) apply (subst (asm) power_Suc [symmetric]) apply (subst (asm) p2_eq_0) apply simp apply (drule (2) word_2p_mult_inc, rule 2) apply simp done lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" apply (simp add: no_plus_overflow_uint_size word_less_alt uint_word_ariths word_size) apply (subst(asm) zmod_zminus1_eq_if) apply (simp split: if_split_asm) done lemma ucast_ucast_eq: fixes x :: "'a::len word" fixes y :: "'b::len word" shows "\ ucast x = (ucast (ucast y::'a::len word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe, simp_all add: word_size to_bl_nth) done lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0") apply simp apply simp apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less) apply simp+ done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" proof - have "Suc (unat ((2::'a word) ^ n - 1)) = unat ((2::'a word) ^ n)" using nv by (metis Suc_pred' power_2_ge_iff unat_gt_0 unat_minus_one word_not_simps(1)) then show ?thesis using nv apply - apply (subst word_le_nat_alt) apply (subst less_Suc_eq_le [symmetric]) apply (erule ssubst) apply (subst word_less_nat_alt) apply (rule refl) done qed lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_unat.Rep_eqD) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < len_of (TYPE('a))" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "\x = z; x < y; Suc (unat y) < 2 ^ LENGTH('a)\ \ [x::'a::len word .e. y] = z # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons) apply (simp) apply (drule unat_mono) apply arith apply (simp only: list.map) apply (subst list.inject) apply rule apply (rule to_from_enum) apply (subst upto_enum_red) apply (rule map_cong [OF _ refl]) apply (rule arg_cong2 [where f = "\x y. [x ..< y]"]) apply unat_arith apply simp done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv by simp qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma mask_shift: "(x && ~~ mask y) >> y = x >> y" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply safe apply (drule test_bit.Rep[simplified, rule_format]) apply (simp add: word_size word_ops_nth_size) done lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_def) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_unat.Rep_eqD) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply (simp add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof cases assume szv: "sz < LENGTH('a::len)" show ?thesis proof (cases "sz = 0") case True then show ?thesis using kv szv by (simp add: unat_of_nat) next case False then have sne: "0 < sz" .. have uk: "unat (of_nat k :: 'a word) = k" apply (subst unat_of_nat) apply (simp add: nat_mod_eq less_trans[OF kv] sne) done show ?thesis using szv apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ done qed next assume "\ sz < LENGTH('a)" with kv show ?thesis by (simp add: not_less power_overflow) qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_def word_mod_by_0) qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_def power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (rule word_eqI) apply (simp add: word_size nth_ucast) apply safe apply (simp add: test_bit.Rep[simplified]) done lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (drule less_mask_eq) apply (rule word_eqI) apply (drule_tac x=n in word_eqD) apply (clarsimp simp: word_size nth_ucast) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" apply (simp add: word_less_nat_alt) apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply simp+ done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp apply assumption done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" apply (erule udvd_minus_le') apply (simp add: udvd_def)+ done lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by (subst unat_of_nat_eq[where x=n, symmetric], simp+) lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: unat_of_nat) done lemma of_nat_0: "\of_nat n = (0::('a::len) word); n < 2 ^ len_of (TYPE('a))\ \ n = 0" by (drule unat_of_nat_eq, simp) lemma minus_one_helper3: "x < y \ x \ (y :: 'a :: len word) - 1" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply clarsimp apply arith done lemma minus_one_helper: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma minus_one_helper5: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" apply (rule notI) apply (erule is_aligned_get_word_bits[where p=y]) apply (simp add: eq_diff_eq[symmetric]) apply (frule minus_one_helper3) apply (drule le_minus'[where a="x" and c="y - x" and b="- 1" for x y, simplified]) apply (simp add: field_simps) apply (frule is_aligned_less_sz[where a=y]) apply clarsimp apply (erule notE) apply (rule minus_one_helper5) apply simp apply (metis is_aligned_no_overflow minus_one_helper3 order_le_less_trans) apply simp done lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by fact next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < len_of (TYPE('a)); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma shiftr_less_t2n': fixes x :: "'a :: len word" shows "\ x && mask (n + m) = x; m < LENGTH('a) \ \ (x >> n) < 2 ^ m" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (rule word_eqI) apply (drule_tac x="na + n" in word_eqD) apply (simp add: nth_shiftr word_size) apply safe done lemma shiftr_less_t2n: fixes x :: "'a :: len word" shows "x < 2 ^ (n + m) \ (x >> n) < 2 ^ m" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ mask m = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_def le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (drule less_mask_eq) apply (rule word_eqI) apply (drule_tac x="na - n" in word_eqD) apply (simp add: nth_shiftl word_size) apply (cases "n \ m") apply safe apply simp apply simp done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" apply (rule word_eqI) apply (simp add: nth_ucast word_size) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma unat_ucast: "unat (ucast x :: ('a :: len0) word) = unat x mod 2 ^ (LENGTH('a))" apply (simp add: unat_def ucast_def) apply (subst word_uint.eq_norm) apply (subst nat_mod_distrib) apply simp apply simp apply (subst nat_power_eq) apply simp apply simp done lemma ucast_less_ucast: "LENGTH('a) < LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (simp add: scast_def) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: nth_rev) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" apply (rule word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule word_eqI) apply (clarsimp simp add: is_aligned_nth) apply (frule(1) nth_bounded) apply simp+ done lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule minus_one_helper) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ mask n) = x - (x && mask n)" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_def word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ mask n) && ~~ mask m = w && ~~ mask (max m n)" apply (simp add: and_not_mask) apply (case_tac "n x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by (rule word_eqI) (simp add: word_size) lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_def] shiftl_shiftr2[unfolded word_size mask_def] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" apply safe apply (rule word_eqI) apply (drule_tac x=n in word_eqD)+ by (auto simp: word_size word_ops_nth_size) lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n) apply simp apply (case_tac n) apply simp_all done lemma word_or_zero: "(a || b = 0) = (a = 0 \ b = 0)" apply (safe, simp_all) apply (rule word_eqI, drule_tac x=n in word_eqD, simp)+ done lemma word_and_1_shiftl: fixes x :: "'a :: len word" shows "x && (1 << n) = (if x !! n then (1 << n) else 0)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl del: shiftl_1) apply auto done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x && (mask n << m) = ((x >> m) && mask n) << m" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply auto done lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = mask 8" by (simp add: mask_def) lemma word_1FF_is_mask: "0x1FF = mask 9" by (simp add: mask_def) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule minus_one_helper3) apply (rule disjCI) apply simp apply simp apply (erule minus_one_helper5) apply fastforce done lemma minus_one_helper2: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed lemma word_div_less: fixes m :: "'a :: len word" shows "m < n \ m div n = 0" apply (rule word_unat.Rep_eqD) apply (simp add: word_less_nat_alt unat_div) done lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" apply (rule ccontr) apply (drule(1) order_trans) apply (drule word_sub_1_le) apply (drule(1) order_antisym) apply simp done lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" apply (subgoal_tac "a \ {a .. b}") apply (frule(1) range_subset_lower) apply (frule(1) range_subset_upper) apply (rule context_conjI, simp) apply (rule word_sub_mono, assumption+) apply (erule word_sub_le) apply (erule word_sub_le) apply simp done lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp apply simp done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemmas if_fun_split = if_apply_def2 lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply (safe, simp_all) apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp add: word_less_nat_alt) apply (subst unat_of_nat_eq) apply (erule order_less_trans) apply simp+ done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" apply (rule word_eqI) apply (simp add: nth_shiftr cong: rev_conj_cong) done lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (induct x) apply clarsimp+ by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt) lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w && mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma shiftr_mask_eq: fixes x :: "'a :: len word" shows "(x >> n) && mask (size x - n) = x >> n" apply (rule word_eqI) apply (clarsimp simp: word_size nth_shiftr) apply (rule iffI) apply clarsimp apply (clarsimp) apply (drule test_bit_size) apply (simp add: word_size) done lemma shiftr_mask_eq': fixes x :: "'a :: len word" shows "m = (size x - n) \ (x >> n) && mask m = x >> n" by (simp add: shiftr_mask_eq) lemma bang_big: "n \ size (x::'a::len0 word) \ (x !! n) = False" by (simp add: test_bit_bl word_size) lemma bang_conj_lt: fixes x :: "'a :: len word" shows "(x !! m \ m < LENGTH('a)) = x !! m" apply (cases "m < LENGTH('a)") apply simp apply (simp add: not_less bang_big word_size) done lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply (simp add: word_unat.norm_eq_iff [symmetric]) done lemma neg_mask_bang: "(~~ mask n :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" apply (cases "m < LENGTH('a)") apply (simp add: word_ops_nth_size word_size not_less) apply (simp add: not_less bang_big word_size) done lemma mask_AND_NOT_mask: "(w && ~~ mask n) && mask n = 0" by (rule word_eqI) (clarsimp simp add: word_size neg_mask_bang) lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ mask n) + (w && mask n) = w" apply (rule word_eqI) apply (rename_tac m) apply (simp add: word_size) apply (cut_tac word_plus_and_or_coroll[of "w && ~~ mask n" "w && mask n"]) apply (simp add: word_ao_dist2[symmetric] word_size neg_mask_bang) apply (rule word_eqI) apply (rename_tac m) apply (simp add: word_size neg_mask_bang) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ mask n = y && ~~ mask n" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size bang_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size bang_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ mask n) !! m)" by (simp add: neg_mask_bang bang_conj_lt) also have "\ = ((y && ~~ mask n) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_bang bang_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: unat_of_nat) done lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply simp done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply (rule word_eqI) apply (simp add: nth_w2p) done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = 1" by (simp add: mask_def) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma neg_mask_mono_le: "(x :: 'a :: len word) \ y \ x && ~~ mask n \ y && ~~ mask n" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False show "y && ~~ mask n < x && ~~ mask n \ False" using False by (simp add: mask_def linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y && ~~ mask n < x && ~~ mask n" have word_bits: "n < LENGTH('a)" using True by assumption have "y \ (y && ~~ mask n) + (y && mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y && ~~ mask n) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp_all add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b) apply (simp_all add: is_aligned_neg_mask) done also have "\ \ x && ~~ mask n" using b apply - apply (subst add.commute, rule le_plus) apply (rule aligned_at_least_t2n_diff, simp_all add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated], simp_all add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" apply (rule iffI) prefer 2 apply simp apply (subgoal_tac "x && mask 1 < 2^1") prefer 2 apply (rule and_mask_less_size) apply (simp add: word_size) apply (simp add: mask_def) apply (drule word_less_cases [where y=2]) apply (erule disjE, simp) apply simp done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" by (simp add: scast_def ucast_def sint_eq_uint) lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" by unat_arith lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: "\ \ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P; \ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P; \ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P \ \ P" apply atomize_elim apply (case_tac "len_of TYPE ('a) = 1") apply clarsimp apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") apply (metis UNIV_I insert_iff singletonE) apply (subst word_unat.univ) apply (clarsimp simp: unats_def image_def) apply (rule set_eqI, rule iffI) apply clarsimp apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) apply clarsimp apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) apply clarsimp apply (metis One_nat_def arith_is_1 le_def len_gt_0) done lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply clarsimp apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0) done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply (metis word_ubin.eq_norm) done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) apply (metis word_ubin.eq_norm) done (* Signed word arithmetic overflow constraints. *) lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply (simp add: sint_word_ariths scast_def) apply (simp add: wi_hom_mult) apply (subst word_sint.Abs_inject, simp_all) apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp apply (simp add: sints_def range_sbintrunc word_size) apply (auto elim: order_less_le_trans order_trans[rotated]) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") - apply (clarsimp simp: sgn_if sign_simps) + apply (clarsimp simp: sgn_if) apply (clarsimp simp: sign_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") - apply (clarsimp simp: sign_simps) + apply clarsimp apply (rule classical) - apply (clarsimp simp: sign_simps sgn_mult not_less) + apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: sign_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if sign_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less sign_simps) apply (clarsimp simp: sign_simps not_less) apply (rule classical) apply (case_tac "b = 0") - apply (clarsimp simp: sign_simps not_less sgn_mult) + apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") - apply (clarsimp simp: sign_simps not_less sgn_mult) + apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) - apply (clarsimp simp: sign_simps not_less sgn_mult) + apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if sign_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] - apply (clarsimp simp: smod_int_alt_def sign_simps sgn_if - abs_if not_less add1_zle_eq [simplified add.commute]) - apply (metis add_le_cancel_left monoid_add_class.add.right_neutral - int_one_le_iff_zero_less less_le_trans mod_minus_right neg_less_0_iff_less - neg_mod_conj not_less pos_mod_conj) + apply (auto simp: smod_int_alt_def algebra_simps sgn_if + abs_if not_less add1_zle_eq [simplified add.commute])[1] + apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) + apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] - apply (clarsimp simp: smod_int_alt_def sign_simps sgn_if + apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (case_tac "a \ - (2 ^ (LENGTH('a) - 1)) \ b \ -1") apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric]) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" apply (clarsimp simp: word_of_int int_word_sint) apply (subst int_mod_eq') apply simp apply (subst (2) power_minus_simp) apply clarsimp apply clarsimp apply clarsimp done lemma of_int_sint [simp]: "of_int (sint a) = a" apply (insert word_sint.Rep [where x=a]) apply (clarsimp simp: word_of_int) done lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_def bang_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (clarsimp simp: word_of_int ucast_def) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_of_int) apply (subst word.abs_eq_iff) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_def sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (clarsimp simp: scast_def) apply (metis down_cast_same is_up_down scast_def ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_bl ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred bang_conj_lt len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) done lemma word_less_nowrapI': "(x :: 'a :: len0 word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = 2 ^ n" by (clarsimp simp: mask_def) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp add: ucast_nat_def [symmetric]) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp add: ucast_nat_def [symmetric]) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) bang_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (len_of (TYPE('a)) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" apply (simp add: to_bool_def) apply (rule iffI) apply (rule classical, erule notE, rule word_eqI) apply clarsimp apply (case_tac n, simp_all)[1] apply (rule notI, drule word_eqD[where x=0]) apply simp done lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr nth_rev field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: is_aligned_nth) apply (drule(1) nth_bounded) apply simp apply simp apply (rule word_eqI) apply (simp add: nth_shiftr) apply safe apply (drule(1) nth_bounded) apply simp+ done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: is_aligned_nth) apply (drule(1) nth_bounded) apply simp apply simp apply (rule word_eqI) apply (simp add: nth_shiftr) apply safe apply (drule(1) nth_bounded) apply simp+ done lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ mask n) + (2 ^ n - 1) = x || mask n" apply (simp add:mask_2pm1[symmetric]) apply (rule word_eqI[rule_format]) apply (rule iffI) apply (clarsimp simp:word_size not_less) apply (cut_tac w = "((x && ~~ mask n) + mask n)" and m = n and n = "na - n" in nth_shiftr[symmetric]) apply clarsimp apply (subst (asm) aligned_shift') apply (simp add:mask_lt_2pn nth_shiftr is_aligned_neg_mask word_size)+ apply (case_tac "na p" apply (rule word_le_minus_cancel[where x = "p && ~~ mask n"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_def word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ mask n) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule minus_one_helper3) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ mask n) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (drule_tac x=na in word_eqD)+ apply (simp add: word_size) apply blast apply (rule word_eqI) apply (drule_tac x=na in word_eqD)+ apply (simp add: word_ops_nth_size word_size) apply blast apply (insert word_plus_and_or_coroll2[where x="p + d" and w="mask n"]) apply simp done lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ mask n) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" apply (rule word_eqI) apply (simp add: word_size conj_comms) done lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ mask m = (ptr + a && ~~ mask m ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: minus_one_helper3) lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ mask n = 0" apply (clarsimp simp:mask_def) by (metis word_bool_alg.conj_cancel_right word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" apply (clarsimp simp:mask_def) by (metis (erased, hide_lams) mask_eq_x_eq_0 shiftl_over_and_dist word_bool_alg.conj_absorb word_bw_assocs(1)) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma mask_1[simp]: "(\x. mask x = 1)" apply (rule_tac x=1 in exI) apply (simp add:mask_def) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_def) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (metis (mono_tags) min_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (metis (mono_tags) max_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (metis word_bool_alg.xor_assoc word_bool_alg.xor_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right) lemma scast_nop_1: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop_2: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop[simp] = scast_nop_1 scast_nop_2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ mask n) && mask n = x && mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) || y && mask n) && ~~ mask m = x && ~~ mask m" apply (subst word_ao_dist) apply (subgoal_tac "(y && mask n) && ~~ mask m = 0") apply simp by (metis (no_types, hide_lams) is_aligned_mask is_aligned_weaken word_and_not word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma mask_twice2: "n \ m \ ((x::'a::len word) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" apply clarsimp apply (subgoal_tac "2 \ uints (LENGTH('a))") apply (subst (asm) Word.word_ubin.set_iff_norm) apply simp apply (subst word_uint.set_iff_norm) apply clarsimp apply (rule int_mod_eq') apply simp apply (rule pow_2_gt) apply simp done lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" apply (case_tac "LENGTH('a) = 1") apply simp apply (subgoal_tac "x = 0 \ x = 1") apply (erule disjE) apply (clarsimp simp:word_div_def)+ apply (metis One_nat_def less_irrefl_nat sint_1_cases) apply clarsimp apply (subst word_div_def) apply clarsimp apply (subst bintrunc_id) apply (subgoal_tac "2 \ LENGTH('a)") apply simp apply (metis (no_types) le_0_eq le_SucE lens_not_0(2) not_less_eq_eq numeral_2_eq_2) apply simp apply (subst bin_rest_def[symmetric]) apply (subst shiftr1_def[symmetric]) apply (clarsimp simp:shiftr1_unfold) done lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (case_tac "n = 0") apply clarsimp apply (subst div_by_0_word) apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) apply (metis (poly_guards_query) One_nat_def less_one nat_neq_iff unat_eq_1(2) unat_eq_zero) apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1 max_word_max plus_1_less) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" by (metis add.commute add_left_cancel max_word_max not_less word_and_1 word_bool_alg.conj_one_right word_bw_comms(1) word_overflow zero_neq_one) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" unfolding word_lsb_def bin_last_def by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" apply (case_tac "LENGTH('a) = 1") apply clarsimp apply (drule_tac x=x in degenerate_word[unfolded One_nat_def]) apply (erule disjE) apply clarsimp+ apply (subst (asm) shiftr1_is_div_2[unfolded One_nat_def])+ apply (subst (asm) word_arith_nat_div)+ apply clarsimp apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (case_tac "x + 1 = 0") apply (clarsimp simp:overflow_imp_lsb) apply (cut_tac x=x in word_overflow_unat) apply clarsimp apply (case_tac "even (unat x)") apply (subgoal_tac "unat x div 2 = Suc (unat x) div 2") apply metis apply (subst numeral_2_eq_2)+ apply simp apply (simp add:odd_iff_lsb) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (clarsimp simp:mask_def) apply (metis max_word_eq max_word_max mult_2_right) apply (metis add_diff_cancel_left' diff_less len_gt_0 mult_2_right) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) || (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x && mask n || x && ~~ mask n = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ mask n = p && ~~ mask n" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (cut_tac n=x and 'a='a in max_word_max, clarsimp simp:max_word_def, simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. ucast (max_word::'b::len word)" shows "ucast ((ucast x)::'b word) = x" proof - { assume a1: "x \ word_of_int (uint (word_of_int (2 ^ len_of (TYPE('b)) - 1)::'b word))" have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ len_of (TYPE('b)::'b itself) + - (1::int) = - 1 + 2 ^ len_of (TYPE('b))" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) } with assms show ?thesis unfolding ucast_def apply (subst word_ubin.eq_norm) apply (subst and_mask_bintr[symmetric]) apply (subst and_mask_eq_iff_le_mask) apply (clarsimp simp: max_word_def mask_def) done qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_def not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (len_of(TYPE('a)) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < len_of(TYPE('a)) \ (~~ (2 ^ n :: ('a::len word))) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x && y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code le_Bits intro: basic_trans_rules)[1] apply (case_tac bit') apply (simp add: le_Bits less_eq_int_code) apply (auto simp: le_Bits less_eq_int_code) done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len0 word) \ v" apply (induct n, simp) apply (unfold shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v") apply simp apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) \ len_of (TYPE('a))" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) \ 0" unfolding word_clz_def by (simp add: max_word_bl[simplified max_word_minus] takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" apply (subst le_less) apply (rule iffI) apply clarsimp apply (subst (asm) add.commute) apply (subst (asm) less_x_plus_1) apply (cut_tac ?'a='a and y=n and x="unat (maxBound::'a word) - 1" in of_nat_mono_maybe') apply clarsimp apply (simp add: less_imp_diff_less) apply (insert bound[unfolded maxBound_word, simplified]) using order_less_trans apply blast apply (metis max_word_minus word_not_simps(3) word_of_nat_less) apply clarsimp apply (erule disjE) apply (subgoal_tac "x < of_nat (1 + n)") prefer 2 apply (cut_tac ?'a='a and y="unat x" and x="1 + n" in of_nat_mono_maybe) apply clarsimp apply (simp add: less_trans_Suc) apply (simp add: less_Suc_eq unat_less_helper) apply clarsimp apply clarsimp apply clarsimp apply (cut_tac y="(of_nat n)::'a word" and x="(of_nat n)::'a word" in less_x_plus_1) apply (cut_tac ?'a='a and y=n and x="unat (maxBound::'a word) - 1" in of_nat_mono_maybe') apply clarsimp apply (simp add: less_imp_diff_less) apply (insert bound[unfolded maxBound_word, simplified]) using order_less_trans apply blast apply (metis max_word_minus word_not_simps(3) word_of_nat_less) apply (clarsimp simp: add.commute) done lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes uc :"LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (simp add: uc) apply simp done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" apply (erule order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) apply assumption apply simp done lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" apply (rule iffI) apply (simp add:unat_ucast_less_no_overflow) apply (simp add:unat_less_helper) done lemma unat_ucast_no_overflow_le: assumes no_overflow : "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::('a::len) word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have ucast_rewrite: "ucast (f::('a::len) word) < ((ucast (ucast b ::('a::len) word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast upward_cast) apply (simp add: ucast_nat_def[symmetric]) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "length (dropWhile Not (to_bl w)) \ LENGTH('s) \ LENGTH('s) \ LENGTH('l) \ (of_bl:: bool list \ 'l::len word) (to_bl ((of_bl:: bool list \ 's::len word) (to_bl w))) = w" apply(rule word_uint_eqI) apply(subst uint_of_bl_is_bl_to_bin) apply(simp; fail) apply(subst to_bl_bin) apply(subst uint_of_bl_is_bl_to_bin_drop) apply blast apply(simp) done (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "length (dropWhile Not (to_bl w)) \ LENGTH('s) \ LENGTH('s) \ LENGTH('l) \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply(subst Word.ucast_bl)+ apply(rule bl_cast_long_short_long_ingoreLeadingZero_generic) apply(simp_all) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ mask n = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) (* simp normal form would be "nat (bintrunc (LENGTH('a)) 2) = 2" *) lemma word_len_min_2: "Suc 0 < LENGTH('a) \ unat (2::'a::len word) = 2" by (metis less_trans_Suc n_less_equal_power_2 numeral_2_eq_2 of_nat_numeral unat_of_nat_len) lemma upper_bits_unset_is_l2p: "n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n') = ((p::'a::len word) < 2 ^ n)" apply (cases "Suc 0 < LENGTH('a)") prefer 2 apply (subgoal_tac "LENGTH('a) = 1", auto simp: word_eq_iff)[1] apply (rule iffI) apply (subst mask_eq_iff_w2p [symmetric]) apply (clarsimp simp: word_size) apply (rule word_eqI, rename_tac n') apply (case_tac "n' < n"; simp add: word_size) apply clarify apply (drule bang_is_le) apply (drule_tac y=p in order_le_less_trans, assumption) apply (drule word_power_increasing; simp add: word_len_min_2[simplified]) done lemma less_2p_is_upper_bits_unset: "((p::'a::len word) < 2 ^ n) = (n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n'))" apply (cases "n < LENGTH('a)") apply (simp add: upper_bits_unset_is_l2p) apply (simp add: power_overflow) done lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" apply (cases "n = 0", simp) apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (subst mult.commute, erule nat_less_power_trans) apply simp apply simp done lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) unfolding is_aligned_def apply clarsimp apply (subst (asm) unat_of_nat_len) apply (metis order_less_trans unat_lt2p unat_power_lower) apply (metis nat_dvd_not_less) done lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt3 apply simp apply (rule diff_less_mono[OF unat_mono, OF lt2]) apply (simp add: word_le_nat_alt[symmetric]) done qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: shows "(max_word :: 'a::len word) = mask (LENGTH('a))" unfolding mask_def by (metis max_word_eq shiftl_1) lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ mask n" unfolding complement_def mask_def by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes al: "is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = a" using sz al unfolding alignUp_def apply (simp add: complement_mask) apply (subst x_power_minus_1) apply (subst neg_mask_is_div) apply (simp only: word_arith_nat_div unat_word_ariths) apply (simp only: unat_power_lower) apply (subst power_mod_div) apply (erule is_alignedE) apply simp apply (subst unat_mult_power_lem) apply simp apply (subst unat_sub) apply (subst unat_arith_simps) apply simp apply (simp add: del: unat_1) apply simp done lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz apply - apply (rule div_less) apply (simp add: unat_minus_one) apply (rule order_less_trans) apply (rule diff_Suc_less) apply (erule contrapos_np) apply (simp add: unat_eq_zero) apply (subst unat_power_lower [symmetric, OF sz]) apply (subst word_less_nat_alt [symmetric]) apply (rule word_mod_less_divisor) apply (simp add: p2_gt_0) done have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz apply - apply (subst td_gal_lt [symmetric]) apply simp apply (simp add: power_add [symmetric]) done have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz apply - apply (rule nat_le_power_trans) apply simp apply (rule Suc_leI [OF lt0]) apply simp done moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz apply - apply (subst td_gal_lt [symmetric]) apply simp apply (simp add: power_add [symmetric]) done have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz apply - apply (rule nat_le_power_trans) apply simp apply (rule Suc_leI [OF lt0]) apply simp done from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz apply (simp add: unat_div word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_less_le_trans [OF kv]) apply simp+ done have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ mask sz" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ mask n = y; n \ m \ \ x && ~~ mask m = y && ~~ mask m" apply (rule word_eqI, rename_tac n') apply (drule_tac x=n' in word_eqD) apply (auto simp: word_ops_nth_size word_size) done lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" apply (rule ccontr,simp add:not_less) apply (drule le_shiftr[where n = n]) apply (simp add: aligned_shift') apply (case_tac "b >> n = a >> n") apply (drule arg_cong[where f = "\x. x< alignUp (w + a) us = w + alignUp a us" apply (clarsimp simp:alignUp_def2 add.assoc) apply (simp add: mask_out_add_aligned field_simps) done lemma mask_lower_twice: "n \ m \ (x && ~~ mask n) && ~~ mask m = x && ~~ mask m" apply (rule word_eqI) apply (simp add: word_size word_ops_nth_size) apply safe apply simp done lemma mask_lower_twice2: "(a && ~~ mask n) && ~~ mask m = a && ~~ mask (max n m)" by (rule word_eqI, simp add: neg_mask_bang conj_comms) lemma ucast_and_neg_mask: "ucast (x && ~~ mask n) = ucast x && ~~ mask n" apply (rule word_eqI) apply (simp add: word_size neg_mask_bang nth_ucast) apply (auto simp add: test_bit_bl word_size) done lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" apply (rule word_eqI) apply (simp add: nth_ucast) apply (auto simp add: test_bit_bl word_size) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" apply (rule word_eqI) apply (simp add: nth_ucast word_size) apply safe apply (simp add: test_bit_bl word_size) done lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" apply (case_tac "LENGTH('a) \ sz") apply (simp add:alignUp_def2 mask_def power_overflow) apply (case_tac "is_aligned q sz") apply (clarsimp simp:alignUp_def2 p_assoc_help) apply (subst mask_out_add_aligned[symmetric],simp)+ apply (simp add:mask_lower_twice word_and_le2) apply (simp add:and_not_mask) apply (subst le_mask_iff[THEN iffD1]) apply (simp add:mask_def) apply simp apply (clarsimp simp:alignUp_def3) apply (subgoal_tac "2 ^ sz - (q - (q - 1 && ~~ mask sz)) \ 2 ^ sz - 1") apply (simp add:field_simps mask_def) apply (rule word_sub_mono) apply simp apply (rule ccontr) apply (clarsimp simp:not_le) apply (drule eq_refl) apply (drule order_trans[OF _ word_and_le2]) apply (subgoal_tac "q \ 0") apply (drule minus_one_helper[rotated]) apply simp apply simp apply (fastforce) apply (simp add: word_sub_le_iff) apply (subgoal_tac "q - 1 && ~~ mask sz = (q - 1) - (q - 1 && mask sz)") apply simp apply (rule order_trans) apply (rule word_add_le_mono2) apply (rule word_and_le1) apply (subst unat_plus_simple[THEN iffD1,symmetric]) apply (simp add:not_le mask_def) apply (rule word_sub_1_le) apply simp apply (rule unat_lt2p) apply (simp add:mask_def) apply (simp add:mask_out_sub_mask) apply (rule word_sub_1_le) apply simp done lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ mask sz) = (p - ((alignUp q sz) && ~~ mask sz))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]) apply simp+ apply (rule sum_to_zero) apply (subst add.commute) apply (subst mask_out_add_aligned) apply (simp add:is_aligned_neg_mask) apply simp apply (subst and_not_mask[where w = "(alignUp q sz && ~~ mask sz) - q "]) apply (subst le_mask_iff[THEN iffD1]) apply (simp add:is_aligned_neg_mask_eq) apply (rule alignUp_distance) apply simp done lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a)" shows "x = ucast y \ ucast x = y" using assms by (simp add: is_down ucast_id ucast_ucast_a) lemma le_ucast_ucast_le: fixes x :: "'a::len word" and y :: "'b::len word" assumes "LENGTH('b) \ LENGTH('a)" shows "x \ ucast y \ ucast x \ y" using assms apply (simp add: word_le_nat_alt unat_ucast_up_simp[where x=y]) apply (simp add: unat_ucast) by (rule le_trans; fastforce) lemma less_ucast_ucast_less: fixes x :: "'a::len word" and y :: "'b::len word" assumes "LENGTH('b) \ LENGTH('a)" shows "x < ucast y \ ucast x < y" using assms apply (simp add: word_less_nat_alt unat_ucast_up_simp[where x=y]) apply (simp add: unat_ucast) by (rule le_less_trans; fastforce) lemma ucast_le_ucast: fixes x :: "'a::len word" and y :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "(ucast x \ (ucast y::'b::len word)) = (x \ y)" using assms apply (simp add: word_le_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done (* High bits w.r.t. mask operations. *) lemma and_neg_mask_eq_iff_not_mask_le: "w && ~~ mask n = ~~ mask n \ ~~ mask n \ w" by (metis (full_types) dual_order.antisym neg_mask_mono_le word_and_le1 word_and_le2 word_bool_alg.conj_absorb) lemma le_mask_high_bits: shows "w \ mask n \ (\ i \ {n ..< size w}. \ w !! i)" by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) lemma neg_mask_le_high_bits: shows "~~ mask n \ w \ (\ i \ {n ..< size w}. w !! i)" by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_bang) lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" unfolding ucast_def Word.bitOR_word.abs_eq uint_or by blast lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" apply (simp add: word_less_nat_alt shiftr_div_2n') apply (blast intro: div_le_dividend le_less_trans) done lemma word_and_notzeroD: "w && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def apply (simp add: word_size) apply (rule_tac y="length (to_bl w)" in order_trans) apply (rule List.length_takeWhile_le) apply simp done lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" using takeWhile_take_has_property[where n="length (to_bl w)" and xs="to_bl w" and P=Not] by simp hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" apply simp apply (subst (asm) to_bl_0[symmetric]) apply (drule Word.word_bl.Rep_eqD, assumption) done with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0[simp]: "max_word \ 0" by (simp add: max_word_minus) lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed lemma unat_max_word_pos[simp]: "0 < unat max_word" by (auto simp: unat_gt_0) (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2^n = y * (2^n::'a::len word)" shows "x = y" proof (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_bang word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if' = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ mask (Suc n) else w && mask (Suc n))" by (rule word_eqI[rule_format]) (auto simp: sign_extend_bitwise_if' word_size word_ops_nth_size dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (rule word_eqI[rule_format], rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size sign_extend_bitwise_if') apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (cases "m < n") (auto intro!: word_eqI simp: word_size sign_extend_bitwise_cases') lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by (rule word_eqI, fastforce dest: word_eqD simp: sign_extend_bitwise_if' word_size) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ mask m)" by (fastforce simp: sign_extended_def word_size neg_mask_bang) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "w BIT c \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num Bit_def) lemma Bit_in_uintsI: "w BIT c \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: "bin_cat a n b \ uints m" if "a \ uints l" "m \ l + n" using that proof (induction n arbitrary: b m) case 0 then have "uints l \ uints m" by (intro uints_monoI) auto then show ?case using 0 by auto next case (Suc n) then show ?case by (auto intro!: Bit_in_uintsI) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) proof (induction m arbitrary: b d) case (Suc m) show ?case using Suc.prems by (auto intro: Suc.IH) qed simp lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (induction n arbitrary: b d) auto lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (induction n arbitrary: b d) auto lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len0 word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len0 word" and b::"'b::len0 word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len0 word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len0 word" and b::"'b::len0 word" by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT mask (LENGTH('a) - len) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT mask (LENGTH('a) - len) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT mask (LENGTH('a) - len) AND a" proof - have f2: "\x\<^sub>0. base AND NOT mask x\<^sub>0 \ a AND NOT mask x\<^sub>0" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT mask x\<^sub>0 \ (base OR mask (LENGTH('a) - len)) AND NOT mask x\<^sub>0" using a neg_mask_mono_le by blast have f4: "base = base AND NOT mask (LENGTH('a) - len)" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT mask (LENGTH('a) - len) = base OR x\<^sub>6 AND NOT mask (LENGTH('a) - len)" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT mask x\<^sub>2 \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT mask x\<^sub>2 \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT mask (LENGTH('a) - len)" using f5 by auto hence "base = a AND NOT mask (LENGTH('a) - len)" using f2 f4 f6 by (metis eq_iff) thus "base = NOT mask (LENGTH('a) - len) AND a" by (metis word_bw_comms(1)) qed qed lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: ('a :: len) word)) \ (- 1 :: ('a::len) word)" apply (cut_tac w=w in word_ctz_le) apply (subst word_unat.Rep_inject[symmetric]) apply (subst unat_of_nat_eq) apply (erule order_le_less_trans, fastforce) apply (subst unat_minus_one_word) apply (rule less_imp_neq) apply (erule order_le_less_trans) apply (subst less_eq_Suc_le) apply (subst le_diff_conv2, fastforce) apply (clarsimp simp: le_diff_conv2 less_eq_Suc_le[symmetric] suc_le_pow_2) done lemmas word_ctz_not_minus_1_32 = word_ctz_not_minus_1[where 'a=32, simplified] end diff --git a/thys/pGCL/Algebra.thy b/thys/pGCL/Algebra.thy --- a/thys/pGCL/Algebra.thy +++ b/thys/pGCL/Algebra.thy @@ -1,1247 +1,1248 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section "The Algebra of pGCL" theory Algebra imports WellDefined begin text_raw \\label{s:Algebra}\ text \Programs in pGCL have a rich algebraic structure, largely mirroring that for GCL. We show that programs form a lattice under refinement, with @{term "a \ b"} and @{term "a \ b"} as the meet and join operators, respectively. We also take advantage of the algebraic structure to establish a framwork for the modular decomposition of proofs.\ subsection \Program Refinement\ text_raw \\label{s:progref}\ text \Refinement in pGCL relates to refinement in GCL exactly as probabilistic entailment relates to implication. It turns out to have a very similar algebra, the rules of which we establish shortly.\ definition refines :: "'s prog \ 's prog \ bool" (infix "\" 70) where "prog \ prog' \ \P. sound P \ wp prog P \ wp prog' P" lemma refinesI[intro]: "\ \P. sound P \ wp prog P \ wp prog' P \ \ prog \ prog'" unfolding refines_def by(simp) lemma refinesD[dest]: "\ prog \ prog'; sound P \ \ wp prog P \ wp prog' P" unfolding refines_def by(simp) text \The equivalence relation below will turn out to be that induced by refinement. It is also the application of @{term equiv_trans} to the weakest precondition.\ definition pequiv :: "'s prog \ 's prog \ bool" (infix "\" 70) where "prog \ prog' \ \P. sound P \ wp prog P = wp prog' P" lemma pequivI[intro]: "\ \P. sound P \ wp prog P = wp prog' P \ \ prog \ prog'" unfolding pequiv_def by(simp) lemma pequivD[dest,simp]: "\ prog \ prog'; sound P \ \ wp prog P = wp prog' P" unfolding pequiv_def by(simp) lemma pequiv_equiv_trans: "a \ b \ equiv_trans (wp a) (wp b)" by(auto) subsection \Simple Identities\ text \The following identities involve only the primitive operations as defined in \autoref{s:syntax}, and refinement as defined above.\ subsubsection \Laws following from the basic arithmetic of the operators seperately\ lemma DC_comm[ac_simps]: "a \ b = b \ a" unfolding DC_def by(simp add:ac_simps) lemma DC_assoc[ac_simps]: "a \ (b \ c) = (a \ b) \ c" unfolding DC_def by(simp add:ac_simps) lemma DC_idem: "a \ a = a" unfolding DC_def by(simp) lemma AC_comm[ac_simps]: "a \ b = b \ a" unfolding AC_def by(simp add:ac_simps) lemma AC_assoc[ac_simps]: "a \ (b \ c) = (a \ b) \ c" unfolding AC_def by(simp add:ac_simps) lemma AC_idem: "a \ a = a" unfolding AC_def by(simp) lemma PC_quasi_comm: "a \<^bsub>p\<^esub>\ b = b \<^bsub>(\s. 1 - p s)\<^esub>\ a" unfolding PC_def by(simp add:algebra_simps) lemma PC_idem: "a \<^bsub>p\<^esub>\ a = a" unfolding PC_def by(simp add:algebra_simps) lemma Seq_assoc[ac_simps]: "A ;; (B ;; C) = A ;; B ;; C" by(simp add:Seq_def o_def) lemma Abort_refines[intro]: "well_def a \ Abort \ a" by(rule refinesI, unfold wp_eval, auto dest!:well_def_wp_healthy) subsubsection \Laws relating demonic choice and refinement\ lemma left_refines_DC: "(a \ b) \ a" by(auto intro!:refinesI simp:wp_eval) lemma right_refines_DC: "(a \ b) \ b" by(auto intro!:refinesI simp:wp_eval) lemma DC_refines: fixes a::"'s prog" and b and c assumes rab: "a \ b" and rac: "a \ c" shows "a \ (b \ c)" proof fix P::"'s \ real" assume sP: "sound P" with assms have "wp a P \ wp b P" and "wp a P \ wp c P" by(auto dest:refinesD) thus "wp a P \ wp (b \ c) P" by(auto simp:wp_eval intro:min.boundedI) qed lemma DC_mono: fixes a::"'s prog" assumes rab: "a \ b" and rcd: "c \ d" shows "(a \ c) \ (b \ d)" proof(rule refinesI, unfold wp_eval, rule le_funI) fix P::"'s \ real" and s::'s assume sP: "sound P" with assms have "wp a P s \ wp b P s" and "wp c P s \ wp d P s" by(auto) thus "min (wp a P s) (wp c P s) \ min (wp b P s) (wp d P s)" by(auto) qed subsubsection \Laws relating angelic choice and refinement\ lemma left_refines_AC: "a \ (a \ b)" by(auto intro!:refinesI simp:wp_eval) lemma right_refines_AC: "b \ (a \ b)" by(auto intro!:refinesI simp:wp_eval) lemma AC_refines: fixes a::"'s prog" and b and c assumes rac: "a \ c" and rbc: "b \ c" shows "(a \ b) \ c" proof fix P::"'s \ real" assume sP: "sound P" with assms have "\s. wp a P s \ wp c P s" and "\s. wp b P s \ wp c P s" by(auto dest:refinesD) thus "wp (a \ b) P \ wp c P" unfolding wp_eval by(auto) qed lemma AC_mono: fixes a::"'s prog" assumes rab: "a \ b" and rcd: "c \ d" shows "(a \ c) \ (b \ d)" proof(rule refinesI, unfold wp_eval, rule le_funI) fix P::"'s \ real" and s::'s assume sP: "sound P" with assms have "wp a P s \ wp b P s" and "wp c P s \ wp d P s" by(auto) thus "max (wp a P s) (wp c P s) \ max (wp b P s) (wp d P s)" by(auto) qed subsubsection \Laws depending on the arithmetic of @{term "a \<^bsub>p\<^esub>\ b"} and @{term "a \ b"} together\ lemma DC_refines_PC: assumes unit: "unitary p" shows "(a \ b) \ (a \<^bsub>p\<^esub>\ b)" proof(rule refinesI, unfold wp_eval, rule le_funI) fix s and P::"'a \ real" assume sound: "sound P" from unit have nn_p: "0 \ p s" by(blast) from unit have "p s \ 1" by(blast) hence nn_np: "0 \ 1 - p s" by(simp) show "min (wp a P s) (wp b P s) \ p s * wp a P s + (1 - p s) * wp b P s" proof(cases "wp a P s \ wp b P s", simp_all add:min.absorb1 min.absorb2) case True note le = this have "wp a P s = (p s + (1 - p s)) * wp a P s" by(simp) also have "... = p s * wp a P s + (1 - p s) * wp a P s" by(simp only: distrib_right) also { from le and nn_np have "(1 - p s) * wp a P s \ (1 - p s) * wp b P s" by(rule mult_left_mono) hence "p s * wp a P s + (1 - p s) * wp a P s \ p s * wp a P s + (1 - p s) * wp b P s" by(rule add_left_mono) } finally show "wp a P s \ p s * wp a P s + (1 - p s) * wp b P s" . next case False then have le: "wp b P s \ wp a P s" by(simp) have "wp b P s = (p s + (1 - p s)) * wp b P s" by(simp) also have "... = p s * wp b P s + (1 - p s) * wp b P s" by(simp only:distrib_right) also { from le and nn_p have "p s * wp b P s \ p s * wp a P s" by(rule mult_left_mono) hence "p s * wp b P s + (1 - p s) * wp b P s \ p s * wp a P s + (1 - p s) * wp b P s" by(rule add_right_mono) } finally show "wp b P s \ p s * wp a P s + (1 - p s) * wp b P s" . qed qed subsubsection \Laws depending on the arithmetic of @{term "a \<^bsub>p\<^esub>\ b"} and @{term "a \ b"} together\ lemma PC_refines_AC: assumes unit: "unitary p" shows "(a \<^bsub>p\<^esub>\ b) \ (a \ b)" proof(rule refinesI, unfold wp_eval, rule le_funI) fix s and P::"'a \ real" assume sound: "sound P" from unit have "p s \ 1" by(blast) hence nn_np: "0 \ 1 - p s" by(simp) show "p s * wp a P s + (1 - p s) * wp b P s \ max (wp a P s) (wp b P s)" proof(cases "wp a P s \ wp b P s") case True note leab = this with unit nn_np have "p s * wp a P s + (1 - p s) * wp b P s \ p s * wp b P s + (1 - p s) * wp b P s" by(auto intro:add_mono mult_left_mono) also have "... = wp b P s" by(auto simp:field_simps) also from leab have "... = max (wp a P s) (wp b P s)" by(auto) finally show ?thesis . next case False note leba = this with unit nn_np have "p s * wp a P s + (1 - p s) * wp b P s \ p s * wp a P s + (1 - p s) * wp a P s" by(auto intro:add_mono mult_left_mono) also have "... = wp a P s" by(auto simp:field_simps) also from leba have "... = max (wp a P s) (wp b P s)" by(auto) finally show ?thesis . qed qed subsubsection \Laws depending on the arithmetic of @{term "a \ b"} and @{term "a \ b"} together \ lemma DC_refines_AC: "(a \ b) \ (a \ b)" by(auto intro!:refinesI simp:wp_eval) subsubsection \Laws Involving Refinement and Equivalence\ lemma pr_trans[trans]: fixes A::"'a prog" assumes prAB: "A \ B" and prBC: "B \ C" shows "A \ C" proof fix P::"'a \ real" assume sP: "sound P" with prAB have "wp A P \ wp B P" by(blast) also from sP and prBC have "... \ wp C P" by(blast) finally show "wp A P \ ..." . qed lemma pequiv_refl[intro!,simp]: "a \ a" by(auto) lemma pequiv_comm[ac_simps]: "a \ b \ b \ a" unfolding pequiv_def by(rule iffI, safe, simp_all) lemma pequiv_pr[dest]: "a \ b \ a \ b" by(auto) lemma pequiv_trans[intro,trans]: "\ a \ b; b \ c \ \ a \ c" unfolding pequiv_def by(auto intro!:order_trans) lemma pequiv_pr_trans[intro,trans]: "\ a \ b; b \ c \ \ a \ c" unfolding pequiv_def refines_def by(simp) lemma pr_pequiv_trans[intro,trans]: "\ a \ b; b \ c \ \ a \ c" unfolding pequiv_def refines_def by(simp) text \Refinement induces equivalence by antisymmetry:\ lemma pequiv_antisym: "\ a \ b; b \ a \ \ a \ b" by(auto intro:antisym) lemma pequiv_DC: "\ a \ c; b \ d \ \ (a \ b) \ (c \ d)" by(auto intro!:DC_mono pequiv_antisym simp:ac_simps) lemma pequiv_AC: "\ a \ c; b \ d \ \ (a \ b) \ (c \ d)" by(auto intro!:AC_mono pequiv_antisym simp:ac_simps) subsection \Deterministic Programs are Maximal\ text \Any sub-additive refinement of a deterministic program is in fact an equivalence. Deterministic programs are thus maximal (under the refinement order) among sub-additive programs. \ lemma refines_determ: fixes a::"'s prog" assumes da: "determ (wp a)" and wa: "well_def a" and wb: "well_def b" and dr: "a \ b" shows "a \ b" txt \Proof by contradiction.\ proof(rule pequivI, rule contrapos_pp) from wb have "feasible (wp b)" by(auto) with wb have sab: "sub_add (wp b)" by(auto dest: sublinear_subadd[OF well_def_wp_sublinear]) fix P::"'s \ real" assume sP: "sound P" txt \Assume that @{term a} and @{term b} are not equivalent:\ assume ne: "wp a P \ wp b P" txt \Find a point at which they differ. As @{term "a \ b"}, @{term "wp b P s"} must by strictly greater than @{term "wp a P s"} here:\ hence "\s. wp a P s < wp b P s" proof(rule contrapos_np) assume "\(\s. wp a P s < wp b P s)" hence "\s. wp b P s \ wp a P s" by(auto simp:not_less) hence "wp b P \ wp a P" by(auto) moreover from sP dr have "wp a P \ wp b P" by(auto) ultimately show "wp a P = wp b P" by(auto) qed then obtain s where less: "wp a P s < wp b P s" by(blast) txt \Take a carefully constructed expectation:\ let ?Pc = "\s. bound_of P - P s" have sPc: "sound ?Pc" proof(rule soundI) from sP have "\s. 0 \ P s" by(auto) hence "\s. ?Pc s \ bound_of P" by(auto) thus "bounded ?Pc" by(blast) from sP have "\s. P s \ bound_of P" by(auto) - hence "\s. 0 \ ?Pc s" by(auto simp:sign_simps) + hence "\s. 0 \ ?Pc s" + by auto thus "nneg ?Pc" by(auto) qed txt \We then show that @{term "wp b"} violates feasibility, and thus healthiness.\ from sP have "0 \ bound_of P" by(auto) with da have "bound_of P = wp a (\s. bound_of P) s" by(simp add:maximalD determ_maximalD) also have "... = wp a (\s. ?Pc s + P s) s" by(simp) also from da sP sPc have "... = wp a ?Pc s + wp a P s" by(subst additiveD[OF determ_additiveD], simp_all add:sP sPc) also from sPc dr have "... \ wp b ?Pc s + wp a P s" by(auto) also from less have "... < wp b ?Pc s + wp b P s" by(auto) also from sab sP sPc have "... \ wp b (\s. ?Pc s + P s) s" by(blast) finally have "\wp b (\s. bound_of P) s \ bound_of P" by(simp) thus "\bounded_by (bound_of P) (wp b (\s. bound_of P))" by(auto) next txt \However,\ fix P::"'s \ real" assume sP: "sound P" hence "nneg (\s. bound_of P)" by(auto) moreover have "bounded_by (bound_of P) (\s. bound_of P)" by(auto) ultimately show "bounded_by (bound_of P) (wp b (\s. bound_of P))" using wb by(auto dest!:well_def_wp_healthy) qed subsection \The Algebraic Structure of Refinement\ text \Well-defined programs form a half-bounded semilattice under refinement, where @{term Abort} is bottom, and @{term "a \ b"} is @{term inf}. There is no unique top element, but all fully-deterministic programs are maximal. The type that we construct here is not especially useful, but serves as a convenient way to express this result.\ quotient_type 's program = "'s prog" / partial : "\a b. a \ b \ well_def a \ well_def b" proof(rule part_equivpI) have "Skip \ Skip" and "well_def Skip" by(auto intro:wd_intros) thus "\x. x \ x \ well_def x \ well_def x" by(blast) show "symp (\a b. a \ b \ well_def a \ well_def b)" proof(rule sympI, safe) fix a::"'a prog" and b assume "a \ b" hence "equiv_trans (wp a) (wp b)" by(simp add:pequiv_equiv_trans) thus "b \ a" by(simp add:ac_simps pequiv_equiv_trans) qed show "transp (\a b. a \ b \ well_def a \ well_def b)" by(rule transpI, safe, rule pequiv_trans) qed instantiation program :: (type) semilattice_inf begin lift_definition less_eq_program :: "'a program \ 'a program \ bool" is refines proof(safe) fix a::"'a prog" and b c d assume "a \ b" hence "b \ a" by(simp add:ac_simps) also assume "a \ c" also assume "c \ d" finally show "b \ d" . next fix a::"'a prog" and b c d assume "a \ b" also assume "b \ d" also assume "c \ d" hence "d \ c" by(simp add:ac_simps) finally show "a \ c" . qed (* XXX - what's up here? *) lift_definition less_program :: "'a program \ 'a program \ bool" is "\a b. a \ b \ \ b \ a" proof(safe) fix a::"'a prog" and b c d assume "a \ b" hence "b \ a" by(simp add:ac_simps) also assume "a \ c" also assume "c \ d" finally show "b \ d" . next fix a::"'a prog" and b c d assume "a \ b" also assume "b \ d" also assume "c \ d" hence "d \ c" by(simp add:ac_simps) finally show "a \ c" . next fix a b and c::"'a prog" and d assume "c \ d" also assume "d \ b" also assume "a \ b" hence "b \ a" by(simp add:ac_simps) finally have "c \ a" . moreover assume "\ c \ a" ultimately show False by(auto) next fix a b and c::"'a prog" and d assume "c \ d" hence "d \ c" by(simp add:ac_simps) also assume "c \ a" also assume "a \ b" finally have "d \ b" . moreover assume "\ d \ b" ultimately show False by(auto) qed lift_definition inf_program :: "'a program \ 'a program \ 'a program" is DC proof(safe) fix a b c d::"'s prog" assume "a \ b" and "c \ d" thus "(a \ c) \ (b \ d)" by(rule pequiv_DC) next fix a c::"'s prog" assume "well_def a" "well_def c" thus "well_def (a \ c)" by(rule wd_intros) next fix a c::"'s prog" assume "well_def a" "well_def c" thus "well_def (a \ c)" by(rule wd_intros) qed instance proof fix x y::"'a program" show "(x < y) = (x \ y \ \ y \ x)" by(transfer, simp) show "x \ x" by(transfer, auto) show "inf x y \ x" by(transfer, rule left_refines_DC) show "inf x y \ y" by(transfer, rule right_refines_DC) assume "x \ y" and "y \ x" thus "x = y" by(transfer, iprover intro:pequiv_antisym) next fix x y z::"'a program" assume "x \ y" and "y \ z" thus "x \ z" by(transfer, iprover intro:pr_trans) next fix x y z::"'a program" assume "x \ y" and "x \ z" thus "x \ inf y z" by(transfer, iprover intro:DC_refines) qed end instantiation program :: (type) bot begin lift_definition bot_program :: "'a program" is Abort by(auto intro:wd_intros) instance .. end lemma eq_det: "\a b::'s prog. \ a \ b; determ (wp a) \ \ determ (wp b)" proof(intro determI additiveI maximalI) fix a b::"'s prog" and P::"'s \ real" and Q::"'s \ real" and s::"'s" assume da: "determ (wp a)" assume sP: "sound P" and sQ: "sound Q" and eq: "a \ b" hence "wp b (\s. P s + Q s) s = wp a (\s. P s + Q s) s" by(simp add:sound_intros) also from da sP sQ have "... = wp a P s + wp a Q s" by(simp add:additiveD determ_additiveD) also from eq sP sQ have "... = wp b P s + wp b Q s" by(simp add:pequivD) finally show "wp b (\s. P s + Q s) s = wp b P s + wp b Q s" . next fix a b::"'s prog" and c::real assume da: "determ (wp a)" assume "a \ b" hence "b \ a" by(simp add:ac_simps) moreover assume nn: "0 \ c" ultimately have "wp b (\_. c) = wp a (\_. c)" by(simp add:pequivD const_sound) also from da nn have "... = (\_. c)" by(simp add:determ_maximalD maximalD) finally show "wp b (\_. c) = (\_. c)" . qed lift_definition pdeterm :: "'s program \ bool" is "\a. determ (wp a)" proof(safe) fix a b::"'s prog" assume "a \ b" and "determ (wp a)" thus "determ (wp b)" by(rule eq_det) next fix a b::"'s prog" assume "a \ b" hence "b \ a" by(simp add:ac_simps) moreover assume "determ (wp b)" ultimately show "determ (wp a)" by(rule eq_det) qed lemma determ_maximal: "\ pdeterm a; a \ x \ \ a = x" by(transfer, auto intro:refines_determ) subsection \Data Refinement\ text \A projective data refinement construction for pGCL. By projective, we mean that the abstract state is always a function (@{term \}) of the concrete state. Refinement may be predicated (@{term G}) on the state.\ definition drefines :: "('b \ 'a) \ ('b \ bool) \ 'a prog \ 'b prog \ bool" where "drefines \ G A B \ \P Q. (unitary P \ unitary Q \ (P \ wp A Q)) \ (\G\ && (P o \) \ wp B (Q o \))" lemma drefinesD[dest]: "\ drefines \ G A B; unitary P; unitary Q; P \ wp A Q \ \ \G\ && (P o \) \ wp B (Q o \)" unfolding drefines_def by(blast) text \We can alternatively use G as an assumption:\ lemma drefinesD2: assumes dr: "drefines \ G A B" and uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" and G: "G s" shows "(P o \) s \ wp B (Q o \) s" proof - from uP have "0 \ (P o \) s" unfolding o_def by(blast) with G have "(P o \) s = (\G\ && (P o \)) s" by(simp add:exp_conj_def) also from assms have "... \ wp B (Q o \) s" by(blast) finally show "(P o \) s \ ..." . qed text \This additional form is sometimes useful:\ lemma drefinesD3: assumes dr: "drefines \ G a b" and G: "G s" and uQ: "unitary Q" and wa: "well_def a" shows "wp a Q (\ s) \ wp b (Q o \) s" proof - let "?L s'" = "wp a Q s'" from uQ wa have sL: "sound ?L" by(blast) from uQ wa have bL: "bounded_by 1 ?L" by(blast) have "?L \ ?L" by(simp) with sL and bL and assms show ?thesis by(blast intro:drefinesD2[OF dr, where P="?L", simplified]) qed lemma drefinesI[intro]: "\ \P Q. \ unitary P; unitary Q; P \ wp A Q \ \ \G\ && (P o \) \ wp B (Q o \) \ \ drefines \ G A B" unfolding drefines_def by(blast) text \Use G as an assumption, when showing refinement:\ lemma drefinesI2: fixes A::"'a prog" and B::"'b prog" and \::"'b \ 'a" and G::"'b \ bool" assumes wB: "well_def B" and withAs: "\P Q s. \ unitary P; unitary Q; G s; P \ wp A Q \ \ (P o \) s \ wp B (Q o \) s" shows "drefines \ G A B" proof fix P and Q assume uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" hence "\s. G s \ (P o \) s \ wp B (Q o \) s" using withAs by(blast) moreover from uQ have "unitary (Q o \)" unfolding o_def by(blast) moreover from uP have "unitary (P o \)" unfolding o_def by(blast) ultimately show "\G\ && (P o \) \ wp B (Q o \)" using wB by(blast intro:entails_pconj_assumption) qed lemma dr_strengthen_guard: fixes a::"'s prog" and b::"'t prog" assumes fg: "\s. F s \ G s" and drab: "drefines \ G a b" shows "drefines \ F a b" proof(intro drefinesI) fix P Q::"'s expect" assume uP: "unitary P" and uQ: "unitary Q" and wp: "P \ wp a Q" from fg have "\s. \F\ s \ \G\ s" by(simp add:embed_bool_def) hence "(\F\ && (P o \)) \ (\G\ && (P o \))" by(auto intro:pconj_mono le_funI simp:exp_conj_def) also from drab uP uQ wp have "... \ wp b (Q o \)" by(auto) finally show "\F\ && (P o \) \ wp b (Q o \)" . qed text \Probabilistic correspondence, @{term pcorres}, is equality on distribution transformers, modulo a guard. It is the analogue, for data refinement, of program equivalence for program refinement.\ definition pcorres :: "('b \ 'a) \ ('b \ bool) \ 'a prog \ 'b prog \ bool" where "pcorres \ G A B \ (\Q. unitary Q \ \G\ && (wp A Q o \) = \G\ && wp B (Q o \))" lemma pcorresI: "\ \Q. unitary Q \ \G\ && (wp A Q o \) = \G\ && wp B (Q o \) \ \ pcorres \ G A B" by(simp add:pcorres_def) text \Often easier to use, as it allows one to assume the precondition.\ lemma pcorresI2[intro]: fixes A::"'a prog" and B::"'b prog" assumes withG: "\Q s. \ unitary Q; G s \ \ wp A Q (\ s)= wp B (Q o \) s" and wA: "well_def A" and wB: "well_def B" shows "pcorres \ G A B" proof(rule pcorresI, rule ext) fix Q::"'a \ real" and s::'b assume uQ: "unitary Q" hence uQ\: "unitary (Q o \)" by(auto) show "(\G\ && (wp A Q \ \)) s = (\G\ && wp B (Q \ \)) s" proof(cases "G s") case True note this moreover from well_def_wp_healthy[OF wA] uQ have "0 \ wp A Q (\ s)" by(blast) moreover from well_def_wp_healthy[OF wB] uQ\ have "0 \ wp B (Q o \) s" by(blast) ultimately show ?thesis using uQ by(simp add:exp_conj_def withG) next case False note this moreover from well_def_wp_healthy[OF wA] uQ have "wp A Q (\ s) \ 1" by(blast) moreover from well_def_wp_healthy[OF wB] uQ\ have "wp B (Q o \) s \ 1" by(blast dest!:healthy_bounded_byD intro:sound_nneg) ultimately show ?thesis by(simp add:exp_conj_def) qed qed lemma pcorresD: "\ pcorres \ G A B; unitary Q \ \ \G\ && (wp A Q o \) = \G\ && wp B (Q o \)" unfolding pcorres_def by(simp) text \Again, easier to use if the precondition is known to hold.\ lemma pcorresD2: assumes pc: "pcorres \ G A B" and uQ: "unitary Q" and wA: "well_def A" and wB: "well_def B" and G: "G s" shows "wp A Q (\ s) = wp B (Q o \) s" proof - from uQ well_def_wp_healthy[OF wA] have "0 \ wp A Q (\ s)" by(auto) with G have "wp A Q (\ s) = \G\ s .& wp A Q (\ s)" by(simp) also { from pc uQ have "\G\ && (wp A Q o \) = \G\ && wp B (Q o \)" by(rule pcorresD) hence "\G\ s .& wp A Q (\ s) = \G\ s .& wp B (Q o \) s" unfolding exp_conj_def o_def by(rule fun_cong) } also { from uQ have "sound Q" by(auto) hence "sound (Q o \)" by(auto intro:sound_intros) with well_def_wp_healthy[OF wB] have "0 \ wp B (Q o \) s" by(auto) with G have "\G\ s .& wp B (Q o \) s = wp B (Q o \) s" by(simp) } finally show ?thesis . qed subsection \The Algebra of Data Refinement\ text \Program refinement implies a trivial data refinement:\ lemma refines_drefines: fixes a::"'s prog" assumes rab: "a \ b" and wb: "well_def b" shows "drefines (\s. s) G a b" proof(intro drefinesI2 wb, simp add:o_def) fix P::"'s \ real" and Q::"'s \ real" and s::'s assume sQ: "unitary Q" assume "P \ wp a Q" hence "P s \ wp a Q s" by(auto) also from rab sQ have "... \ wp b Q s" by(auto) finally show "P s \ wp b Q s" . qed text \Data refinement is transitive:\ lemma dr_trans[trans]: fixes A::"'a prog" and B::"'b prog" and C::"'c prog" assumes drAB: "drefines \ G A B" and drBC: "drefines \' G' B C" and Gimp: "\s. G' s \ G (\' s)" shows "drefines (\ o \') G' A C" proof(rule drefinesI) fix P::"'a \ real" and Q::"'a \ real" and s::'a assume uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" have "\G'\ && \G o \'\ = \G'\" proof(rule ext, unfold exp_conj_def) fix x show "\G'\ x .& \G o \'\ x = \G'\ x" (is ?X) proof(cases "G' x") case False then show ?X by(simp) next case True moreover with Gimp have "(G o \') x" by(simp add:o_def) ultimately show ?X by(simp) qed qed with uP have "\G'\ && (P o (\ o \')) = \G'\ && ((\G\ && (P o \)) o \')" by(simp add:exp_conj_assoc o_assoc) also { from uP uQ wpA and drAB have "\G\ && (P o \) \ wp B (Q o \)" by(blast intro:drefinesD) with drBC and uP uQ have "\G'\ && ((\G\ && (P o \)) o \') \ wp C ((Q o \) o \')" by(blast intro:unitary_intros drefinesD) } finally show "\G'\ && (P o (\ o \')) \ wp C (Q o (\ o \'))" by(simp add:o_assoc) qed text \Data refinement composes with program refinement:\ lemma pr_dr_trans[trans]: assumes prAB: "A \ B" and drBC: "drefines \ G B C" shows "drefines \ G A C" proof(rule drefinesI) fix P and Q assume uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" note wpA also from uQ and prAB have "wp A Q \ wp B Q" by(blast) finally have "P \ wp B Q" . with uP uQ drBC show "\G\ && (P o \) \ wp C (Q o \)" by(blast intro:drefinesD) qed lemma dr_pr_trans[trans]: assumes drAB: "drefines \ G A B" assumes prBC: "B \ C" shows "drefines \ G A C" proof(rule drefinesI) fix P and Q assume uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" with drAB have "\G\ && (P o \) \ wp B (Q o \)" by(blast intro:drefinesD) also from uQ prBC have "... \ wp C (Q o \)" by(blast) finally show "\G\ && (P o \) \ ..." . qed text \If the projection @{term \} commutes with the transformer, then data refinement is reflexive:\ lemma dr_refl: assumes wa: "well_def a" and comm: "\Q. unitary Q \ wp a Q o \ \ wp a (Q o \)" shows "drefines \ G a a" proof(intro drefinesI2 wa) fix P and Q and s assume wp: "P \ wp a Q" assume uQ: "unitary Q" have "(P o \) s = P (\ s)" by(simp) also from wp have "... \ wp a Q (\ s)" by(blast) also { from comm uQ have "wp a Q o \ \ wp a (Q o \)" by(blast) hence "(wp a Q o \) s \ wp a (Q o \) s" by(blast) hence "wp a Q (\ s) \ ..." by(simp) } finally show "(P o \) s \ wp a (Q o \) s" . qed text \Correspondence implies data refinement\ lemma pcorres_drefine: assumes corres: "pcorres \ G A C" and wC: "well_def C" shows "drefines \ G A C" proof fix P and Q assume uP: "unitary P" and uQ: "unitary Q" and wpA: "P \ wp A Q" from wpA have "P o \ \ wp A Q o \" by(simp add:o_def le_fun_def) hence "\G\ && (P o \) \ \G\ && (wp A Q o \)" by(rule exp_conj_mono_right) also from corres uQ have "... = \G\ && (wp C (Q o \))" by(rule pcorresD) also have "... \ wp C (Q o \)" proof(rule le_funI) fix s from uQ have "unitary (Q o \)" by(rule unitary_intros) with well_def_wp_healthy[OF wC] have nn_wpC: "0 \ wp C (Q o \) s" by(blast) show "(\G\ && wp C (Q o \)) s \ wp C (Q o \) s" proof(cases "G s") case True with nn_wpC show ?thesis by(simp add:exp_conj_def) next case False note this moreover { from uQ have "unitary (Q o \)" by(simp) with well_def_wp_healthy[OF wC] have "wp C (Q o \) s \ 1" by(auto) } moreover note nn_wpC ultimately show ?thesis by(simp add:exp_conj_def) qed qed finally show "\G\ && (P o \) \ wp C (Q o \)" . qed text \Any \emph{data} refinement of a deterministic program is correspondence. This is the analogous result to that relating program refinement and equivalence.\ lemma drefines_determ: fixes a::"'a prog" and b::"'b prog" assumes da: "determ (wp a)" and wa: "well_def a" and wb: "well_def b" and dr: "drefines \ G a b" shows "pcorres \ G a b" txt \The proof follows exactly the same form as that for program refinement: Assuming that correspondence \emph{doesn't} hold, we show that @{term "wp b"} is not feasible, and thus not healthy, contradicting the assumption.\ proof(rule pcorresI, rule contrapos_pp) from wb show "feasible (wp b)" by(auto) note ha = well_def_wp_healthy[OF wa] note hb = well_def_wp_healthy[OF wb] from wb have "sublinear (wp b)" by(auto) moreover from hb have "feasible (wp b)" by(auto) ultimately have sab: "sub_add (wp b)" by(rule sublinear_subadd) fix Q::"'a \ real" assume uQ: "unitary Q" hence uQ\: "unitary (Q o \)" by(auto) assume ne: "\G\ && (wp a Q o \) \ \G\ && wp b (Q o \)" hence ne': "wp a Q o \ \ wp b (Q o \)" by(auto) txt \From refinement, @{term "\G\ && (wp a Q o \)"} lies below @{term "\G\ && wp b (Q o \)"}.\ from ha uQ have gle: "\G\ && (wp a Q o \) \ wp b (Q o \)" by(blast intro!:drefinesD[OF dr]) have le: "\G\ && (wp a Q o \) \ \G\ && wp b (Q o \)" unfolding exp_conj_def proof(rule le_funI) fix s from gle have "\G\ s .& (wp a Q o \) s \ wp b (Q o \) s" unfolding exp_conj_def by(auto) hence "\G\ s .& (\G\ s .& (wp a Q o \) s) \ \G\ s .& wp b (Q o \) s" by(auto intro:pconj_mono) moreover from uQ ha have "wp a Q (\ s) \ 1" by(auto dest:healthy_bounded_byD) moreover from uQ ha have "0 \ wp a Q (\ s)" by(auto) ultimately show "\ G \ s .& (wp a Q \ \) s \ \ G \ s .& wp b (Q \ \) s" by(simp add:pconj_assoc) qed txt \If the programs do not correspond, the terms must differ somewhere, and given the previous result, the second must be somewhere strictly larger than the first:\ have nle: "\s. (\G\ && (wp a Q o \)) s < (\G\ && wp b (Q o \)) s" proof(rule contrapos_np[OF ne], rule ext, rule antisym) fix s from le show "(\G\ && (wp a Q o \)) s \ (\G\ && wp b (Q o \)) s" by(blast) next fix s assume "\ (\s. (\G\ && (wp a Q \ \)) s < (\G\ && wp b (Q \ \)) s)" thus " (\G\ && (wp b (Q \ \))) s \ (\G\ && (wp a Q \ \)) s" by(simp add:not_less) qed from this obtain s where less_s: "(\G\ && (wp a Q \ \)) s < (\G\ && wp b (Q \ \)) s" by(blast) txt \The transformers themselves must differ at this point:\ hence larger: "wp a Q (\ s) < wp b (Q \ \) s" proof(cases "G s") case True moreover from ha uQ have "0 \ wp a Q (\ s)" by(blast) moreover from hb uQ\ have "0 \ wp b (Q o \) s" by(blast) moreover note less_s ultimately show ?thesis by(simp add:exp_conj_def) next case False moreover from ha uQ have "wp a Q (\ s) \ 1" by(blast) moreover { from uQ have "bounded_by 1 (Q o \)" by(blast) moreover from unitary_sound[OF uQ] have "sound (Q o \)" by(auto) ultimately have "wp b (Q o \) s \ 1" using hb by(auto) } moreover note less_s ultimately show ?thesis by(simp add:exp_conj_def) qed from less_s have "(\G\ && (wp a Q \ \)) s \ (\G\ && wp b (Q \ \)) s" by(force) txt \@{term G} must also hold, as otherwise both would be zero.\ hence G_s: "G s" proof(rule contrapos_np) assume nG: "\ G s" moreover from ha uQ have "wp a Q (\ s) \ 1" by(blast) moreover { from uQ have "bounded_by 1 (Q o \)" by(blast) moreover from unitary_sound[OF uQ] have "sound (Q o \)" by(auto) ultimately have "wp b (Q o \) s \ 1" using hb by(auto) } ultimately show "(\G\ && (wp a Q \ \)) s = (\G\ && wp b (Q \ \)) s" by(simp add:exp_conj_def) qed txt \Take a carefully constructed expectation:\ let ?Qc = "\s. bound_of Q - Q s" have bQc: "bounded_by 1 ?Qc" proof(rule bounded_byI) fix s from uQ have "bound_of Q \ 1" and "0 \ Q s" by(auto) thus "bound_of Q - Q s \ 1" by(auto) qed have sQc: "sound ?Qc" proof(rule soundI) from bQc show "bounded ?Qc" by(auto) show "nneg ?Qc" proof(rule nnegI) fix s from uQ have "Q s \ bound_of Q" by(auto) thus "0 \ bound_of Q - Q s" by(auto) qed qed txt \By the maximality of @{term "wp a"}, @{term "wp b"} must violate feasibility, by mapping @{term s} to something strictly greater than @{term "bound_of Q"}.\ from uQ have "0 \ bound_of Q" by(auto) with da have "bound_of Q = wp a (\s. bound_of Q) (\ s)" by(simp add:maximalD determ_maximalD) also have "wp a (\s. bound_of Q) (\ s) = wp a (\s. Q s + ?Qc s) (\ s)" by(simp) also { from da have "additive (wp a)" by(blast) with uQ sQc have "wp a (\s. Q s + ?Qc s) (\ s) = wp a Q (\ s) + wp a ?Qc (\ s)" by(subst additiveD, blast+) } also { from ha and sQc and bQc have "\G\ && (wp a ?Qc o \) \ wp b (?Qc o \)" by(blast intro!:drefinesD[OF dr]) hence "(\G\ && (wp a ?Qc o \)) s \ wp b (?Qc o \) s" by(blast) moreover from sQc and ha have "0 \ wp a (\s. bound_of Q - Q s) (\ s)" by(blast) ultimately have "wp a ?Qc (\ s) \ wp b (?Qc o \) s" using G_s by(simp add:exp_conj_def) hence "wp a Q (\ s) + wp a ?Qc (\ s) \ wp a Q (\ s) + wp b (?Qc o \) s" by(rule add_left_mono) also with larger have "wp a Q (\ s) + wp b (?Qc o \) s < wp b (Q o \) s + wp b (?Qc o \) s" by(auto) finally have "wp a Q (\ s) + wp a ?Qc (\ s) < wp b (Q o \) s + wp b (?Qc o \) s" . } also from sab and unitary_sound[OF uQ] and sQc have "wp b (Q o \) s + wp b (?Qc o \) s \ wp b (\s. (Q o \) s + (?Qc o \) s) s" by(blast) also have "... = wp b (\s. bound_of Q) s" by(simp) finally show "\ feasible (wp b)" proof(rule contrapos_pn) assume fb: "feasible (wp b)" have "bounded_by (bound_of Q) (\s. bound_of Q)" by(blast) hence "bounded_by (bound_of Q) (wp b (\s. bound_of Q))" using uQ by(blast intro:feasible_boundedD[OF fb]) hence "wp b (\s. bound_of Q) s \ bound_of Q" by(blast) thus "\ bound_of Q < wp b (\s. bound_of Q) s" by(simp) qed qed subsection \Structural Rules for Correspondence\ lemma pcorres_Skip: "pcorres \ G Skip Skip" by(simp add:pcorres_def wp_eval) text \Correspondence composes over sequential composition.\ lemma pcorres_Seq: fixes A::"'b prog" and B::"'c prog" and C::"'b prog" and D::"'c prog" and \::"'c \ 'b" assumes pcAB: "pcorres \ G A B" and pcCD: "pcorres \ H C D" and wA: "well_def A" and wB: "well_def B" and wC: "well_def C" and wD: "well_def D" and p3p2: "\Q. unitary Q \ \I\ && wp B Q = wp B (\H\ && Q)" and p1p3: "\s. G s \ I s" shows "pcorres \ G (A;;C) (B;;D)" proof(rule pcorresI) fix Q::"'b \ real" assume uQ: "unitary Q" with well_def_wp_healthy[OF wC] have uCQ: "unitary (wp C Q)" by(auto) from uQ well_def_wp_healthy[OF wD] have uDQ: "unitary (wp D (Q o \))" by(auto dest:unitary_comp) have p3p1: "\R S. \ unitary R; unitary S; \I\ && R = \I\ && S \ \ \G\ && R = \G\ && S" proof(rule ext) fix R::"'c \ real" and S::"'c \ real" and s::'c assume a3: "\I\ && R = \I\ && S" and uR: "unitary R" and uS: "unitary S" show "(\G\ && R) s = (\G\ && S) s" proof(simp add:exp_conj_def, cases "G s") case False note this moreover from uR have "R s \ 1" by(blast) moreover from uS have "S s \ 1" by(blast) ultimately show "\G\ s .& R s = \G\ s .& S s" by(simp) next case True note p1 = this with p1p3 have "I s" by(blast) with fun_cong[OF a3, where x=s] have "1 .& R s = 1 .& S s" by(simp add:exp_conj_def) with p1 show "\G\ s .& R s = \G\ s .& S s" by(simp) qed qed show "\G\ && (wp (A;;C) Q o \) = \G\ && wp (B;;D) (Q o \)" proof(simp add:wp_eval) from uCQ pcAB have "\G\ && (wp A (wp C Q) \ \) = \G\ && wp B ((wp C Q) \ \)" by(auto dest:pcorresD) also have "\G\ && wp B ((wp C Q) \ \) = \G\ && wp B (wp D (Q \ \))" proof(rule p3p1) from uCQ well_def_wp_healthy[OF wB] show "unitary (wp B (wp C Q \ \))" by(auto intro:unitary_comp) from uDQ well_def_wp_healthy[OF wB] show "unitary (wp B (wp D (Q \ \)))" by(auto) from uQ have "\ H \ && (wp C Q \ \) = \ H \ && wp D (Q \ \)" by(blast intro:pcorresD[OF pcCD]) thus "\ I \ && wp B (wp C Q \ \) = \ I \ && wp B (wp D (Q \ \))" by(simp add:p3p2 uCQ uDQ) qed finally show "\G\ && (wp A (wp C Q) \ \) = \G\ && wp B (wp D (Q \ \))" . qed qed subsection \Structural Rules for Data Refinement\ lemma dr_Skip: fixes \::"'c \ 'b" shows "drefines \ G Skip Skip" proof(intro drefinesI2 wd_intros) fix P::"'b \ real" and Q::"'b \ real" and s::'c assume "P \ wp Skip Q" hence "(P o \) s \ wp Skip Q (\ s)" by(simp, blast) thus "(P o \) s \ wp Skip (Q o \) s" by(simp add:wp_eval) qed lemma dr_Abort: fixes \::"'c \ 'b" shows "drefines \ G Abort Abort" proof(intro drefinesI2 wd_intros) fix P::"'b \ real" and Q::"'b \ real" and s::'c assume "P \ wp Abort Q" hence "(P o \) s \ wp Abort Q (\ s)" by(auto) thus "(P o \) s \ wp Abort (Q o \) s" by(simp add:wp_eval) qed lemma dr_Apply: fixes \::"'c \ 'b" assumes commutes: "f o \ = \ o g" shows "drefines \ G (Apply f) (Apply g)" proof(intro drefinesI2 wd_intros) fix P::"'b \ real" and Q::"'b \ real" and s::'c assume wp: "P \ wp (Apply f) Q" hence "P \ (Q o f)" by(simp add:wp_eval) hence "P (\ s) \ (Q o f) (\ s)" by(blast) also have "... = Q ((f o \) s)" by(simp) also with commutes have "... = ((Q o \) o g) s" by(simp) also have "... = wp (Apply g) (Q o \) s" by(simp add:wp_eval) finally show "(P o \) s \ wp (Apply g) (Q o \) s" by(simp) qed lemma dr_Seq: assumes drAB: "drefines \ P A B" and drBC: "drefines \ Q C D" and wpB: "\P\ \ wp B \Q\" and wB: "well_def B" and wC: "well_def C" and wD: "well_def D" shows "drefines \ P (A;;C) (B;;D)" proof fix R and S assume uR: "unitary R" and uS: "unitary S" and wpAC: "R \ wp (A;;C) S" from uR have "\P\ && (R o \) = \P\ && (\P\ && (R o \))" by(simp add:exp_conj_assoc) also { from well_def_wp_healthy[OF wC] uR uS and wpAC[unfolded eval_wp_Seq o_def] have "\P\ && (R o \) \ wp B (wp C S o \)" by(auto intro:drefinesD[OF drAB]) with wpB well_def_wp_healthy[OF wC] uS sublinear_sub_conj[OF well_def_wp_sublinear, OF wB] have "\P\ && (\P\ && (R o \)) \ wp B (\Q\ && (wp C S o \))" by(auto intro!:entails_combine dest!:unitary_sound) } also { from uS well_def_wp_healthy[OF wC] have "\Q\ && (wp C S o \) \ wp D (S o \)" by(auto intro!:drefinesD[OF drBC]) with well_def_wp_healthy[OF wB] well_def_wp_healthy[OF wC] well_def_wp_healthy[OF wD] and unitary_sound[OF uS] have "wp B (\Q\ && (wp C S o \)) \ wp B (wp D (S o \))" by(blast intro!:mono_transD) } finally show "\P\ && (R o \) \ wp (B;;D) (S o \)" unfolding wp_eval o_def . qed lemma dr_repeat: fixes \ :: "'a \ 'b" assumes dr_ab: "drefines \ G a b" and Gpr: "\G\ \ wp b \G\" and wa: "well_def a" and wb: "well_def b" shows "drefines \ G (repeat n a) (repeat n b)" (is "?X n") proof(induct n) show "?X 0" by(simp add:dr_Skip) fix n assume IH: "?X n" thus "?X (Suc n)" by(auto intro!:dr_Seq Gpr assms wd_intros) qed end diff --git a/thys/pGCL/Continuity.thy b/thys/pGCL/Continuity.thy --- a/thys/pGCL/Continuity.thy +++ b/thys/pGCL/Continuity.thy @@ -1,1010 +1,1016 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section \Continuity\ theory Continuity imports Healthiness begin text \We rely on one additional healthiness property, continuity, which is shown here seperately, as its proof relies, in general, on healthiness. It is only relevant when a program appears in an inductive context i.e.~inside a loop.\ text \A continuous transformer preserves limits (or the suprema of ascending chains).\ definition bd_cts :: "'s trans \ bool" where "bd_cts t = (\M. (\i. (M i \ M (Suc i)) \ sound (M i)) \ (\b. \i. bounded_by b (M i)) \ t (Sup_exp (range M)) = Sup_exp (range (t o M)))" lemma bd_ctsD: "\ bd_cts t; \i. M i \ M (Suc i); \i. sound (M i); \i. bounded_by b (M i) \ \ t (Sup_exp (range M)) = Sup_exp (range (t o M))" unfolding bd_cts_def by(auto) lemma bd_ctsI: "(\b M. (\i. M i \ M (Suc i)) \ (\i. sound (M i)) \ (\i. bounded_by b (M i)) \ t (Sup_exp (range M)) = Sup_exp (range (t o M))) \ bd_cts t" unfolding bd_cts_def by(auto) text \A generalised property for transformers of transformers.\ definition bd_cts_tr :: "('s trans \ 's trans) \ bool" where "bd_cts_tr T = (\M. (\i. le_trans (M i) (M (Suc i)) \ feasible (M i)) \ equiv_trans (T (Sup_trans (M ` UNIV))) (Sup_trans ((T o M) ` UNIV)))" lemma bd_cts_trD: "\ bd_cts_tr T; \i. le_trans (M i) (M (Suc i)); \i. feasible (M i) \ \ equiv_trans (T (Sup_trans (M ` UNIV))) (Sup_trans ((T o M) ` UNIV))" by(simp add:bd_cts_tr_def) lemma bd_cts_trI: "(\M. (\i. le_trans (M i) (M (Suc i))) \ (\i. feasible (M i)) \ equiv_trans (T (Sup_trans (M ` UNIV))) (Sup_trans ((T o M) ` UNIV))) \ bd_cts_tr T" by(simp add:bd_cts_tr_def) subsection \Continuity of Primitives\ lemma cts_wp_Abort: "bd_cts (wp (Abort::'s prog))" proof - have X: "range (\(i::nat) (s::'s). 0) = {\s. 0}" by(auto) show ?thesis by(intro bd_ctsI, simp add:wp_eval o_def Sup_exp_def X) qed lemma cts_wp_Skip: "bd_cts (wp Skip)" by(rule bd_ctsI, simp add:wp_def Skip_def o_def) lemma cts_wp_Apply: "bd_cts (wp (Apply f))" proof - have X: "\M s. {P (f s) |P. P \ range M} = {P s |P. P \ range (\i s. M i (f s))}" by(auto) show ?thesis by(intro bd_ctsI ext, simp add:wp_eval o_def Sup_exp_def X) qed lemma cts_wp_Bind: fixes a::"'a \ 's prog" assumes ca: "\s. bd_cts (wp (a (f s)))" shows "bd_cts (wp (Bind f a))" proof(rule bd_ctsI) fix M::"nat \ 's expect" and c::real assume chain: "\i. M i \ M (Suc i)" and sM: "\i. sound (M i)" and bM: "\i. bounded_by c (M i)" with bd_ctsD[OF ca] have "\s. wp (a (f s)) (Sup_exp (range M)) = Sup_exp (range (wp (a (f s)) o M))" by(auto) moreover have "\s. {fa s |fa. fa \ range (\x. wp (a (f s)) (M x))} = {fa s |fa. fa \ range (\x s. wp (a (f s)) (M x) s)}" by(auto) ultimately show "wp (Bind f a) (Sup_exp (range M)) = Sup_exp (range (wp (Bind f a) \ M))" by(simp add:wp_eval o_def Sup_exp_def) qed text \The first nontrivial proof. We transform the suprema into limits, and appeal to the continuity of the underlying operation (here infimum). This is typical of the remainder of the nonrecursive elements.\ lemma cts_wp_DC: fixes a b::"'s prog" assumes ca: "bd_cts (wp a)" and cb: "bd_cts (wp b)" and ha: "healthy (wp a)" and hb: "healthy (wp b)" shows "bd_cts (wp (a \ b))" proof(rule bd_ctsI, rule antisym) fix M::"nat \ 's expect" and c::real assume chain: "\i. M i \ M (Suc i)" and sM: "\i. sound (M i)" and bM: "\i. bounded_by c (M i)" from ha hb have hab: "healthy (wp (a \ b))" by(rule healthy_intros) from bM have leSup: "\i. M i \ Sup_exp (range M)" by(auto intro:Sup_exp_upper) from sM bM have sSup: "sound (Sup_exp (range M))" by(auto intro:Sup_exp_sound) show "Sup_exp (range (wp (a \ b) \ M)) \ wp (a \ b) (Sup_exp (range M))" proof(rule Sup_exp_least, clarsimp, rule le_funI) fix i s from mono_transD[OF healthy_monoD[OF hab]] leSup sM sSup have "wp (a \ b) (M i) \ wp (a \ b) (Sup_exp (range M))" by(auto) thus "wp (a \ b) (M i) s \ wp (a \ b) (Sup_exp (range M)) s" by(auto) from hab sSup have "sound (wp (a \ b) (Sup_exp (range M)))" by(auto) thus "nneg (wp (a \ b) (Sup_exp (range M)))" by(auto) qed from sM bM ha have "\i. bounded_by c (wp a (M i))" by(auto) hence baM: "\i s. wp a (M i) s \ c" by(auto) from sM bM hb have "\i. bounded_by c (wp b (M i))" by(auto) hence bbM: "\i s. wp b (M i) s \ c" by(auto) show "wp (a \ b) (Sup_exp (range M)) \ Sup_exp (range (wp (a \ b) \ M))" proof(simp add:wp_eval o_def, rule le_funI) fix s::'s from bd_ctsD[OF ca, of M, OF chain sM bM] bd_ctsD[OF cb, of M, OF chain sM bM] have "min (wp a (Sup_exp (range M)) s) (wp b (Sup_exp (range M)) s) = min (Sup_exp (range (wp a o M)) s) (Sup_exp (range (wp b o M)) s)" by(simp) also { have "{f s |f. f \ range (\x. wp a (M x))} = range (\i. wp a (M i) s)" "{f s |f. f \ range (\x. wp b (M x))} = range (\i. wp b (M i) s)" by(auto) hence "min (Sup_exp (range (wp a o M)) s) (Sup_exp (range (wp b o M)) s) = min (Sup (range (\i. wp a (M i) s))) (Sup (range (\i. wp b (M i) s)))" by(simp add:Sup_exp_def o_def) } also { have "(\i. wp a (M i) s) \ Sup (range (\i. wp a (M i) s))" proof(rule increasing_LIMSEQ) fix n from mono_transD[OF healthy_monoD, OF ha] sM chain show "wp a (M n) s \ wp a (M (Suc n)) s" by(auto intro:le_funD) from baM show "wp a (M n) s \ Sup (range (\i. wp a (M i) s))" by(intro cSup_upper bdd_aboveI, auto) fix e::real assume pe: "0 < e" from baM have cSup: "Sup (range (\i. wp a (M i) s)) \ closure (range (\i. wp a (M i) s))" by(blast intro:closure_contains_Sup) with pe obtain y where yin: "y \ (range (\i. wp a (M i) s))" and dy: "dist y (Sup (range (\i. wp a (M i) s))) < e" by(blast dest:iffD1[OF closure_approachable]) from yin obtain i where "y = wp a (M i) s" by(auto) with dy have "dist (wp a (M i) s) (Sup (range (\i. wp a (M i) s))) < e" by(simp) moreover from baM have "wp a (M i) s \ Sup (range (\i. wp a (M i) s))" by(intro cSup_upper bdd_aboveI, auto) ultimately have "Sup (range (\i. wp a (M i) s)) \ wp a (M i) s + e" by(simp add:dist_real_def) thus "\i. Sup (range (\i. wp a (M i) s)) \ wp a (M i) s + e" by(auto) qed moreover have "(\i. wp b (M i) s) \ Sup (range (\i. wp b (M i) s))" proof(rule increasing_LIMSEQ) fix n from mono_transD[OF healthy_monoD, OF hb] sM chain show "wp b (M n) s \ wp b (M (Suc n)) s" by(auto intro:le_funD) from bbM show "wp b (M n) s \ Sup (range (\i. wp b (M i) s))" by(intro cSup_upper bdd_aboveI, auto) fix e::real assume pe: "0 < e" from bbM have cSup: "Sup (range (\i. wp b (M i) s)) \ closure (range (\i. wp b (M i) s))" by(blast intro:closure_contains_Sup) with pe obtain y where yin: "y \ (range (\i. wp b (M i) s))" and dy: "dist y (Sup (range (\i. wp b (M i) s))) < e" by(blast dest:iffD1[OF closure_approachable]) from yin obtain i where "y = wp b (M i) s" by(auto) with dy have "dist (wp b (M i) s) (Sup (range (\i. wp b (M i) s))) < e" by(simp) moreover from bbM have "wp b (M i) s \ Sup (range (\i. wp b (M i) s))" by(intro cSup_upper bdd_aboveI, auto) ultimately have "Sup (range (\i. wp b (M i) s)) \ wp b (M i) s + e" by(simp add:dist_real_def) thus "\i. Sup (range (\i. wp b (M i) s)) \ wp b (M i) s + e" by(auto) qed ultimately have "(\i. min (wp a (M i) s) (wp b (M i) s)) \ min (Sup (range (\i. wp a (M i) s))) (Sup (range (\i. wp b (M i) s)))" by(rule tendsto_min) moreover have "bdd_above (range (\i. min (wp a (M i) s) (wp b (M i) s)))" proof(intro bdd_aboveI, clarsimp) fix i have "min (wp a (M i) s) (wp b (M i) s) \ wp a (M i) s" by(auto) also { from ha sM bM have "bounded_by c (wp a (M i))" by(auto) hence "wp a (M i) s \ c" by(auto) } finally show "min (wp a (M i) s) (wp b (M i) s) \ c" . qed ultimately have "min (Sup (range (\i. wp a (M i) s))) (Sup (range (\i. wp b (M i) s))) \ Sup (range (\i. min (wp a (M i) s) (wp b (M i) s)))" by(blast intro:LIMSEQ_le_const2 cSup_upper min.mono[OF baM bbM]) } also { have "range (\i. min (wp a (M i) s) (wp b (M i) s)) = {f s |f. f \ range (\i s. min (wp a (M i) s) (wp b (M i) s))}" by(auto) hence "Sup (range (\i. min (wp a (M i) s) (wp b (M i) s))) = Sup_exp (range (\i s. min (wp a (M i) s) (wp b (M i) s))) s" by (simp add: Sup_exp_def cong del: SUP_cong_simp) } finally show "min (wp a (Sup_exp (range M)) s) (wp b (Sup_exp (range M)) s) \ Sup_exp (range (\i s. min (wp a (M i) s) (wp b (M i) s))) s" . qed qed lemma cts_wp_Seq: fixes a b::"'s prog" assumes ca: "bd_cts (wp a)" and cb: "bd_cts (wp b)" and hb: "healthy (wp b)" shows "bd_cts (wp (a ;; b))" proof(rule bd_ctsI, simp add:o_def wp_eval) fix M::"nat \ 's expect" and c::real assume chain: "\i. M i \ M (Suc i)" and sM: "\i. sound (M i)" and bM: "\i. bounded_by c (M i)" hence "wp a (wp b (Sup_exp (range M))) = wp a (Sup_exp (range (wp b o M)))" by(subst bd_ctsD[OF cb], auto) also { from sM hb have "\i. sound ((wp b o M) i)" by(auto) moreover from chain sM have "\i. (wp b o M) i \ (wp b o M) (Suc i)" by(auto intro:mono_transD[OF healthy_monoD, OF hb]) moreover from sM bM hb have "\i. bounded_by c ((wp b o M) i)" by(auto) ultimately have "wp a (Sup_exp (range (wp b o M))) = Sup_exp (range (wp a o (wp b o M)))" by(subst bd_ctsD[OF ca], auto) } also have "Sup_exp (range (wp a o (wp b o M))) = Sup_exp (range (\i. wp a (wp b (M i))))" by(simp add:o_def) finally show "wp a (wp b (Sup_exp (range M))) = Sup_exp (range (\i. wp a (wp b (M i))))" . qed lemma cts_wp_PC: fixes a b::"'s prog" assumes ca: "bd_cts (wp a)" and cb: "bd_cts (wp b)" and ha: "healthy (wp a)" and hb: "healthy (wp b)" and up: "unitary p" shows "bd_cts (wp (PC a p b))" proof(rule bd_ctsI, rule ext, simp add:o_def wp_eval) fix M::"nat \ 's expect" and c::real and s::'s assume chain: "\i. M i \ M (Suc i)" and sM: "\i. sound (M i)" and bM: "\i. bounded_by c (M i)" from sM have "\i. nneg (M i)" by(auto) with bM have nc: "0 \ c" by(auto) from chain sM bM have "wp a (Sup_exp (range M)) = Sup_exp (range (wp a o M))" by(rule bd_ctsD[OF ca]) hence "wp a (Sup_exp (range M)) s = Sup_exp (range (wp a o M)) s" by(simp) also { have "{f s |f. f \ range (\x. wp a (M x))} = range (\i. wp a (M i) s)" by(auto) hence "Sup_exp (range (wp a o M)) s = Sup (range (\i. wp a (M i) s))" by(simp add:Sup_exp_def o_def) } finally have "p s * wp a (Sup_exp (range M)) s = p s * Sup (range (\i. wp a (M i) s))" by(simp) also have "... = Sup {p s * x |x. x \ range (\i. wp a (M i) s)}" proof(rule cSup_mult, blast, clarsimp) from up show "0 \ p s" by(auto) fix i from sM bM ha have "bounded_by c (wp a (M i))" by(auto) thus "wp a (M i) s \ c" by(auto) qed also { have "{p s * x |x. x \ range (\i. wp a (M i) s)} = range (\i. p s * wp a (M i) s)" by(auto) hence "Sup {p s * x |x. x \ range (\i. wp a (M i) s)} = Sup (range (\i. p s * wp a (M i) s))" by(simp) } finally have "p s * wp a (Sup_exp (range M)) s = Sup (range (\i. p s * wp a (M i) s))" . moreover { from chain sM bM have "wp b (Sup_exp (range M)) = Sup_exp (range (wp b o M))" by(rule bd_ctsD[OF cb]) hence "wp b (Sup_exp (range M)) s = Sup_exp (range (wp b o M)) s" by(simp) also { have "{f s |f. f \ range (\x. wp b (M x))} = range (\i. wp b (M i) s)" by(auto) hence "Sup_exp (range (wp b o M)) s = Sup (range (\i. wp b (M i) s))" by(simp add:Sup_exp_def o_def) } finally have "(1 - p s) * wp b (Sup_exp (range M)) s = (1 - p s) * Sup (range (\i. wp b (M i) s))" by(simp) also have "... = Sup {(1 - p s) * x |x. x \ range (\i. wp b (M i) s)}" proof(rule cSup_mult, blast, clarsimp) - from up show "0 \ 1 - p s" by(auto simp:sign_simps) + from up show "0 \ 1 - p s" + by auto fix i from sM bM hb have "bounded_by c (wp b (M i))" by(auto) thus "wp b (M i) s \ c" by(auto) qed also { have "{(1 - p s) * x |x. x \ range (\i. wp b (M i) s)} = range (\i. (1 - p s) * wp b (M i) s)" by(auto) hence "Sup {(1 - p s) * x |x. x \ range (\i. wp b (M i) s)} = Sup (range (\i. (1 - p s) * wp b (M i) s))" by(simp) } finally have "(1 - p s) * wp b (Sup_exp (range M)) s = Sup (range (\i. (1 - p s) * wp b (M i) s))" . } ultimately have "p s * wp a (Sup_exp (range M)) s + (1 - p s) * wp b (Sup_exp (range M)) s = Sup (range (\i. p s * wp a (M i) s)) + Sup (range (\i. (1 - p s) * wp b (M i) s))" by(simp) also { from bM sM ha have "\i. bounded_by c (wp a (M i))" by(auto) hence "\i. wp a (M i) s \ c" by(auto) moreover from up have "0 \ p s" by(auto) ultimately have "\i. p s * wp a (M i) s \ p s * c" by(auto intro:mult_left_mono) also from up nc have "p s * c \ 1 * c" by(blast intro:mult_right_mono) also have "... = c" by(simp) finally have baM: "\i. p s * wp a (M i) s \ c" . have lima: "(\i. p s * wp a (M i) s) \ Sup (range (\i. p s * wp a (M i) s))" proof(rule increasing_LIMSEQ) fix n from sM chain healthy_monoD[OF ha] have "wp a (M n) \ wp a (M (Suc n))" by(auto) with up show "p s * wp a (M n) s \ p s * wp a (M (Suc n)) s" by(blast intro:mult_left_mono) from baM show "p s * wp a (M n) s \ Sup (range (\i. p s * wp a (M i) s))" by(intro cSup_upper bdd_aboveI, auto) next fix e::real assume pe: "0 < e" from baM have "Sup (range (\i. p s * wp a (M i) s)) \ closure (range (\i. p s * wp a (M i) s))" by(blast intro:closure_contains_Sup) thm closure_approachable with pe obtain y where yin: "y \ range (\i. p s * wp a (M i) s)" and dy: "dist y (Sup (range (\i. p s * wp a (M i) s))) < e" by(blast dest:iffD1[OF closure_approachable]) from yin obtain i where "y = p s * wp a (M i) s" by(auto) with dy have "dist (p s * wp a (M i) s) (Sup (range (\i. p s * wp a (M i) s))) < e" by(simp) moreover from baM have "p s * wp a (M i) s \ Sup (range (\i. p s * wp a (M i) s))" by(intro cSup_upper bdd_aboveI, auto) ultimately have "Sup (range (\i. p s * wp a (M i) s)) \ p s * wp a (M i) s + e" by(simp add:dist_real_def) thus "\i. Sup (range (\i. p s * wp a (M i) s)) \ p s * wp a (M i) s + e" by(auto) qed from bM sM hb have "\i. bounded_by c (wp b (M i))" by(auto) hence "\i. wp b (M i) s \ c" by(auto) - moreover from up have "0 \ (1 - p s)" by(auto simp:sign_simps) + moreover from up have "0 \ (1 - p s)" + by auto ultimately have "\i. (1 - p s) * wp b (M i) s \ (1 - p s) * c" by(auto intro:mult_left_mono) also { from up have "1 - p s \ 1" by(auto) with nc have "(1 - p s) * c \ 1 * c" by(blast intro:mult_right_mono) } also have "1 * c = c" by(simp) finally have bbM: "\i. (1 - p s) * wp b (M i) s \ c" . have limb: "(\i. (1 - p s) * wp b (M i) s) \ Sup (range (\i. (1 - p s) * wp b (M i) s))" proof(rule increasing_LIMSEQ) fix n from sM chain healthy_monoD[OF hb] have "wp b (M n) \ wp b (M (Suc n))" by(auto) - moreover from up have "0 \ 1 - p s" by(auto simp:sign_simps) + moreover from up have "0 \ 1 - p s" + by auto ultimately show "(1 - p s) * wp b (M n) s \ (1 - p s) * wp b (M (Suc n)) s" by(blast intro:mult_left_mono) from bbM show "(1 - p s) * wp b (M n) s \ Sup (range (\i. (1 - p s) * wp b (M i) s))" by(intro cSup_upper bdd_aboveI, auto) next fix e::real assume pe: "0 < e" from bbM have "Sup (range (\i. (1 - p s) * wp b (M i) s)) \ closure (range (\i. (1 - p s) * wp b (M i) s))" by(blast intro:closure_contains_Sup) with pe obtain y where yin: "y \ range (\i. (1 - p s) * wp b (M i) s)" and dy: "dist y (Sup (range (\i. (1 - p s) * wp b (M i) s))) < e" by(blast dest:iffD1[OF closure_approachable]) from yin obtain i where "y = (1 - p s) * wp b (M i) s" by(auto) with dy have "dist ((1 - p s) * wp b (M i) s) (Sup (range (\i. (1 - p s) * wp b (M i) s))) < e" by(simp) moreover from bbM have "(1 - p s) * wp b (M i) s \ Sup (range (\i. (1 - p s) * wp b (M i) s))" by(intro cSup_upper bdd_aboveI, auto) ultimately have "Sup (range (\i. (1 - p s) * wp b (M i) s)) \ (1 - p s) * wp b (M i) s + e" by(simp add:dist_real_def) thus "\i. Sup (range (\i. (1 - p s) * wp b (M i) s)) \ (1 - p s) * wp b (M i) s + e" by(auto) qed from lima limb have "(\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s) \ Sup (range (\i. p s * wp a (M i) s)) + Sup (range (\i. (1 - p s) * wp b (M i) s))" by(rule tendsto_add) moreover from add_mono[OF baM bbM] have "\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s \ Sup (range (\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s))" by(intro cSup_upper bdd_aboveI, auto) ultimately have "Sup (range (\i. p s * wp a (M i) s)) + Sup (range (\i. (1 - p s) * wp b (M i) s)) \ Sup (range (\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s))" by(blast intro: LIMSEQ_le_const2) } also { have "range (\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s) = {f s |f. f \ range (\x s. p s * wp a (M x) s + (1 - p s) * wp b (M x) s)}" by(auto) hence "Sup (range (\i. p s * wp a (M i) s + (1 - p s) * wp b (M i) s)) = Sup_exp (range (\x s. p s * wp a (M x) s + (1 - p s) * wp b (M x) s)) s" by (simp add: Sup_exp_def cong del: SUP_cong_simp) } finally have "p s * wp a (Sup_exp (range M)) s + (1 - p s) * wp b (Sup_exp (range M)) s \ Sup_exp (range (\i s. p s * wp a (M i) s + (1 - p s) * wp b (M i) s)) s" . moreover have "Sup_exp (range (\i s. p s * wp a (M i) s + (1 - p s) * wp b (M i) s)) s \ p s * wp a (Sup_exp (range M)) s + (1 - p s) * wp b (Sup_exp (range M)) s" proof(rule le_funD[OF Sup_exp_least], clarsimp, rule le_funI) fix i::nat and s::'s from bM have leSup: "M i \ Sup_exp (range M)" by(blast intro: Sup_exp_upper) moreover from sM bM have sSup: "sound (Sup_exp (range M))" by(auto intro:Sup_exp_sound) moreover note healthy_monoD[OF ha] sM ultimately have "wp a (M i) \ wp a (Sup_exp (range M))" by(auto) hence "wp a (M i) s \ wp a (Sup_exp (range M)) s" by(auto) moreover { from leSup sSup healthy_monoD[OF hb] sM have "wp b (M i) \ wp b (Sup_exp (range M))" by(auto) hence "wp b (M i) s \ wp b (Sup_exp (range M)) s" by(auto) } - moreover from up have "0 \ p s" "0 \ 1 - p s" by(auto simp:sign_simps) + moreover from up have "0 \ p s" "0 \ 1 - p s" + by auto ultimately show "p s * wp a (M i) s + (1 - p s) * wp b (M i) s \ p s * wp a (Sup_exp (range M)) s + (1 - p s) * wp b (Sup_exp (range M)) s" by(blast intro:add_mono mult_left_mono) from sSup ha hb have "sound (wp a (Sup_exp (range M)))" "sound (wp b (Sup_exp (range M)))" by(auto) hence "\s. 0 \ wp a (Sup_exp (range M)) s" "\s. 0 \ wp b (Sup_exp (range M)) s" by(auto) - moreover from up have "\s. 0 \ p s" "\s. 0 \ 1 - p s" by(auto simp:sign_simps) + moreover from up have "\s. 0 \ p s" "\s. 0 \ 1 - p s" + by auto ultimately show "nneg (\c. p c * wp a (Sup_exp (range M)) c + (1 - p c) * wp b (Sup_exp (range M)) c)" by(blast intro:add_nonneg_nonneg mult_nonneg_nonneg) qed ultimately show "p s * wp a (Sup_exp (range M)) s + (1 - p s) * wp b (Sup_exp (range M)) s = Sup_exp (range (\x s. p s * wp a (M x) s + (1 - p s) * wp b (M x) s)) s" by(auto) qed text \Both set-based choice operators are only continuous for finite sets (probabilistic choice \emph{can} be extended infinitely, but we have not done so). The proofs for both are inductive, and rely on the above results on binary operators.\ lemma SetPC_Bind: "SetPC a p = Bind p (\p. SetPC a (\_. p))" by(intro ext, simp add:SetPC_def Bind_def Let_def) lemma SetPC_remove: assumes nz: "p x \ 0" and n1: "p x \ 1" and fsupp: "finite (supp p)" shows "SetPC a (\_. p) = PC (a x) (\_. p x) (SetPC a (\_. dist_remove p x))" proof(intro ext, simp add:SetPC_def PC_def) fix ab P s from nz have "x \ supp p" by(simp add:supp_def) hence "supp p = insert x (supp p - {x})" by(auto) hence "(\x\supp p. p x * a x ab P s) = (\x\insert x (supp p - {x}). p x * a x ab P s)" by(simp) also from fsupp have "... = p x * a x ab P s + (\x\supp p - {x}. p x * a x ab P s)" by(blast intro:sum.insert) also from n1 have "... = p x * a x ab P s + (1 - p x) * ((\x\supp p - {x}. p x * a x ab P s) / (1 - p x))" by(simp add:field_simps) also have "... = p x * a x ab P s + (1 - p x) * ((\y\supp p - {x}. (p y / (1 - p x)) * a y ab P s))" by(simp add:sum_divide_distrib) also have "... = p x * a x ab P s + (1 - p x) * ((\y\supp p - {x}. dist_remove p x y * a y ab P s))" by(simp add:dist_remove_def) also from nz n1 have "... = p x * a x ab P s + (1 - p x) * ((\y\supp (dist_remove p x). dist_remove p x y * a y ab P s))" by(simp add:supp_dist_remove) finally show "(\x\supp p. p x * a x ab P s) = p x * a x ab P s + (1 - p x) * (\y\supp (dist_remove p x). dist_remove p x y * a y ab P s)" . qed lemma cts_bot: "bd_cts (\(P::'s expect) (s::'s). 0::real)" proof - have X: "\s::'s. {(P::'s expect) s |P. P \ range (\P s. 0)} = {0}" by(auto) show ?thesis by(intro bd_ctsI, simp add:Sup_exp_def o_def X) qed lemma wp_SetPC_nil: "wp (SetPC a (\s a. 0)) = (\P s. 0)" by(intro ext, simp add:wp_eval) lemma SetPC_sgl: "supp p = {x} \ SetPC a (\_. p) = (\ab P s. p x * a x ab P s)" by(simp add:SetPC_def) lemma bd_cts_scale: fixes a::"'s trans" assumes ca: "bd_cts a" and ha: "healthy a" and nnc: "0 \ c" shows "bd_cts (\P s. c * a P s)" proof(intro bd_ctsI ext, simp add:o_def) fix M::"nat \ 's expect" and d::real and s::'s assume chain: "\i. M i \ M (Suc i)" and sM: "\i. sound (M i)" and bM: "\i. bounded_by d (M i)" from sM have "\i. nneg (M i)" by(auto) with bM have nnd: "0 \ d" by(auto) from sM bM have sSup: "sound (Sup_exp (range M))" by(auto intro:Sup_exp_sound) with healthy_scalingD[OF ha] nnc have "c * a (Sup_exp (range M)) s = a (\s. c * Sup_exp (range M) s) s" by(auto intro:scalingD) also { have "\s. {f s |f. f \ range M} = range (\i. M i s)" by(auto) hence "a (\s. c * Sup_exp (range M) s) s = a (\s. c * Sup (range (\i. M i s))) s" by(simp add:Sup_exp_def) } also { from bM have "\x s. x \ range (\i. M i s) \ x \ d" by(auto) with nnc have "a (\s. c * Sup (range (\i. M i s))) s = a (\s. Sup {c*x |x. x \ range (\i. M i s)}) s" by(subst cSup_mult, blast+) } also { have X: "\s. {c * x |x. x \ range (\i. M i s)} = range (\i. c * M i s)" by(auto) have "a (\s. Sup {c * x |x. x \ range (\i. M i s)}) s = a (\s. Sup (range (\i. c * M i s))) s" by(simp add:X) } also { have "\s. range (\i. c * M i s) = {f s |f. f \ range (\i s. c * M i s)}" by(auto) hence "(\s. Sup (range (\i. c * M i s))) = Sup_exp (range (\i s. c * M i s))" by (simp add: Sup_exp_def cong del: SUP_cong_simp) hence "a (\s. Sup (range (\i. c * M i s))) s = a (Sup_exp (range (\i s. c * M i s))) s" by(simp) } also { from le_funD[OF chain] nnc have "\i. (\s. c * M i s) \ (\s. c * M (Suc i) s)" by(auto intro:le_funI[OF mult_left_mono]) moreover from sM nnc have "\i. sound (\s. c * M i s)" by(auto intro:sound_intros) moreover from bM nnc have "\i. bounded_by (c * d) (\s. c * M i s)" by(auto intro:mult_left_mono) ultimately have "a (Sup_exp (range (\i s. c * M i s))) = Sup_exp (range (a o (\i s. c * M i s)))" by(rule bd_ctsD[OF ca]) hence "a (Sup_exp (range (\i s. c * M i s))) s = Sup_exp (range (a o (\i s. c * M i s))) s" by(auto) } also have "Sup_exp (range (a o (\i s. c * M i s))) s = Sup_exp (range (\x. a (\s. c * M x s))) s" by(simp add:o_def) also { from nnc sM have "\x. a (\s. c * M x s) = (\s. c * a (M x) s)" by(auto intro:scalingD[OF healthy_scalingD, OF ha, symmetric]) hence "Sup_exp (range (\x. a (\s. c * M x s))) s = Sup_exp (range (\x s. c * a (M x) s)) s" by(simp) } finally show "c * a (Sup_exp (range M)) s = Sup_exp (range (\x s. c * a (M x) s)) s" . qed lemma cts_wp_SetPC_const: fixes a::"'a \ 's prog" assumes ca: "\x. x \ (supp p) \ bd_cts (wp (a x))" and ha: "\x. x \ (supp p) \ healthy (wp (a x))" and up: "unitary p" and sump: "sum p (supp p) \ 1" and fsupp: "finite (supp p)" shows "bd_cts (wp (SetPC a (\_. p)))" proof(cases "supp p = {}", simp add:supp_empty SetPC_def wp_def cts_bot) assume nesupp: "supp p \ {}" from fsupp have "unitary p \ sum p (supp p) \ 1 \ (\x\supp p. bd_cts (wp (a x))) \ (\x\supp p. healthy (wp (a x))) \ bd_cts (wp (SetPC a (\_. p)))" proof(induct "supp p" arbitrary:p, simp add:supp_empty wp_SetPC_nil cts_bot, clarify) fix x::'a and F::"'a set" and p::"'a \ real" assume fF: "finite F" assume "insert x F = supp p" hence pstep: "supp p = insert x F" by(simp) hence xin: "x \ supp p" by(auto) assume up: "unitary p" and ca: "\x\supp p. bd_cts (wp (a x))" and ha: "\x\supp p. healthy (wp (a x))" and sump: "sum p (supp p) \ 1" and xni: "x \ F" assume IH: "\p. F = supp p \ unitary p \ sum p (supp p) \ 1 \ (\x\supp p. bd_cts (wp (a x))) \ (\x\supp p. healthy (wp (a x))) \ bd_cts (wp (SetPC a (\_. p)))" from fF pstep have fsupp: "finite (supp p)" by(auto) from xin have nzp: "p x \ 0" by(simp add:supp_def) have xy_le_sum: "\y. y \ supp p \ y \ x \ p x + p y \ sum p (supp p)" proof - fix y assume yin: "y \ supp p" and yne: "y \ x" from up have "0 \ sum p (supp p - {x,y})" by(auto intro:sum_nonneg) hence "p x + p y \ p x + p y + sum p (supp p - {x,y})" by(auto) also { from yin yne fsupp have "p y + sum p (supp p - {x,y}) = sum p (supp p - {x})" by(subst sum.insert[symmetric], (blast intro!:sum.cong)+) moreover from xin fsupp have "p x + sum p (supp p - {x}) = sum p (supp p)" by(subst sum.insert[symmetric], (blast intro!:sum.cong)+) ultimately have "p x + p y + sum p (supp p - {x, y}) = sum p (supp p)" by(simp) } finally show "p x + p y \ sum p (supp p)" . qed have n1p: "\y. y \ supp p \ y \ x \ p x \ 1" proof(rule ccontr, simp) assume px1: "p x = 1" fix y assume yin: "y \ supp p" and yne: "y \ x" from up have "0 \ p y" by(auto) with yin have "0 < p y" by(auto simp:supp_def) hence "0 + p x < p y + p x" by(rule add_strict_right_mono) with px1 have "1 < p x + p y" by(simp) also from yin yne have "p x + p y \ sum p (supp p)" by(rule xy_le_sum) finally show False using sump by(simp) qed show "bd_cts (wp (SetPC a (\_. p)))" proof(cases "F = {}") case True with pstep have "supp p = {x}" by(simp) hence "wp (SetPC a (\_. p)) = (\P s. p x * wp (a x) P s)" by(simp add:SetPC_sgl wp_def) moreover { from up ca ha xin have "bd_cts (wp (a x))" "healthy (wp (a x))" "0 \ p x" by(auto) hence "bd_cts (\P s. p x * wp (a x) P s)" by(rule bd_cts_scale) } ultimately show ?thesis by(simp) next assume neF: "F \ {}" then obtain y where yinF: "y \ F" by(auto) with xni have yne: "y \ x" by(auto) from yinF pstep have yin: "y \ supp p" by(auto) from supp_dist_remove[of p x, OF nzp n1p, OF yin yne] have supp_sub: "supp (dist_remove p x) \ supp p" by(auto) from xin ca have cax: "bd_cts (wp (a x))" by(auto) from xin ha have hax: "healthy (wp (a x))" by(auto) from supp_sub ha have hra: "\x\supp (dist_remove p x). healthy (wp (a x))" by(auto) from supp_sub ca have cra: "\x\supp (dist_remove p x). bd_cts (wp (a x))" by(auto) from supp_dist_remove[of p x, OF nzp n1p, OF yin yne] pstep xni have Fsupp: "F = supp (dist_remove p x)" by(simp) have udp: "unitary (dist_remove p x)" proof(intro unitaryI2 nnegI bounded_byI) fix y show "0 \ dist_remove p x y" proof(cases "y=x", simp_all add:dist_remove_def) - from up have "0 \ p y" "0 \ 1 - p x" by(auto simp:sign_simps) + from up have "0 \ p y" "0 \ 1 - p x" + by auto thus "0 \ p y / (1 - p x)" by(rule divide_nonneg_nonneg) qed show "dist_remove p x y \ 1" proof(cases "y=x", simp_all add:dist_remove_def, cases "y\supp p", simp_all add:nsupp_zero) assume yne: "y \ x" and yin: "y \ supp p" hence "p x + p y \ sum p (supp p)" by(auto intro:xy_le_sum) also note sump finally have "p y \ 1 - p x" by(auto) moreover from up have "p x \ 1" by(auto) moreover from yin yne have "p x \ 1" by(rule n1p) ultimately show "p y / (1 - p x) \ 1" by(auto) qed qed from xin have pxn0: "p x \ 0" by(auto simp:supp_def) from yin yne have pxn1: "p x \ 1" by(rule n1p) from pxn0 pxn1 have "sum (dist_remove p x) (supp (dist_remove p x)) = sum (dist_remove p x) (supp p - {x})" by(simp add:supp_dist_remove) also have "... = (\y\supp p - {x}. p y / (1 - p x))" by(simp add:dist_remove_def) also have "... = (\y\supp p - {x}. p y) / (1 - p x)" by(simp add:sum_divide_distrib) also { from xin have "insert x (supp p) = supp p" by(auto) with fsupp have "p x + (\y\supp p - {x}. p y) = sum p (supp p)" by(simp add:sum.insert[symmetric]) also note sump finally have "sum p (supp p - {x}) \ 1 - p x" by(auto) moreover { from up have "p x \ 1" by(auto) with pxn1 have "p x < 1" by(auto) hence "0 < 1 - p x" by(auto) } ultimately have "sum p (supp p - {x}) / (1 - p x) \ 1" by(auto) } finally have sdp: "sum (dist_remove p x) (supp (dist_remove p x)) \ 1" . from Fsupp udp sdp hra cra IH have cts_dr: "bd_cts (wp (SetPC a (\_. dist_remove p x)))" by(auto) from up have upx: "unitary (\_. p x)" by(auto) from pxn0 pxn1 fsupp hra show ?thesis by(simp add:SetPC_remove, blast intro:cts_wp_PC cax cts_dr hax healthy_intros unitary_sound[OF udp] sdp upx) qed qed with assms show ?thesis by(auto) qed lemma cts_wp_SetPC: fixes a::"'a \ 's prog" assumes ca: "\x s. x \ (supp (p s)) \ bd_cts (wp (a x))" and ha: "\x s. x \ (supp (p s)) \ healthy (wp (a x))" and up: "\s. unitary (p s)" and sump: "\s. sum (p s) (supp (p s)) \ 1" and fsupp: "\s. finite (supp (p s))" shows "bd_cts (wp (SetPC a p))" proof - from assms have "bd_cts (wp (Bind p (\p. SetPC a (\_. p))))" by(iprover intro!:cts_wp_Bind cts_wp_SetPC_const) thus ?thesis by(simp add:SetPC_Bind[symmetric]) qed lemma wp_SetDC_Bind: "SetDC a S = Bind S (\S. SetDC a (\_. S))" by(intro ext, simp add:SetDC_def Bind_def) lemma SetDC_finite_insert: assumes fS: "finite S" and neS: "S \ {}" shows "SetDC a (\_. insert x S) = a x \ SetDC a (\_. S)" proof (intro ext, simp add: SetDC_def DC_def cong del: image_cong_simp cong add: INF_cong_simp) fix ab P s from fS have A: "finite (insert (a x ab P s) ((\x. a x ab P s) ` S))" and B: "finite (((\x. a x ab P s) ` S))" by(auto) from neS have C: "insert (a x ab P s) ((\x. a x ab P s) ` S) \ {}" and D: "(\x. a x ab P s) ` S \ {}" by(auto) from A C have "Inf (insert (a x ab P s) ((\x. a x ab P s) ` S)) = Min (insert (a x ab P s) ((\x. a x ab P s) ` S))" by(auto intro:cInf_eq_Min) also from B D have "... = min (a x ab P s) (Min ((\x. a x ab P s) ` S))" by(auto intro:Min_insert) also from B D have "... = min (a x ab P s) (Inf ((\x. a x ab P s) ` S))" by(simp add:cInf_eq_Min) finally show "(INF x\insert x S. a x ab P s) = min (a x ab P s) (INF x\S. a x ab P s)" by (simp cong del: INF_cong_simp) qed lemma SetDC_singleton: "SetDC a (\_. {x}) = a x" by (simp add: SetDC_def cong del: INF_cong_simp) lemma cts_wp_SetDC_const: fixes a::"'a \ 's prog" assumes ca: "\x. x \ S \ bd_cts (wp (a x))" and ha: "\x. x \ S \ healthy (wp (a x))" and fS: "finite S" and neS: "S \ {}" shows "bd_cts (wp (SetDC a (\_. S)))" proof - have "finite S \ S \ {} \ (\x\S. bd_cts (wp (a x))) \ (\x\S. healthy (wp (a x))) \ bd_cts (wp (SetDC a (\_. S)))" proof(induct S rule:finite_induct, simp, clarsimp) fix x::'a and F::"'a set" assume fF: "finite F" and IH: "F \ {} \ bd_cts (wp (SetDC a (\_. F)))" and cax: "bd_cts (wp (a x))" and hax: "healthy (wp (a x))" and haF: "\x\F. healthy (wp (a x))" show "bd_cts (wp (SetDC a (\_. insert x F)))" proof(cases "F = {}", simp add:SetDC_singleton cax) assume "F \ {}" with fF cax hax haF IH show "bd_cts (wp (SetDC a (\_. insert x F)))" by(auto intro!:cts_wp_DC healthy_intros simp:SetDC_finite_insert) qed qed with assms show ?thesis by(auto) qed lemma cts_wp_SetDC: fixes a::"'a \ 's prog" assumes ca: "\x s. x \ S s \ bd_cts (wp (a x))" and ha: "\x s. x \ S s \ healthy (wp (a x))" and fS: "\s. finite (S s)" and neS: "\s. S s \ {}" shows "bd_cts (wp (SetDC a S))" proof - from assms have "bd_cts (wp (Bind S (\S. SetDC a (\_. S))))" by(iprover intro!:cts_wp_Bind cts_wp_SetDC_const) thus ?thesis by(simp add:wp_SetDC_Bind[symmetric]) qed lemma cts_wp_repeat: "bd_cts (wp a) \ healthy (wp a) \ bd_cts (wp (repeat n a))" by(induct n, auto intro:cts_wp_Skip cts_wp_Seq healthy_intros) lemma cts_wp_Embed: "bd_cts t \ bd_cts (wp (Embed t))" by(simp add:wp_eval) subsection \Continuity of a Single Loop Step\ text \A single loop iteration is continuous, in the more general sense defined above for transformer transformers.\ lemma cts_wp_loopstep: fixes body::"'s prog" assumes hb: "healthy (wp body)" and cb: "bd_cts (wp body)" shows "bd_cts_tr (\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip))" (is "bd_cts_tr ?F") proof(rule bd_cts_trI, rule le_trans_antisym) fix M::"nat \ 's trans" and b::real assume chain: "\i. le_trans (M i) (M (Suc i))" and fM: "\i. feasible (M i)" show fw: "le_trans (Sup_trans (range (?F o M))) (?F (Sup_trans (range M)))" proof(rule le_transI[OF Sup_trans_least2], clarsimp) fix P Q::"'s expect" and t assume sP: "sound P" assume nQ: "nneg Q" and bP: "bounded_by (bound_of P) Q" hence sQ: "sound Q" by(auto) from fM have fSup: "feasible (Sup_trans (range M))" by(auto intro:feasible_Sup_trans) from sQ fM have "M t Q \ Sup_trans (range M) Q" by(auto intro:Sup_trans_upper2) moreover from sQ fM fSup have sMtP: "sound (M t Q)" "sound (Sup_trans (range M) Q)" by(auto) ultimately have "wp body (M t Q) \ wp body (Sup_trans (range M) Q)" using healthy_monoD[OF hb] by(auto) hence "\s. wp body (M t Q) s \ wp body (Sup_trans (range M) Q) s" by(rule le_funD) thus "?F (M t) Q \ ?F (Sup_trans (range M)) Q" by(intro le_funI, simp add:wp_eval mult_left_mono) show "nneg (wp (body ;; Embed (Sup_trans (range M)) \<^bsub>\ G \\<^esub>\ Skip) Q)" proof(rule nnegI, simp add:wp_eval) fix s::'s from fSup sQ have "sound (Sup_trans (range M) Q)" by(auto) with hb have "sound (wp body (Sup_trans (range M) Q))" by(auto) hence "0 \ wp body (Sup_trans (range M) Q) s" by(auto) moreover from sQ have "0 \ Q s" by(auto) ultimately show "0 \ \G\ s * wp body (Sup_trans (range M) Q) s + (1 - \G\ s) * Q s" by(auto intro:add_nonneg_nonneg mult_nonneg_nonneg) qed next fix P::"'s expect" assume sP: "sound P" thus "nneg P" "bounded_by (bound_of P) P" by(auto) show "\u\range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M). \R. nneg R \ bounded_by (bound_of P) R \ nneg (u R) \ bounded_by (bound_of P) (u R)" proof(clarsimp, intro conjI nnegI bounded_byI, simp_all add:wp_eval) fix u::nat and R::"'s expect" and s::'s assume nR: "nneg R" and bR: "bounded_by (bound_of P) R" hence sR: "sound R" by(auto) with fM have sMuR: "sound (M u R)" by(auto) with hb have "sound (wp body (M u R))" by(auto) hence "0 \ wp body (M u R) s" by(auto) moreover from nR have "0 \ R s" by(auto) ultimately show "0 \ \G\ s * wp body (M u R) s + (1 - \G\ s) * R s" by(auto intro:add_nonneg_nonneg mult_nonneg_nonneg) from sR bR fM have "bounded_by (bound_of P) (M u R)" by(auto) with sMuR hb have "bounded_by (bound_of P) (wp body (M u R))" by(auto) hence "wp body (M u R) s \ bound_of P" by(auto) moreover from bR have "R s \ bound_of P" by(auto) ultimately have "\G\ s * wp body (M u R) s + (1 - \G\ s) * R s \ \G\ s * bound_of P + (1 - \G\ s) * bound_of P" by(auto intro:add_mono mult_left_mono) also have "... = bound_of P" by(simp add:algebra_simps) finally show "\G\ s * wp body (M u R) s + (1 - \G\ s) * R s \ bound_of P" . qed qed show "le_trans (?F (Sup_trans (range M))) (Sup_trans (range (?F o M)))" proof(rule le_transI, rule le_funI, simp add: wp_eval cong del: image_cong_simp) fix P::"'s expect" and s::'s assume sP: "sound P" have "{t P |t. t \ range M} = range (\i. M i P)" by(blast) hence "wp body (Sup_trans (range M) P) s = wp body (Sup_exp (range (\i. M i P))) s" by(simp add:Sup_trans_def) also { from sP fM have "\i. sound (M i P)" by(auto) moreover from sP chain have "\i. M i P \ M (Suc i) P" by(auto) moreover { from sP have "bounded_by (bound_of P) P" by(auto) with sP fM have "\i. bounded_by (bound_of P) (M i P)" by(auto) } ultimately have "wp body (Sup_exp (range (\i. M i P))) s = Sup_exp (range (\i. wp body (M i P))) s" by(subst bd_ctsD[OF cb], auto simp:o_def) } also have "Sup_exp (range (\i. wp body (M i P))) s = Sup {f s |f. f \ range (\i. wp body (M i P))}" by(simp add:Sup_exp_def) finally have "\G\ s * wp body (Sup_trans (range M) P) s + (1 - \G\ s) * P s = \G\ s * Sup {f s |f. f \ range (\i. wp body (M i P))} + (1 - \G\ s) * P s" by(simp) also { from sP fM have "\i. sound (M i P)" by(auto) moreover from sP fM have "\i. bounded_by (bound_of P) (M i P)" by(auto) ultimately have "\i. bounded_by (bound_of P) (wp body (M i P))" using hb by(auto) hence bound: "\i. wp body (M i P) s \ bound_of P" by(auto) moreover have "{\ G \ s * x |x. x \ {f s |f. f \ range (\i. wp body (M i P))}} = {\ G \ s * f s |f. f \ range (\i. wp body (M i P))}" by(blast) ultimately have "\G\ s * Sup {f s |f. f \ range (\i. wp body (M i P))} = Sup {\G\ s * f s |f. f \ range (\i. wp body (M i P))}" by(subst cSup_mult, auto) moreover { have "{x + (1-\G\ s) * P s |x. x \ {\G\ s * f s |f. f \ range (\i. wp body (M i P))}} = {\G\ s * f s + (1-\G\ s) * P s |f. f \ range (\i. wp body (M i P))}" by(blast) moreover from bound sP have "\i. \G\ s * wp body (M i P) s \ bound_of P" by(cases "G s", auto) ultimately have "Sup {\G\ s * f s |f. f \ range (\i. wp body (M i P))} + (1-\G\ s) * P s = Sup {\G\ s * f s + (1-\G\ s) * P s |f. f \ range (\i. wp body (M i P))}" by(subst cSup_add, auto) } ultimately have "\G\ s * Sup {f s |f. f \ range (\i. wp body (M i P))} + (1-\G\ s) * P s = Sup {\G\ s * f s + (1-\G\ s) * P s |f. f \ range (\i. wp body (M i P))}" by(simp) } also { have "\i. \G\ s * wp body (M i P) s + (1-\G\ s) * P s = ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M) i P s" by(simp add:wp_eval) also have "\i. ... i \ Sup {f s |f. f \ {t P |t. t \ range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M)}}" proof(intro cSup_upper bdd_aboveI, blast, clarsimp simp:wp_eval) fix i from sP have bP: "bounded_by (bound_of P) P" by(auto) with sP fM have "sound (M i P)" "bounded_by (bound_of P) (M i P)" by(auto) with hb have "bounded_by (bound_of P) (wp body (M i P))" by(auto) with bP have "wp body (M i P) s \ bound_of P" "P s \ bound_of P" by(auto) hence "\G\ s * wp body (M i P) s + (1-\G\ s) * P s \ \G\ s * (bound_of P) + (1-\G\ s) * (bound_of P)" by(auto intro:add_mono mult_left_mono) also have "... = bound_of P" by(simp add:algebra_simps) finally show "\G\ s * wp body (M i P) s + (1-\G\ s) * P s \ bound_of P" . qed finally have "Sup {\G\ s * f s + (1-\G\ s) * P s |f. f \ range (\i. wp body (M i P))} \ Sup {f s |f. f \ {t P |t. t \ range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M)}}" by(blast intro:cSup_least) } also have "Sup {f s |f. f \ {t P |t. t \ range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M)}} = Sup_trans (range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M)) P s" by(simp add:Sup_trans_def Sup_exp_def) finally show "\G\ s * wp body (Sup_trans (range M) P) s + (1-\G\ s) * P s \ Sup_trans (range ((\x. wp (body ;; Embed x \<^bsub>\ G \\<^esub>\ Skip)) \ M)) P s" . qed qed end diff --git a/thys/pGCL/Healthiness.thy b/thys/pGCL/Healthiness.thy --- a/thys/pGCL/Healthiness.thy +++ b/thys/pGCL/Healthiness.thy @@ -1,1122 +1,1126 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section \Healthiness\ theory Healthiness imports Embedding begin subsection \The Healthiness of the Embedding\ text \Healthiness is mostly derived by structural induction using the simplifier. @{term Abort}, @{term Skip} and @{term Apply} form base cases.\ lemma healthy_wp_Abort: "healthy (wp Abort)" proof(rule healthy_parts) fix b and P::"'a \ real" assume nP: "nneg P" and bP: "bounded_by b P" thus "bounded_by b (wp Abort P)" unfolding wp_eval by(blast) show "nneg (wp Abort P)" unfolding wp_eval by(blast) next fix P Q::"'a expect" show "wp Abort P \ wp Abort Q" unfolding wp_eval by(blast) next fix P and c and s::'a show "c * wp Abort P s = wp Abort (\s. c * P s) s" unfolding wp_eval by(auto) qed lemma nearly_healthy_wlp_Abort: "nearly_healthy (wlp Abort)" proof(rule nearly_healthyI) fix P::"'s \ real" show "unitary (wlp Abort P)" by(simp add:wp_eval) next fix P Q :: "'s expect" assume "P \ Q" and "unitary P" and "unitary Q" thus "wlp Abort P \ wlp Abort Q" unfolding wp_eval by(blast) qed lemma healthy_wp_Skip: "healthy (wp Skip)" by(force intro!:healthy_parts simp:wp_eval) lemma nearly_healthy_wlp_Skip: "nearly_healthy (wlp Skip)" by(auto simp:wp_eval) lemma healthy_wp_Seq: fixes t::"'s prog" and u assumes ht: "healthy (wp t)" and hu: "healthy (wp u)" shows "healthy (wp (t ;; u))" proof(rule healthy_parts, simp_all add:wp_eval) fix b and P::"'s \ real" assume "bounded_by b P" and "nneg P" with hu have "bounded_by b (wp u P)" and "nneg (wp u P)" by(auto) with ht show "bounded_by b (wp t (wp u P))" and "nneg (wp t (wp u P))" by(auto) next fix P::"'s \ real" and Q assume "sound P" and "sound Q" and "P \ Q" with hu have "sound (wp u P)" and "sound (wp u Q)" and "wp u P \ wp u Q" by(auto) with ht show "wp t (wp u P) \ wp t (wp u Q)" by(auto) next fix P::"'s \ real" and c::real and s assume pos: "0 \ c" and sP: "sound P" with ht and hu have "c * wp t (wp u P) s = wp t (\s. c * wp u P s) s" by(auto intro!:scalingD) also with hu and pos and sP have "... = wp t (wp u (\s. c * P s)) s" by(simp add:scalingD[OF healthy_scalingD]) finally show "c * wp t (wp u P) s = wp t (wp u (\s. c * P s)) s" . qed lemma nearly_healthy_wlp_Seq: fixes t::"'s prog" and u assumes ht: "nearly_healthy (wlp t)" and hu: "nearly_healthy (wlp u)" shows "nearly_healthy (wlp (t ;; u))" proof(rule nearly_healthyI, simp_all add:wp_eval) fix b and P::"'s \ real" assume "unitary P" with hu have "unitary (wlp u P)" by(auto) with ht show "unitary (wlp t (wlp u P))" by(auto) next fix P Q::"'s \ real" assume "unitary P" and "unitary Q" and "P \ Q" with hu have "unitary (wlp u P)" and "unitary (wlp u Q)" and "wlp u P \ wlp u Q" by(auto) with ht show "wlp t (wlp u P) \ wlp t (wlp u Q)" by(auto) qed lemma healthy_wp_PC: fixes f::"'s prog" assumes hf: "healthy (wp f)" and hg: "healthy (wp g)" and uP: "unitary P" shows "healthy (wp (f \<^bsub>P\<^esub>\ g))" proof(intro healthy_parts bounded_byI nnegI le_funI, simp_all add:wp_eval) fix b and Q::"'s \ real" and s::'s assume nQ: "nneg Q" and bQ: "bounded_by b Q" txt \Non-negative:\ from nQ and bQ and hf have "0 \ wp f Q s" by(auto) with uP have "0 \ P s * ..." by(auto intro:mult_nonneg_nonneg) moreover { - from uP have "0 \ 1 - P s" by(auto simp:sign_simps) + from uP have "0 \ 1 - P s" + by auto with nQ and bQ and hg have "0 \ ... * wp g Q s" by (metis healthy_nnegD2 mult_nonneg_nonneg nneg_def) } ultimately show "0 \ P s * wp f Q s + (1 - P s) * wp g Q s" by(auto intro:mult_nonneg_nonneg) txt \Bounded:\ from nQ bQ hf have "wp f Q s \ b" by(auto) with uP nQ bQ hf have "P s * wp f Q s \ P s * b" by(blast intro!:mult_mono) moreover { from nQ bQ hg uP - have "wp g Q s \ b" and "0 \ 1 - P s" by(auto simp:sign_simps) + have "wp g Q s \ b" and "0 \ 1 - P s" + by auto with nQ bQ hg have "(1 - P s) * wp g Q s \ (1 - P s) * b" by(blast intro!:mult_mono) } ultimately have "P s * wp f Q s + (1 - P s) * wp g Q s \ P s * b + (1 - P s) * b" by(blast intro:add_mono) also have "... = b" by(auto simp:algebra_simps) finally show "P s * wp f Q s + (1 - P s) * wp g Q s \ b" . next txt \Monotonic:\ fix Q R::"'s \ real" and s assume sQ: "sound Q" and sR: "sound R" and le: "Q \ R" with hf have "wp f Q s \ wp f R s" by(blast dest:mono_transD) with uP have "P s * wp f Q s \ P s * wp f R s" by(auto intro:mult_left_mono) moreover { from sQ sR le hg have "wp g Q s \ wp g R s" by(blast dest:mono_transD) - moreover from uP have "0 \ 1 - P s" by(auto simp:sign_simps) + moreover from uP have "0 \ 1 - P s" + by auto ultimately have "(1 - P s) * wp g Q s \ (1 - P s) * wp g R s" by(auto intro:mult_left_mono) } ultimately show "P s * wp f Q s + (1 - P s) * wp g Q s \ P s * wp f R s + (1 - P s) * wp g R s" by(auto) next txt \Scaling:\ fix Q::"'s \ real" and c::real and s::'s assume sQ: "sound Q" and pos: "0 \ c" have "c * (P s * wp f Q s + (1 - P s) * wp g Q s) = P s * (c * wp f Q s) + (1 - P s) * (c * wp g Q s)" by(simp add:distrib_left) also have "... = P s * wp f (\s. c * Q s) s + (1 - P s) * wp g (\s. c * Q s) s" using hf hg sQ pos by(simp add:scalingD[OF healthy_scalingD]) finally show "c * (P s * wp f Q s + (1 - P s) * wp g Q s) = P s * wp f (\s. c * Q s) s + (1 - P s) * wp g (\s. c * Q s) s" . qed lemma nearly_healthy_wlp_PC: fixes f::"'s prog" assumes hf: "nearly_healthy (wlp f)" and hg: "nearly_healthy (wlp g)" and uP: "unitary P" shows "nearly_healthy (wlp (f \<^bsub>P\<^esub>\ g))" proof(intro nearly_healthyI unitaryI2 nnegI bounded_byI le_funI, simp_all add:wp_eval) fix Q::"'s expect" and s::'s assume uQ: "unitary Q" from uQ hf hg have utQ: "unitary (wlp f Q)" "unitary (wlp g Q)" by(auto) - from uP have nnP: "0 \ P s" "0 \ 1 - P s" by(auto simp:sign_simps) + from uP have nnP: "0 \ P s" "0 \ 1 - P s" + by auto moreover from utQ have "0 \ wlp f Q s" "0 \ wlp g Q s" by(auto) ultimately show "0 \ P s * wlp f Q s + (1 - P s) * wlp g Q s" by(auto intro:add_nonneg_nonneg mult_nonneg_nonneg) from utQ have "wlp f Q s \ 1" "wlp g Q s \ 1" by(auto) with nnP have "P s * wlp f Q s + (1 - P s) * wlp g Q s \ P s * 1 + (1 - P s) * 1" by(blast intro:add_mono mult_left_mono) thus "P s * wlp f Q s + (1 - P s) * wlp g Q s \ 1" by(simp) fix R::"'s expect" assume uR: "unitary R" and le: "Q \ R" with uQ have "wlp f Q s \ wlp f R s" by(auto intro:le_funD[OF nearly_healthy_monoD, OF hf]) with nnP have "P s * wlp f Q s \ P s * wlp f R s" by(auto intro:mult_left_mono) moreover { from uQ uR le have "wlp g Q s \ wlp g R s" by(auto intro:le_funD[OF nearly_healthy_monoD, OF hg]) with nnP have "(1 - P s) * wlp g Q s \ (1 - P s) * wlp g R s" by(auto intro:mult_left_mono) } ultimately show "P s * wlp f Q s + (1 - P s) * wlp g Q s \ P s * wlp f R s + (1 - P s) * wlp g R s" by(auto) qed lemma healthy_wp_DC: fixes f::"'s prog" assumes hf: "healthy (wp f)" and hg: "healthy (wp g)" shows "healthy (wp (f \ g))" proof(intro healthy_parts bounded_byI nnegI le_funI, simp_all only:wp_eval) fix b and P::"'s \ real" and s::'s assume nP: "nneg P" and bP: "bounded_by b P" with hf have "bounded_by b (wp f P)" by(auto) hence "wp f P s \ b" by(blast) thus "min (wp f P s) (wp g P s) \ b" by(auto) from nP bP assms show "0 \ min (wp f P s) (wp g P s)" by(auto) next fix P::"'s \ real" and Q and s::'s from assms have mf: "mono_trans (wp f)" and mg: "mono_trans (wp g)" by(auto) assume sP: "sound P" and sQ: "sound Q" and le: "P \ Q" hence "wp f P s \ wp f Q s" and "wp g P s \ wp g Q s" by(auto intro:le_funD[OF mono_transD[OF mf]] le_funD[OF mono_transD[OF mg]]) thus "min (wp f P s) (wp g P s) \ min (wp f Q s) (wp g Q s)" by(auto) next fix P::"'s \ real" and c::real and s::'s assume sP: "sound P" and pos: "0 \ c" from assms have sf: "scaling (wp f)" and sg: "scaling (wp g)" by(auto) from pos have "c * min (wp f P s) (wp g P s) = min (c * wp f P s) (c * wp g P s)" by(simp add:min_distrib) also from sP and pos have "... = min (wp f (\s. c * P s) s) (wp g (\s. c * P s) s)" by(simp add:scalingD[OF sf] scalingD[OF sg]) finally show "c * min (wp f P s) (wp g P s) = min (wp f (\s. c * P s) s) (wp g (\s. c * P s) s)" . qed lemma nearly_healthy_wlp_DC: fixes f::"'s prog" assumes hf: "nearly_healthy (wlp f)" and hg: "nearly_healthy (wlp g)" shows "nearly_healthy (wlp (f \ g))" proof(intro nearly_healthyI bounded_byI nnegI le_funI unitaryI2, simp_all add:wp_eval, safe) fix P::"'s \ real" and s::'s assume uP: "unitary P" with hf hg have utP: "unitary (wlp f P)" "unitary (wlp g P)" by(auto) thus "0 \ wlp f P s" "0 \ wlp g P s" by(auto) have "min (wlp f P s) (wlp g P s) \ wlp f P s" by(auto) also from utP have "... \ 1" by(auto) finally show "min (wlp f P s) (wlp g P s) \ 1" . fix Q::"'s \ real" assume uQ: "unitary Q" and le: "P \ Q" have "min (wlp f P s) (wlp g P s) \ wlp f P s" by(auto) also from uP uQ le have "... \ wlp f Q s" by(auto intro:le_funD[OF nearly_healthy_monoD, OF hf]) finally show "min (wlp f P s) (wlp g P s) \ wlp f Q s" . have "min (wlp f P s) (wlp g P s) \ wlp g P s" by(auto) also from uP uQ le have "... \ wlp g Q s" by(auto intro:le_funD[OF nearly_healthy_monoD, OF hg]) finally show "min (wlp f P s) (wlp g P s) \ wlp g Q s" . qed lemma healthy_wp_AC: fixes f::"'s prog" assumes hf: "healthy (wp f)" and hg: "healthy (wp g)" shows "healthy (wp (f \ g))" proof(intro healthy_parts bounded_byI nnegI le_funI, simp_all only:wp_eval) fix b and P::"'s \ real" and s::'s assume nP: "nneg P" and bP: "bounded_by b P" with hf have "bounded_by b (wp f P)" by(auto) hence "wp f P s \ b" by(blast) moreover { from bP nP hg have "bounded_by b (wp g P)" by(auto) hence "wp g P s \ b" by(blast) } ultimately show "max (wp f P s) (wp g P s) \ b" by(auto) from nP bP assms have "0 \ wp f P s" by(auto) thus "0 \ max (wp f P s) (wp g P s)" by(auto) next fix P::"'s \ real" and Q and s::'s from assms have mf: "mono_trans (wp f)" and mg: "mono_trans (wp g)" by(auto) assume sP: "sound P" and sQ: "sound Q" and le: "P \ Q" hence "wp f P s \ wp f Q s" and "wp g P s \ wp g Q s" by(auto intro:le_funD[OF mono_transD, OF mf] le_funD[OF mono_transD, OF mg]) thus "max (wp f P s) (wp g P s) \ max (wp f Q s) (wp g Q s)" by(auto) next fix P::"'s \ real" and c::real and s::'s assume sP: "sound P" and pos: "0 \ c" from assms have sf: "scaling (wp f)" and sg: "scaling (wp g)" by(auto) from pos have "c * max (wp f P s) (wp g P s) = max (c * wp f P s) (c * wp g P s)" by(simp add:max_distrib) also from sP and pos have "... = max (wp f (\s. c * P s) s) (wp g (\s. c * P s) s)" by(simp add:scalingD[OF sf] scalingD[OF sg]) finally show "c * max (wp f P s) (wp g P s) = max (wp f (\s. c * P s) s) (wp g (\s. c * P s) s)" . qed lemma nearly_healthy_wlp_AC: fixes f::"'s prog" assumes hf: "nearly_healthy (wlp f)" and hg: "nearly_healthy (wlp g)" shows "nearly_healthy (wlp (f \ g))" proof(intro nearly_healthyI bounded_byI nnegI unitaryI2 le_funI, simp_all only:wp_eval) fix b and P::"'s \ real" and s::'s assume uP: "unitary P" with hf have "wlp f P s \ 1" by(auto) moreover from uP hg have "unitary (wlp g P)" by(auto) hence "wlp g P s \ 1" by(auto) ultimately show "max (wlp f P s) (wlp g P s) \ 1" by(auto) from uP hf have "unitary (wlp f P)" by(auto) hence "0 \ wlp f P s" by(auto) thus "0 \ max (wlp f P s) (wlp g P s)" by(auto) next fix P::"'s \ real" and Q and s::'s assume uP: "unitary P" and uQ: "unitary Q" and le: "P \ Q" hence "wlp f P s \ wlp f Q s" and "wlp g P s \ wlp g Q s" by(auto intro:le_funD[OF nearly_healthy_monoD, OF hf] le_funD[OF nearly_healthy_monoD, OF hg]) thus "max (wlp f P s) (wlp g P s) \ max (wlp f Q s) (wlp g Q s)" by(auto) qed lemma healthy_wp_Embed: "healthy t \ healthy (wp (Embed t))" unfolding wp_def Embed_def by(simp) lemma nearly_healthy_wlp_Embed: "nearly_healthy t \ nearly_healthy (wlp (Embed t))" unfolding wlp_def Embed_def by(simp) lemma healthy_wp_repeat: assumes h_a: "healthy (wp a)" shows "healthy (wp (repeat n a))" (is "?X n") proof(induct n) show "?X 0" by(auto simp:wp_eval) next fix n assume IH: "?X n" thus "?X (Suc n)" by(simp add:healthy_wp_Seq h_a) qed lemma nearly_healthy_wlp_repeat: assumes h_a: "nearly_healthy (wlp a)" shows "nearly_healthy (wlp (repeat n a))" (is "?X n") proof(induct n) show "?X 0" by(simp add:wp_eval) next fix n assume IH: "?X n" thus "?X (Suc n)" by(simp add:nearly_healthy_wlp_Seq h_a) qed lemma healthy_wp_SetDC: fixes prog::"'b \ 'a prog" and S::"'a \ 'b set" assumes healthy: "\x s. x \ S s \ healthy (wp (prog x))" and nonempty: "\s. \x. x \ S s" shows "healthy (wp (SetDC prog S))" (is "healthy ?T") proof(intro healthy_parts bounded_byI nnegI le_funI, simp_all only:wp_eval) fix b and P::"'a \ real" and s::'a assume bP: "bounded_by b P" and nP: "nneg P" hence sP: "sound P" by(auto) from nonempty obtain x where xin: "x \ (\a. wp (prog a) P s) ` S s" by(blast) moreover from sP and healthy have "\x\(\a. wp (prog a) P s) ` S s. 0 \ x" by(auto) ultimately have "Inf ((\a. wp (prog a) P s) ` S s) \ x" by(intro cInf_lower bdd_belowI, auto) also from xin and healthy and sP and bP have "x \ b" by(blast) finally show "Inf ((\a. wp (prog a) P s) ` S s) \ b" . from xin and sP and healthy show "0 \ Inf ((\a. wp (prog a) P s) ` S s)" by(blast intro:cInf_greatest) next fix P::"'a \ real" and Q and s::'a assume sP: "sound P" and sQ: "sound Q" and le: "P \ Q" from nonempty obtain x where xin: "x \ (\a. wp (prog a) P s) ` S s" by(blast) moreover from sP and healthy have "\x\(\a. wp (prog a) P s) ` S s. 0 \ x" by(auto) moreover have "\x\(\a. wp (prog a) Q s) ` S s. \y\(\a. wp (prog a) P s) ` S s. y \ x" proof(rule ballI, clarify, rule bexI) fix x and a assume ain: "a \ S s" with healthy and sP and sQ and le show "wp (prog a) P s \ wp (prog a) Q s" by(auto dest:mono_transD[OF healthy_monoD]) from ain show "wp (prog a) P s \ (\a. wp (prog a) P s) ` S s" by(simp) qed ultimately show "Inf ((\a. wp (prog a) P s) ` S s) \ Inf ((\a. wp (prog a) Q s) ` S s)" by(intro cInf_mono, blast+) next fix P::"'a \ real" and c::real and s::'a assume sP: "sound P" and pos: "0 \ c" from nonempty obtain x where xin: "x \ (\a. wp (prog a) P s) ` S s" by(blast) have "c * Inf ((\a. wp (prog a) P s) ` S s) = Inf ((*) c ` ((\a. wp (prog a) P s) ` S s))" (is "?U = ?V") proof(rule antisym) show "?U \ ?V" proof(rule cInf_greatest) from nonempty show "(*) c ` (\a. wp (prog a) P s) ` S s \ {}" by(auto) fix x assume "x \ (*) c ` (\a. wp (prog a) P s) ` S s" then obtain y where yin: "y \ (\a. wp (prog a) P s) ` S s" and rwx: "x = c * y" by(auto) have "Inf ((\a. wp (prog a) P s) ` S s) \ y" proof(intro cInf_lower[OF yin] bdd_belowI) fix z assume zin: "z \ (\a. wp (prog a) P s) ` S s" then obtain a where "a \ S s" and "z = wp (prog a) P s" by(auto) with sP show "0 \ z" by(auto dest:healthy) qed with pos rwx show "c * Inf ((\a. wp (prog a) P s) ` S s) \ x" by(auto intro:mult_left_mono) qed show "?V \ ?U" proof(cases) assume cz: "c = 0" moreover { from nonempty obtain c where "c \ S s" by(auto) hence "\x. \xa\S s. x = wp (prog xa) P s" by(auto) } ultimately show ?thesis by(simp add:image_def) next assume "c \ 0" from nonempty have "S s \ {}" by blast then have "inverse c * (INF x\S s. c * wp (prog x) P s) \ (INF a\S s. wp (prog a) P s)" proof (rule cINF_greatest) fix x assume "x \ S s" have "bdd_below ((\x. c * wp (prog x) P s) ` S s)" proof (rule bdd_belowI [of _ 0]) fix z assume "z \ (\x. c * wp (prog x) P s) ` S s" then obtain b where "b \ S s" and rwz: "z = c * wp (prog b) P s" by auto with sP have "0 \ wp (prog b) P s" by (auto dest: healthy) with pos show "0 \ z" by (auto simp: rwz intro: mult_nonneg_nonneg) qed then have "(INF x\S s. c * wp (prog x) P s) \ c * wp (prog x) P s" using \x \ S s\ by (rule cINF_lower) with \c \ 0\ show "inverse c * (INF x\S s. c * wp (prog x) P s) \ wp (prog x) P s" by (simp add: mult_div_mono_left pos) qed with \c \ 0\ have "inverse c * ?V \ inverse c * ?U" by (simp add: mult.assoc [symmetric] image_comp) with pos have "c * (inverse c * ?V) \ c * (inverse c * ?U)" by(auto intro:mult_left_mono) with \c \ 0\ show ?thesis by (simp add:mult.assoc [symmetric]) qed qed also have "... = Inf ((\a. c * wp (prog a) P s) ` S s)" by (simp add: image_comp) also from sP and pos have "... = Inf ((\a. wp (prog a) (\s. c * P s) s) ` S s)" by(simp add:scalingD[OF healthy_scalingD, OF healthy] cong:image_cong) finally show "c * Inf ((\a. wp (prog a) P s) ` S s) = Inf ((\a. wp (prog a) (\s. c * P s) s) ` S s)" . qed lemma nearly_healthy_wlp_SetDC: fixes prog::"'b \ 'a prog" and S::"'a \ 'b set" assumes healthy: "\x s. x \ S s \ nearly_healthy (wlp (prog x))" and nonempty: "\s. \x. x \ S s" shows "nearly_healthy (wlp (SetDC prog S))" (is "nearly_healthy ?T") proof(intro nearly_healthyI unitaryI2 bounded_byI nnegI le_funI, simp_all only:wp_eval) fix b and P::"'a \ real" and s::'a assume uP: "unitary P" from nonempty obtain x where xin: "x \ (\a. wlp (prog a) P s) ` S s" by(blast) moreover { from uP healthy have "\x\(\a. wlp (prog a) P) ` S s. unitary x" by(auto) hence "\x\(\a. wlp (prog a) P) ` S s. 0 \ x s" by(auto) hence "\y\(\a. wlp (prog a) P s) ` S s. 0 \ y" by(auto) } ultimately have "Inf ((\a. wlp (prog a) P s) ` S s) \ x" by(intro cInf_lower bdd_belowI, auto) also from xin healthy uP have "x \ 1" by(blast) finally show "Inf ((\a. wlp (prog a) P s) ` S s) \ 1" . from xin uP healthy show "0 \ Inf ((\a. wlp (prog a) P s) ` S s)" by(blast dest!:unitary_sound[OF nearly_healthy_unitaryD[OF _ uP]] intro:cInf_greatest) next fix P::"'a \ real" and Q and s::'a assume uP: "unitary P" and uQ: "unitary Q" and le: "P \ Q" from nonempty obtain x where xin: "x \ (\a. wlp (prog a) P s) ` S s" by(blast) moreover { from uP healthy have "\x\(\a. wlp (prog a) P) ` S s. unitary x" by(auto) hence "\x\(\a. wlp (prog a) P) ` S s. 0 \ x s" by(auto) hence "\y\(\a. wlp (prog a) P s) ` S s. 0 \ y" by(auto) } moreover have "\x\(\a. wlp (prog a) Q s) ` S s. \y\(\a. wlp (prog a) P s) ` S s. y \ x" proof(rule ballI, clarify, rule bexI) fix x and a assume ain: "a \ S s" from uP uQ le show "wlp (prog a) P s \ wlp (prog a) Q s" by(auto intro:le_funD[OF nearly_healthy_monoD[OF healthy, OF ain]]) from ain show "wlp (prog a) P s \ (\a. wlp (prog a) P s) ` S s" by(simp) qed ultimately show "Inf ((\a. wlp (prog a) P s) ` S s) \ Inf ((\a. wlp (prog a) Q s) ` S s)" by(intro cInf_mono, blast+) qed lemma healthy_wp_SetPC: fixes p::"'s \ 'a \ real" and f::"'a \ 's prog" assumes healthy: "\a s. a \ supp (p s) \ healthy (wp (f a))" and sound: "\s. sound (p s)" and sub_dist: "\s. (\a\supp (p s). p s a) \ 1" shows "healthy (wp (SetPC f p))" (is "healthy ?X") proof(intro healthy_parts bounded_byI nnegI le_funI, simp_all add:wp_eval) fix b and P::"'s \ real" and s::'s assume bP: "bounded_by b P" and nP: "nneg P" hence sP: "sound P" by(auto) from sP and bP and healthy have "\a s. a \ supp (p s) \ wp (f a) P s \ b" by(blast dest:healthy_bounded_byD) with sound have "(\a\supp (p s). p s a * wp (f a) P s) \ (\a\supp (p s). p s a * b)" by(blast intro:sum_mono mult_left_mono) also have "... = (\a\supp (p s). p s a) * b" by(simp add:sum_distrib_right) also { from bP and nP have "0 \ b" by(blast) with sub_dist have "(\a\supp (p s). p s a) * b \ 1 * b" by(rule mult_right_mono) } also have "1 * b = b" by(simp) finally show "(\a\supp (p s). p s a * wp (f a) P s) \ b" . show "0 \ (\a\supp (p s). p s a * wp (f a) P s)" proof(rule sum_nonneg [OF mult_nonneg_nonneg]) fix x from sound show "0 \ p s x" by(blast) assume "x \ supp (p s)" with sP and healthy show "0 \ wp (f x) P s" by(blast) qed next fix P::"'s \ real" and Q::"'s \ real" and s assume s_P: "sound P" and s_Q: "sound Q" and ent: "P \ Q" with healthy have "\a. a \ supp (p s) \ wp (f a) P s \ wp (f a) Q s" by(blast) with sound show "(\a\supp (p s). p s a * wp (f a) P s) \ (\a\supp (p s). p s a * wp (f a) Q s)" by(blast intro:sum_mono mult_left_mono) next fix P::"'s \ real" and c::real and s::'s assume sound: "sound P" and pos: "0 \ c" have "c * (\a\supp (p s). p s a * wp (f a) P s) = (\a\supp (p s). p s a * (c * wp (f a) P s))" (is "?A = ?B") by(simp add:sum_distrib_left ac_simps) also from sound and pos and healthy have "... = (\a\supp (p s). p s a * wp (f a) (\s. c * P s) s)" by(auto simp:scalingD[OF healthy_scalingD]) finally show "?A = ..." . qed lemma nearly_healthy_wlp_SetPC: fixes p::"'s \ 'a \ real" and f::"'a \ 's prog" assumes healthy: "\a s. a \ supp (p s) \ nearly_healthy (wlp (f a))" and sound: "\s. sound (p s)" and sub_dist: "\s. (\a\supp (p s). p s a) \ 1" shows "nearly_healthy (wlp (SetPC f p))" (is "nearly_healthy ?X") proof(intro nearly_healthyI unitaryI2 bounded_byI nnegI le_funI, simp_all only:wp_eval) fix b and P::"'s \ real" and s::'s assume uP: "unitary P" from uP healthy have "\a. a \ supp (p s) \ unitary (wlp (f a) P)" by(auto) hence "\a. a \ supp (p s) \ wlp (f a) P s \ 1" by(auto) with sound have "(\a\supp (p s). p s a * wlp (f a) P s) \ (\a\supp (p s). p s a * 1)" by(blast intro:sum_mono mult_left_mono) also have "... = (\a\supp (p s). p s a)" by(simp add:sum_distrib_right) also note sub_dist finally show "(\a\supp (p s). p s a * wlp (f a) P s) \ 1" . show "0 \ (\a\supp (p s). p s a * wlp (f a) P s)" proof(rule sum_nonneg [OF mult_nonneg_nonneg]) fix x from sound show "0 \ p s x" by(blast) assume "x \ supp (p s)" with uP healthy show "0 \ wlp (f x) P s" by(blast) qed next fix P::"'s expect" and Q::"'s expect" and s assume uP: "unitary P" and uQ: "unitary Q" and le: "P \ Q" hence "\a. a \ supp (p s) \ wlp (f a) P s \ wlp (f a) Q s" by(blast intro:le_funD[OF nearly_healthy_monoD, OF healthy]) with sound show "(\a\supp (p s). p s a * wlp (f a) P s) \ (\a\supp (p s). p s a * wlp (f a) Q s)" by(blast intro:sum_mono mult_left_mono) qed lemma healthy_wp_Apply: "healthy (wp (Apply f))" unfolding Apply_def wp_def by(blast) lemma nearly_healthy_wlp_Apply: "nearly_healthy (wlp (Apply f))" by(intro nearly_healthyI unitaryI2 nnegI bounded_byI, auto simp:o_def wp_eval) lemma healthy_wp_Bind: fixes f::"'s \ 'a" assumes hsub: "\s. healthy (wp (p (f s)))" shows "healthy (wp (Bind f p))" proof(intro healthy_parts nnegI bounded_byI le_funI, simp_all only:wp_eval) fix b and P::"'s expect" and s::'s assume bP: "bounded_by b P" and nP: "nneg P" with hsub have "bounded_by b (wp (p (f s)) P)" by(auto) thus "wp (p (f s)) P s \ b" by(auto) from bP nP hsub have "nneg (wp (p (f s)) P)" by(auto) thus "0 \ wp (p (f s)) P s" by(auto) next fix P Q::"'s expect" and s::'s assume "sound P" "sound Q" "P \ Q" thus "wp (p (f s)) P s \ wp (p (f s)) Q s" by(rule le_funD[OF mono_transD, OF healthy_monoD, OF hsub]) next fix P::"'s expect" and c::real and s::'s assume "sound P" and "0 \ c" thus "c * wp (p (f s)) P s = wp (p (f s)) (\s. c * P s) s" by(simp add:scalingD[OF healthy_scalingD, OF hsub]) qed lemma nearly_healthy_wlp_Bind: fixes f::"'s \ 'a" assumes hsub: "\s. nearly_healthy (wlp (p (f s)))" shows "nearly_healthy (wlp (Bind f p))" proof(intro nearly_healthyI unitaryI2 nnegI bounded_byI le_funI, simp_all only:wp_eval) fix P::"'s expect" and s::'s assume uP: "unitary P" with hsub have "unitary (wlp (p (f s)) P)" by(auto) thus "0 \ wlp (p (f s)) P s" "wlp (p (f s)) P s \ 1" by(auto) fix Q::"'s expect" assume "unitary Q" "P \ Q" with uP show "wlp (p (f s)) P s \ wlp (p (f s)) Q s" by(blast intro:le_funD[OF nearly_healthy_monoD, OF hsub]) qed subsection \Healthiness for Loops\ lemma wp_loop_step_mono: fixes t u::"'s trans" assumes hb: "healthy (wp body)" and le: "le_trans t u" and ht: "\P. sound P \ sound (t P)" and hu: "\P. sound P \ sound (u P)" shows "le_trans (wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip)) (wp (body ;; Embed u \<^bsub>\ G \\<^esub>\ Skip))" proof(intro le_transI le_funI, simp add:wp_eval) fix P::"'s expect" and s::'s assume sP: "sound P" with le have "t P \ u P" by(auto) moreover from sP ht hu have "sound (t P)" "sound (u P)" by(auto) ultimately have "wp body (t P) s \ wp body (u P) s" by(auto intro:le_funD[OF mono_transD, OF healthy_monoD, OF hb]) thus "\G\ s * wp body (t P) s \ \G\ s * wp body (u P) s " by(auto intro:mult_left_mono) qed lemma wlp_loop_step_mono: fixes t u::"'s trans" assumes mb: "nearly_healthy (wlp body)" and le: "le_utrans t u" and ht: "\P. unitary P \ unitary (t P)" and hu: "\P. unitary P \ unitary (u P)" shows "le_utrans (wlp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip)) (wlp (body ;; Embed u \<^bsub>\ G \\<^esub>\ Skip))" proof(intro le_utransI le_funI, simp add:wp_eval) fix P::"'s expect" and s::'s assume uP: "unitary P" with le have "t P \ u P" by(auto) moreover from uP ht hu have "unitary (t P)" "unitary (u P)" by(auto) ultimately have "wlp body (t P) s \ wlp body (u P) s" by(rule le_funD[OF nearly_healthy_monoD[OF mb]]) thus "\G\ s * wlp body (t P) s \ \G\ s * wlp body (u P) s " by(auto intro:mult_left_mono) qed text \For each sound expectation, we have a pre fixed point of the loop body. This lets us use the relevant fixed-point lemmas.\ lemma lfp_loop_fp: assumes hb: "healthy (wp body)" and sP: "sound P" shows "\s. \G\ s * wp body (\s. bound_of P) s + \\ G\ s * P s \ \s. bound_of P" proof(rule le_funI) fix s from sP have "sound (\s. bound_of P)" by(auto) moreover hence "bounded_by (bound_of P) (\s. bound_of P)" by(auto) ultimately have "bounded_by (bound_of P) (wp body (\s. bound_of P))" using hb by(auto) hence "wp body (\s. bound_of P) s \ bound_of P" by(auto) moreover from sP have "P s \ bound_of P" by(auto) ultimately have "\G\ s * wp body (\a. bound_of P) s + (1 - \G\ s) * P s \ \G\ s * bound_of P + (1 - \G\ s) * bound_of P" by(blast intro:add_mono mult_left_mono) thus "\G\ s * wp body (\a. bound_of P) s + \\ G\ s * P s \ bound_of P" by(simp add:algebra_simps negate_embed) qed lemma lfp_loop_greatest: fixes P::"'s expect" assumes lb: "\R. \s. \G\ s * wp body R s + \\ G\ s * P s \ R \ sound R \ Q \ R" and hb: "healthy (wp body)" and sP: "sound P" and sQ: "sound Q" shows "Q \ lfp_exp (\Q s. \G\ s * wp body Q s + \\ G\ s * P s)" using sP by(auto intro!:lfp_exp_greatest[OF lb sQ] sP lfp_loop_fp hb) lemma lfp_loop_sound: fixes P::"'s expect" assumes hb: "healthy (wp body)" and sP: "sound P" shows "sound (lfp_exp (\Q s. \G\ s * wp body Q s + \\ G\ s * P s))" using assms by(auto intro!:lfp_exp_sound lfp_loop_fp) lemma wlp_loop_step_unitary: fixes t u::"'s trans" assumes hb: "nearly_healthy (wlp body)" and ht: "\P. unitary P \ unitary (t P)" and uP: "unitary P" shows "unitary (wlp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) P)" proof(intro unitaryI2 nnegI bounded_byI, simp_all add:wp_eval) fix s::'s from ht uP have utP: "unitary (t P)" by(auto) with hb have "unitary (wlp body (t P))" by(auto) hence "0 \ wlp body (t P) s" by(auto) with uP show "0 \ \ G \ s * wlp body (t P) s + (1 - \ G \ s) * P s" by(auto intro!:add_nonneg_nonneg mult_nonneg_nonneg) from ht uP have "bounded_by 1 (t P)" by(auto) with utP hb have "bounded_by 1 (wlp body (t P))" by(auto) hence "wlp body (t P) s \ 1" by(auto) with uP have "\G\ s * wlp body (t P) s + (1 - \G\ s) * P s \ \G\ s * 1 + (1 - \G\ s) * 1" by(blast intro:add_mono mult_left_mono) also have "... = 1" by(simp) finally show "\G\ s * wlp body (t P) s + (1 - \G\ s) * P s \ 1" . qed lemma wp_loop_step_sound: fixes t u::"'s trans" assumes hb: "healthy (wp body)" and ht: "\P. sound P \ sound (t P)" and sP: "sound P" shows "sound (wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) P)" proof(intro soundI2 nnegI bounded_byI, simp_all add:wp_eval) fix s::'s from ht sP have stP: "sound (t P)" by(auto) with hb have "0 \ wp body (t P) s" by(auto) with sP show "0 \ \ G \ s * wp body (t P) s + (1 - \ G \ s) * P s" by(auto intro!:add_nonneg_nonneg mult_nonneg_nonneg) from ht sP have "sound (t P)" by(auto) moreover hence "bounded_by (bound_of (t P)) (t P)" by(auto) ultimately have "wp body (t P) s \ bound_of (t P)" using hb by(auto) hence "wp body (t P) s \ max (bound_of P) (bound_of (t P))" by(auto) moreover { from sP have "P s \ bound_of P" by(auto) hence "P s \ max (bound_of P) (bound_of (t P))" by(auto) } ultimately have "\G\ s * wp body (t P) s + (1 - \G\ s) * P s \ \G\ s * max (bound_of P) (bound_of (t P)) + (1 - \G\ s) * max (bound_of P) (bound_of (t P))" by(blast intro:add_mono mult_left_mono) also have "... = max (bound_of P) (bound_of (t P))" by(simp add:algebra_simps) finally show "\G\ s * wp body (t P) s + (1 - \G\ s) * P s \ max (bound_of P) (bound_of (t P))" . qed text \This gives the equivalence with the alternative definition for loops\citep[\S7, p.~198, footnote 23]{McIver_M_04}.\ lemma wlp_Loop1: fixes body :: "'s prog" assumes unitary: "unitary P" and healthy: "nearly_healthy (wlp body)" shows "wlp (do G \ body od) P = gfp_exp (\Q s. \G\ s * wlp body Q s + \\ G\ s * P s)" (is "?X = gfp_exp (?Y P)") proof - let "?Z u" = "(body ;; Embed u \<^bsub>\ G \\<^esub>\ Skip)" show ?thesis proof(simp only: wp_eval, intro gfp_pulldown assms le_funI) fix u P show "wlp (?Z u) P = ?Y P (u P)" by(simp add:wp_eval negate_embed) next fix t::"'s trans" and P::"'s expect" assume ut: "\Q. unitary Q \ unitary (t Q)" and uP: "unitary P" thus "unitary (wlp (?Z t) P)" by(rule wlp_loop_step_unitary[OF healthy]) next fix P Q::"'s expect" assume uP: "unitary P" and uQ: "unitary Q" show "unitary (\a. \ G \ a * wlp body Q a + \ \ G \ a * P a)" proof(intro unitaryI2 nnegI bounded_byI) fix s::'s from healthy uQ have "unitary (wlp body Q)" by(auto) hence "0 \ wlp body Q s" by(auto) with uP show "0 \ \G\ s * wlp body Q s + \\ G\ s * P s" by(auto intro!:add_nonneg_nonneg mult_nonneg_nonneg) from healthy uQ have "bounded_by 1 (wlp body Q)" by(auto) with uP have "\G\ s * wlp body Q s + (1 - \G\ s) * P s \ \G\ s * 1 + (1 - \G\ s) * 1" by(blast intro:add_mono mult_left_mono) also have "... = 1" by(simp) finally show "\G\ s * wlp body Q s + \\ G\ s * P s \ 1" by(simp add:negate_embed) qed next fix P Q R::"'s expect" and s::'s assume uP: "unitary P" and uQ: "unitary Q" and uR: "unitary R" and le: "Q \ R" hence "wlp body Q s \ wlp body R s" by(blast intro:le_funD[OF nearly_healthy_monoD, OF healthy]) thus "\G\ s * wlp body Q s + \\ G\ s * P s \ \G\ s * wlp body R s + \\ G\ s * P s" by(auto intro:mult_left_mono) next fix t u::"'s trans" assume "le_utrans t u" "\P. unitary P \ unitary (t P)" "\P. unitary P \ unitary (u P)" thus "le_utrans (wlp (?Z t)) (wlp (?Z u))" by(blast intro!:wlp_loop_step_mono[OF healthy]) qed qed lemma wp_loop_sound: assumes sP: "sound P" and hb: "healthy (wp body)" shows "sound (wp do G \ body od P)" proof(simp only: wp_eval, intro lfp_trans_sound sP) let ?v = "\P s. bound_of P" show "le_trans (wp (body ;; Embed ?v \<^bsub>\ G \\<^esub>\ Skip)) ?v" by(intro le_transI, simp add:wp_eval lfp_loop_fp[unfolded negate_embed] hb) show "\P. sound P \ sound (?v P)" by(auto) qed text \Likewise, we can rewrite strict loops.\ lemma wp_Loop1: fixes body :: "'s prog" assumes sP: "sound P" and healthy: "healthy (wp body)" shows "wp (do G \ body od) P = lfp_exp (\Q s. \G\ s * wp body Q s + \\ G\ s * P s)" (is "?X = lfp_exp (?Y P)") proof - let "?Z u" = "(body ;; Embed u \<^bsub>\ G \\<^esub>\ Skip)" show ?thesis proof(simp only: wp_eval, intro lfp_pulldown assms le_funI sP mono_transI) fix u P show "wp (?Z u) P = ?Y P (u P)" by(simp add:wp_eval negate_embed) next fix t::"'s trans" and P::"'s expect" assume ut: "\Q. sound Q \ sound (t Q)" and uP: "sound P" with healthy show "sound (wp (?Z t) P)" by(rule wp_loop_step_sound) next fix P Q::"'s expect" assume sP: "sound P" and sQ: "sound Q" show "sound (\a. \ G \ a * wp body Q a + \ \ G \ a * P a)" proof(intro soundI2 nnegI bounded_byI) fix s::'s from sQ have "nneg Q" "bounded_by (bound_of Q) Q" by(auto) with healthy have "bounded_by (bound_of Q) (wp body Q)" by(auto) hence "wp body Q s \ bound_of Q" by(auto) hence "wp body Q s \ max (bound_of P) (bound_of Q)" by(auto) moreover { from sP have "P s \ bound_of P" by(auto) hence "P s \ max (bound_of P) (bound_of Q)" by(auto) } ultimately have "\G\ s * wp body Q s + \\ G\ s * P s \ \G\ s * max (bound_of P) (bound_of Q) + \\ G\ s * max (bound_of P) (bound_of Q)" by(auto intro!:add_mono mult_left_mono) also have "... = max (bound_of P) (bound_of Q)" by(simp add:algebra_simps negate_embed) finally show "\G\ s * wp body Q s + \\ G\ s * P s \ max (bound_of P) (bound_of Q)" . from sP have "0 \ P s" by(auto) moreover from sQ healthy have "0 \ wp body Q s" by(auto) ultimately show "0 \ \G\ s * wp body Q s + \\ G\ s * P s" by(auto intro:add_nonneg_nonneg mult_nonneg_nonneg) qed next fix P Q R::"'s expect" and s::'s assume sQ: "sound Q" and sR: "sound R" and le: "Q \ R" hence "wp body Q s \ wp body R s" by(blast intro:le_funD[OF mono_transD, OF healthy_monoD, OF healthy]) thus "\G\ s * wp body Q s + \\ G\ s * P s \ \G\ s * wp body R s + \\ G\ s * P s" by(auto intro:mult_left_mono) next fix t u::"'s trans" assume le: "le_trans t u" and st: "\P. sound P \ sound (t P)" and su: "\P. sound P \ sound (u P)" with healthy show "le_trans (wp (?Z t)) (wp (?Z u))" by(rule wp_loop_step_mono) next from healthy show "le_trans (wp (?Z (\P s. bound_of P))) (\P s. bound_of P)" by(intro le_transI, simp add:wp_eval lfp_loop_fp[unfolded negate_embed]) next fix P::"'s expect" and s::'s assume "sound P" thus "sound (\s. bound_of P)" by(auto) qed qed lemma nearly_healthy_wlp_loop: fixes body::"'s prog" assumes hb: "nearly_healthy (wlp body)" shows "nearly_healthy (wlp (do G \ body od))" proof(intro nearly_healthyI unitaryI2 nnegI2 bounded_byI2, simp_all add:wlp_Loop1 hb) fix P::"'s expect" assume uP: "unitary P" let "?X R" = "\Q s. \ G \ s * wlp body Q s + \ \ G \ s * R s" show "\s. 0 \ gfp_exp (?X P)" proof(rule gfp_exp_upperbound) show "unitary (\s. 0::real)" by(auto) with hb have "unitary (wlp body (\s. 0))" by(auto) with uP show "\s. 0 \ (?X P (\s. 0))" by(blast intro!:le_funI add_nonneg_nonneg mult_nonneg_nonneg) qed show "gfp_exp (?X P) \ \s. 1" proof(rule gfp_exp_least) show "unitary (\s. 1::real)" by(auto) fix Q::"'s expect" assume "unitary Q" thus "Q \ \s. 1" by(auto) qed fix Q::"'s expect" assume uQ: "unitary Q" and le: "P \ Q" show "gfp_exp (?X P) \ gfp_exp (?X Q)" proof(rule gfp_exp_least) fix R::"'s expect" assume uR: "unitary R" assume fp: "R \ ?X P R" also from le have "... \ ?X Q R" by(blast intro:add_mono mult_left_mono le_funI) finally show "R \ gfp_exp (?X Q)" using uR by(auto intro:gfp_exp_upperbound) next show "unitary (gfp_exp (?X Q))" proof(rule gfp_exp_unitary, intro unitaryI2 nnegI bounded_byI) fix R::"'s expect" and s::'s assume uR: "unitary R" with hb have ubP: "unitary (wlp body R)" by(auto) with uQ show "0 \ \ G \ s * wlp body R s + \ \ G \ s * Q s" by(blast intro:add_nonneg_nonneg mult_nonneg_nonneg) from ubP uQ have "wlp body R s \ 1" "Q s \ 1" by(auto) hence "\ G \ s * wlp body R s + \ \ G \ s * Q s \ \G\ s * 1 + \\ G\ s * 1" by(blast intro:add_mono mult_left_mono) thus "\ G \ s * wlp body R s + \ \ G \ s * Q s \ 1" by(simp add:negate_embed) qed qed qed text \We show healthiness by appealing to the properties of expectation fixed points, applied to the alternative loop definition.\ lemma healthy_wp_loop: fixes body::"'s prog" assumes hb: "healthy (wp body)" shows "healthy (wp (do G \ body od))" proof - let "?X P" = "(\Q s. \G\ s * wp body Q s + \\ G\ s * P s)" show ?thesis proof(intro healthy_parts bounded_byI2 nnegI2, simp_all add:wp_Loop1 hb soundI2 sound_intros) fix P::"'s expect" and c::real and s::'s assume sP: "sound P" and nnc: "0 \ c" show "c * (lfp_exp (?X P)) s = lfp_exp (?X (\s. c * P s)) s" proof(cases) assume "c = 0" thus ?thesis proof(simp, intro antisym) from hb have fp: "\s. \G\ s * wp body (\_. 0) s \ \s. 0" by(simp) hence "lfp_exp (\P s. \G\ s * wp body P s) \ \s. 0" by(auto intro:lfp_exp_lowerbound) thus "lfp_exp (\P s. \G\ s * wp body P s) s \ 0" by(auto) have "\s. 0 \ lfp_exp (\P s. \G\ s * wp body P s)" by(auto intro:lfp_exp_greatest fp) thus "0 \ lfp_exp (\P s. \G\ s * wp body P s) s" by(auto) qed next have onesided: "\P c. c \ 0 \ 0 \ c \ sound P \ \a. c * lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) a \ lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * (c * P b))" proof - fix P::"'s expect" and c::real assume cnz: "c \ 0" and nnc: "0 \ c" and sP: "sound P" with nnc have cpos: "0 < c" by(auto) hence nnic: "0 \ inverse c" by(auto) show "\a. c * lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) a \ lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * (c * P b))" proof(rule lfp_exp_greatest) fix Q::"'s expect" assume sQ: "sound Q" and fp: "\b. \G\ b * wp body Q b + \\ G\ b * (c * P b) \ Q" hence "\s. \G\ s * wp body Q s + \\ G\ s * (c * P s) \ Q s" by(auto) with nnic have "\s. inverse c * (\G\ s * wp body Q s + \\ G\ s * (c * P s)) \ inverse c * Q s" by(auto intro:mult_left_mono) hence "\s. \G\ s * (inverse c * wp body Q s) + (inverse c * c) * \\ G\ s * P s \ inverse c * Q s" by(simp add:algebra_simps) hence "\s. \G\ s * wp body (\s. inverse c * Q s) s + \\ G\ s * P s \ inverse c * Q s" by(simp add:cnz scalingD[OF healthy_scalingD, OF hb sQ nnic]) hence "\s. \G\ s * wp body (\s. inverse c * Q s) s + \\ G\ s * P s \ \s. inverse c * Q s" by(rule le_funI) moreover from nnic sQ have "sound (\s. inverse c * Q s)" by(iprover intro:sound_intros) ultimately have "lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) \ \s. inverse c * Q s" by(rule lfp_exp_lowerbound) hence "\s. lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) s \ inverse c * Q s" by(rule le_funD) with nnc have "\s. c * lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) s \ c * (inverse c * Q s)" by(auto intro:mult_left_mono) also from cnz have "\s. ... s = Q s" by(simp) finally show "\a. c * lfp_exp (\a b. \G\ b * wp body a b + \\ G\ b * P b) a \ Q" by(rule le_funI) next from sP have "sound (\s. bound_of P)" by(auto) with hb sP have "sound (lfp_exp (?X P))" by(blast intro:lfp_exp_sound lfp_loop_fp) with nnc show "sound (\s. c * lfp_exp (?X P) s)" by(auto intro!:sound_intros) from hb sP nnc show "\s. \G\ s * wp body (\s. bound_of (\s. c * P s)) s + \\ G\ s * (c * P s) \ \s. bound_of (\s. c * P s)" by(iprover intro:lfp_loop_fp sound_intros) from sP nnc show "sound (\s. bound_of (\s. c * P s))" by(auto intro!:sound_intros) qed qed assume nzc: "c \ 0" show ?thesis (is "?X P c s = ?Y P c s") proof(rule fun_cong[where x=s], rule antisym) from nzc nnc sP show "?X P c \ ?Y P c" by(rule onesided) from nzc have nzic: "inverse c \ 0" by(auto) moreover with nnc have nnic: "0 \ inverse c" by(auto) moreover from nnc sP have scP: "sound (\s. c * P s)" by(auto intro!:sound_intros) ultimately have "?X (\s. c * P s) (inverse c) \ ?Y (\s. c * P s) (inverse c)" by(rule onesided) with nnc have "\s. c * ?X (\s. c * P s) (inverse c) s \ \s. c * ?Y (\s. c * P s) (inverse c) s" by(blast intro:mult_left_mono) with nzc show "?Y P c \ ?X P c" by(simp add:mult.assoc[symmetric]) qed qed next fix P::"'s expect" and b::real assume bP: "bounded_by b P" and nP: "nneg P" show "lfp_exp (\Q s. \G\ s * wp body Q s + \\ G\ s * P s) \ \s. b" proof(intro lfp_exp_lowerbound le_funI) fix s::'s from bP nP hb have "bounded_by b (wp body (\s. b))" by(auto) hence "wp body (\s. b) s \ b" by(auto) moreover from bP have "P s \ b" by(auto) ultimately have "\G\ s * wp body (\s. b) s + \\ G\ s * P s \ \G\ s * b + \\ G\ s * b" by(auto intro!:add_mono mult_left_mono) also have "... = b" by(simp add:negate_embed field_simps) finally show "\G\ s * wp body (\s. b) s + \\ G\ s * P s \ b" . from bP nP have "0 \ b" by(auto) thus "sound (\s. b)" by(auto) qed from hb bP nP show "\s. 0 \ lfp_exp (\Q s. \G\ s * wp body Q s + \\ G\ s * P s)" by(auto dest!:sound_nneg intro!:lfp_loop_greatest) next fix P Q::"'s expect" assume sP: "sound P" and sQ: "sound Q" and le: "P \ Q" show "lfp_exp (?X P) \ lfp_exp (?X Q)" proof(rule lfp_exp_greatest) fix R::"'s expect" assume sR: "sound R" and fp: "\s. \G\ s * wp body R s + \\ G\ s * Q s \ R" from le have "\s. \G\ s * wp body R s + \\ G\ s * P s \ \s. \G\ s * wp body R s + \\ G\ s * Q s" by(auto intro:le_funI add_left_mono mult_left_mono) also note fp finally show "lfp_exp (\R s. \G\ s * wp body R s + \\ G\ s * P s) \ R" using sR by(auto intro:lfp_exp_lowerbound) next from hb sP show "sound (lfp_exp (\R s. \G\ s * wp body R s + \\ G\ s * P s))" by(rule lfp_loop_sound) from hb sQ show "\s. \G\ s * wp body (\s. bound_of Q) s + \\ G\ s * Q s \ \s. bound_of Q" by(rule lfp_loop_fp) from sQ show "sound (\s. bound_of Q)" by(auto) qed qed qed text \Use 'simp add:healthy\_intros' or 'blast intro:healthy\_intros' as appropriate to discharge healthiness side-contitions for primitive programs automatically.\ lemmas healthy_intros = healthy_wp_Abort nearly_healthy_wlp_Abort healthy_wp_Skip nearly_healthy_wlp_Skip healthy_wp_Seq nearly_healthy_wlp_Seq healthy_wp_PC nearly_healthy_wlp_PC healthy_wp_DC nearly_healthy_wlp_DC healthy_wp_AC nearly_healthy_wlp_AC healthy_wp_Embed nearly_healthy_wlp_Embed healthy_wp_Apply nearly_healthy_wlp_Apply healthy_wp_SetDC nearly_healthy_wlp_SetDC healthy_wp_SetPC nearly_healthy_wlp_SetPC healthy_wp_Bind nearly_healthy_wlp_Bind healthy_wp_repeat nearly_healthy_wlp_repeat healthy_wp_loop nearly_healthy_wlp_loop end diff --git a/thys/pGCL/Sublinearity.thy b/thys/pGCL/Sublinearity.thy --- a/thys/pGCL/Sublinearity.thy +++ b/thys/pGCL/Sublinearity.thy @@ -1,461 +1,461 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section \Sublinearity\ theory Sublinearity imports Embedding Healthiness LoopInduction begin subsection \Nonrecursive Primitives\ text \Sublinearity of non-recursive programs is generally straightforward, and follows from the alebraic properties of the underlying operations, together with healthiness.\ lemma sublinear_wp_Skip: "sublinear (wp Skip)" by(auto simp:wp_eval) lemma sublinear_wp_Abort: "sublinear (wp Abort)" by(auto simp:wp_eval) lemma sublinear_wp_Apply: "sublinear (wp (Apply f))" by(auto simp:wp_eval) lemma sublinear_wp_Seq: fixes x::"'s prog" assumes slx: "sublinear (wp x)" and sly: "sublinear (wp y)" and hx: "healthy (wp x)" and hy: "healthy (wp y)" shows "sublinear (wp (x ;; y))" proof(rule sublinearI, simp add:wp_eval) fix P::"'s \ real" and Q::"'s \ real" and s::'s and a::real and b::real and c::real assume sP: "sound P" and sQ: "sound Q" and nna: "0 \ a" and nnb: "0 \ b" and nnc: "0 \ c" with slx hy have "a * wp x (wp y P) s + b * wp x (wp y Q) s \ c \ wp x (\s. a * wp y P s + b * wp y Q s \ c) s" by(blast intro:sublinearD) also { from sP sQ nna nnb nnc sly have "\s. a * wp y P s + b * wp y Q s \ c \ wp y (\s. a * P s + b * Q s \ c) s" by(blast intro:sublinearD) moreover from sP sQ hy have "sound (wp y P)" and "sound (wp y Q)" by(auto) moreover with nna nnb nnc have "sound (\s. a * wp y P s + b * wp y Q s \ c)" by(auto intro!:sound_intros tminus_sound) moreover from sP sQ nna nnb nnc have "sound (\s. a * P s + b * Q s \ c)" by(auto intro!:sound_intros tminus_sound) moreover with hy have "sound (wp y (\s. a * P s + b * Q s \ c))" by(blast) ultimately have "wp x (\s. a * wp y P s + b * wp y Q s \ c) s \ wp x (wp y (\s. a * P s + b * Q s \ c)) s" by(blast intro!:le_funD[OF mono_transD[OF healthy_monoD[OF hx]]]) } finally show "a * wp x (wp y P) s + b * wp x (wp y Q) s \ c \ wp x (wp y (\s. a * P s + b * Q s \ c)) s" . qed lemma sublinear_wp_PC: fixes x::"'s prog" assumes slx: "sublinear (wp x)" and sly: "sublinear (wp y)" and uP: "unitary P" shows "sublinear (wp (x \<^bsub>P\<^esub>\ y))" proof(rule sublinearI, simp add:wp_eval) fix R::"'s \ real" and Q::"'s \ real" and s::'s and a::real and b::real and c::real assume sR: "sound R" and sQ: "sound Q" and nna: "0 \ a" and nnb: "0 \ b" and nnc: "0 \ c" have "a * (P s * wp x Q s + (1 - P s) * wp y Q s) + b * (P s * wp x R s + (1 - P s) * wp y R s) \ c = (P s * a * wp x Q s + (1 - P s) * a * wp y Q s) + (P s * b * wp x R s + (1 - P s) * b * wp y R s) \ c" by(simp add:field_simps) also have "... = (P s * a * wp x Q s + P s * b * wp x R s) + ((1 - P s) * a * wp y Q s + (1 - P s) * b * wp y R s) \ c" by(simp add:ac_simps) also have "... = P s * (a * wp x Q s + b * wp x R s) + (1 - P s) * (a * wp y Q s + b * wp y R s) \ (P s * c + (1 - P s) * c)" by(simp add:field_simps) also have "... \ (P s * (a * wp x Q s + b * wp x R s) \ P s * c) + ((1 - P s) * (a * wp y Q s + b * wp y R s) \ (1 - P s) * c)" by(rule tminus_add_mono) also { from uP have "0 \ P s" and "0 \ 1 - P s" - by(auto simp:sign_simps) + by auto hence "(P s * (a * wp x Q s + b * wp x R s) \ P s * c) + ((1 - P s) * (a * wp y Q s + b * wp y R s) \ (1 - P s) * c) = P s * (a * wp x Q s + b * wp x R s \ c) + (1 - P s) * (a * wp y Q s + b * wp y R s \ c)" by(simp add:tminus_left_distrib) } also { from sQ sR nna nnb nnc slx have "a * wp x Q s + b * wp x R s \ c \ wp x (\s. a * Q s + b * R s \ c) s" by(blast) moreover from sQ sR nna nnb nnc sly have "a * wp y Q s + b * wp y R s \ c \ wp y (\s. a * Q s + b * R s \ c) s" by(blast) moreover from uP have "0 \ P s" and "0 \ 1 - P s" - by(auto simp:sign_simps) + by auto ultimately have "P s * (a * wp x Q s + b * wp x R s \ c) + (1 - P s) * (a * wp y Q s + b * wp y R s \ c) \ P s * wp x (\s. a * Q s + b * R s \ c) s + (1 - P s) * wp y (\s. a * Q s + b * R s \ c) s" by(blast intro:add_mono mult_left_mono) } finally show " a * (P s * wp x Q s + (1 - P s) * wp y Q s) + b * (P s * wp x R s + (1 - P s) * wp y R s) \ c \ P s * wp x (\s. a * Q s + b * R s \ c) s + (1 - P s) * wp y (\s. a * Q s + b * R s \ c) s" . qed lemma sublinear_wp_DC: fixes x::"'s prog" assumes slx: "sublinear (wp x)" and sly: "sublinear (wp y)" shows "sublinear (wp (x \ y))" proof(rule sublinearI, simp only:wp_eval) fix R::"'s \ real" and Q::"'s \ real" and s::'s and a::real and b::real and c::real assume sR: "sound R" and sQ: "sound Q" and nna: "0 \ a" and nnb: "0 \ b" and nnc: "0 \ c" from nna nnb have "a * min (wp x Q s) (wp y Q s) + b * min (wp x R s) (wp y R s) \ c = min (a * wp x Q s) (a * wp y Q s) + min (b * wp x R s) (b * wp y R s) \ c" by(simp add:min_distrib) also have "... \ min (a * wp x Q s + b * wp x R s) (a * wp y Q s + b * wp y R s) \ c" by(auto intro!:tminus_left_mono) also have "... = min (a * wp x Q s + b * wp x R s \ c) (a * wp y Q s + b * wp y R s \ c)" by(rule min_tminus_distrib) also { from slx sQ sR nna nnb nnc have "a * wp x Q s + b * wp x R s \ c \ wp x (\s. a * Q s + b * R s \ c) s" by(blast) moreover from sly sQ sR nna nnb nnc have "a * wp y Q s + b * wp y R s \ c \ wp y (\s. a * Q s + b * R s \ c) s" by(blast) ultimately have "min (a * wp x Q s + b * wp x R s \ c) (a * wp y Q s + b * wp y R s \ c) \ min (wp x (\s. a * Q s + b * R s \ c) s) (wp y (\s. a * Q s + b * R s \ c) s)" by(auto) } finally show "a * min (wp x Q s) (wp y Q s) + b * min (wp x R s) (wp y R s) \ c \ min (wp x (\s. a * Q s + b * R s \ c) s) (wp y (\s. a * Q s + b * R s \ c) s)" . qed text \As for continuity, we insist on a finite support.\ lemma sublinear_wp_SetPC: fixes p::"'a \ 's prog" assumes slp: "\s a. a \ supp (P s) \ sublinear (wp (p a))" and sum: "\s. (\a\supp (P s). P s a) \ 1" and nnP: "\s a. 0 \ P s a" and fin: "\s. finite (supp (P s))" shows "sublinear (wp (SetPC p P))" proof(rule sublinearI, simp add:wp_eval) fix R::"'s \ real" and Q::"'s \ real" and s::'s and a::real and b::real and c::real assume sR: "sound R" and sQ: "sound Q" and nna: "0 \ a" and nnb: "0 \ b" and nnc: "0 \ c" have "a * (\a'\supp (P s). P s a' * wp (p a') Q s) + b * (\a'\supp (P s). P s a' * wp (p a') R s) \ c = (\a'\supp (P s). P s a' * (a * wp (p a') Q s + b * wp (p a') R s)) \ c" by(simp add:field_simps sum_distrib_left sum.distrib) also have "... \ (\a'\supp (P s). P s a' * (a * wp (p a') Q s + b * wp (p a') R s)) \ (\a'\supp (P s). P s a' * c)" proof(rule tminus_right_antimono) have "(\a'\supp (P s). P s a' * c) \ (\a'\supp (P s). P s a') * c" by(simp add:sum_distrib_right) also from sum and nnc have "... \ 1 * c" by(rule mult_right_mono) finally show "(\a'\supp (P s). P s a' * c) \ c" by(simp) qed also from fin have "... \ (\a'\supp (P s). P s a' * (a * wp (p a') Q s + b * wp (p a') R s) \ P s a' * c)" by(blast intro:tminus_sum_mono) also have "... = (\a'\supp (P s). P s a' * (a * wp (p a') Q s + b * wp (p a') R s \ c))" by(simp add:nnP tminus_left_distrib) also { from slp sQ sR nna nnb nnc have "\a'. a' \ supp (P s) \ a * wp (p a') Q s + b * wp (p a') R s \ c \ wp (p a') (\s. a * Q s + b * R s \ c) s" by(blast) with nnP have "(\a'\supp (P s). P s a' * (a * wp (p a') Q s + b * wp (p a') R s \ c)) \ (\a'\supp (P s). P s a' * wp (p a') (\s. a * Q s + b * R s \ c) s)" by(blast intro:sum_mono mult_left_mono) } finally show "a * (\a'\supp (P s). P s a' * wp (p a') Q s) + b * (\a'\supp (P s). P s a' * wp (p a') R s) \ c \ (\a'\supp (P s). P s a' * wp (p a') (\s. a * Q s + b * R s \ c) s)" . qed lemma sublinear_wp_SetDC: fixes p::"'a \ 's prog" assumes slp: "\s a. a \ S s \ sublinear (wp (p a))" and hp: "\s a. a \ S s \ healthy (wp (p a))" and ne: "\s. S s \ {}" shows "sublinear (wp (SetDC p S))" proof(rule sublinearI, simp add:wp_eval, rule cInf_greatest) fix P::"'s \ real" and Q::"'s \ real" and s::'s and x y and a::real and b::real and c::real assume sP: "sound P" and sQ: "sound Q" and nna: "0 \ a" and nnb: "0 \ b" and nnc: "0 \ c" from ne show "(\pr. wp (p pr) (\s. a * P s + b * Q s \ c) s) ` S s \ {}" by(auto) assume yin: "y \ (\pr. wp (p pr) (\s. a * P s + b * Q s \ c) s) ` S s" then obtain x where xin: "x \ S s" and rwy: "y = wp (p x) (\s. a * P s + b * Q s \ c) s" by(auto) from xin hp sP nna have "a * Inf ((\a. wp (p a) P s) ` S s) \ a * wp (p x) P s" by(intro mult_left_mono[OF cInf_lower] bdd_belowI[where m=0], blast+) moreover from xin hp sQ nnb have "b * Inf ((\a. wp (p a) Q s) ` S s) \ b * wp (p x) Q s" by(intro mult_left_mono[OF cInf_lower] bdd_belowI[where m=0], blast+) ultimately have "a * Inf ((\a. wp (p a) P s) ` S s) + b * Inf ((\a. wp (p a) Q s) ` S s) \ c \ a * wp (p x) P s + b * wp (p x) Q s \ c" by(blast intro:tminus_left_mono add_mono) also from xin slp sP sQ nna nnb nnc have "... \ wp (p x) (\s. a * P s + b * Q s \ c) s" by(blast) finally show "a * Inf ((\a. wp (p a) P s) ` S s) + b * Inf ((\a. wp (p a) Q s) ` S s) \ c \ y" by(simp add:rwy) qed lemma sublinear_wp_Embed: "sublinear t \ sublinear (wp (Embed t))" by(simp add:wp_eval) lemma sublinear_wp_repeat: "\ sublinear (wp p); healthy (wp p) \ \ sublinear (wp (repeat n p))" by(induct n, simp_all add:sublinear_wp_Seq sublinear_wp_Skip healthy_wp_repeat) lemma sublinear_wp_Bind: "\ \s. sublinear (wp (a (f s))) \ \ sublinear (wp (Bind f a))" by(rule sublinearI, simp add:wp_eval, auto) subsection \Sublinearity for Loops\ text \We break the proof of sublinearity loops into separate proofs of sub-distributivity and sub-additivity. The first follows by transfinite induction.\ lemma sub_distrib_wp_loop: fixes body::"'s prog" assumes sdb: "sub_distrib (wp body)" and hb: "healthy (wp body)" and nhb: "nearly_healthy (wlp body)" shows "sub_distrib (wp (do G \ body od))" proof - have "\P s. sound P \ wp (do G \ body od) P s \ 1 \ wp (do G \ body od) (\s. P s \ 1) s" proof(rule loop_induct[OF hb nhb], safe) fix S::"('s trans \ 's trans) set" and P::"'s expect" and s::'s assume saS: "\x\S. \P s. sound P \ fst x P s \ 1 \ fst x (\s. P s \ 1) s" and sP: "sound P" and fS: "\x\S. feasible (fst x)" from sP have sPm: "sound (\s. P s \ 1)" by(auto intro:tminus_sound) have nnSup: "\s. 0 \ Sup_trans (fst ` S) (\s. P s \ 1) s" proof(cases "S={}", simp add:Sup_trans_def Sup_exp_def) fix s assume "S \ {}" then obtain x where xin: "x\S" by(auto) with fS sPm have "0 \ fst x (\s. P s \ 1) s" by(auto) also from xin fS sPm have "... \ Sup_trans (fst ` S) (\s. P s \ 1) s" by(auto intro!: le_funD[OF Sup_trans_upper2]) finally show "?thesis s" . qed have "\x s. fst x P s \ (fst x P s \ 1) + 1" by(simp add:tminus_def) also from saS sP have "\x s. x\S \ (fst x P s \ 1) + 1 \ fst x (\s. P s \ 1) s + 1" by(auto intro:add_right_mono) also { from sP have "sound (\s. P s \ 1)" by(auto intro:tminus_sound) with fS have "\x s. x\S \ fst x (\s. P s \ 1) s + 1 \ Sup_trans (fst ` S) (\s. P s \ 1) s + 1" by(blast intro!: add_right_mono le_funD[OF Sup_trans_upper2]) } finally have le: "\s. \x\S. fst x P s \ Sup_trans (fst ` S) (\s. P s \ 1) s + 1" by(auto) moreover from nnSup have nn: "\s. 0 \ Sup_trans (fst ` S) (\s. P s \ 1) s + 1" by(auto intro:add_nonneg_nonneg) ultimately have leSup: "Sup_trans (fst ` S) P s \ Sup_trans (fst ` S) (\s. P s \ 1) s + 1" unfolding Sup_trans_def by(intro le_funD[OF Sup_exp_least], auto) show "Sup_trans (fst ` S) P s \ 1 \ Sup_trans (fst ` S) (\s. P s \ 1) s" proof(cases "Sup_trans (fst ` S) P s \ 1", simp_all add:nnSup) from leSup have "Sup_trans (fst ` S) P s - 1 \ Sup_trans (fst ` S) (\s. P s \ 1) s + 1 - 1" by(auto) thus "Sup_trans (fst ` S) P s - 1 \ Sup_trans (fst ` S) (\s. P s \ 1) s" by(simp) qed next fix t::"'s trans" and P::"'s expect" and s::'s assume IH: "\P s. sound P \ t P s \ 1 \ t (\a. P a \ 1) s" and ft: "feasible t" and sP: "sound P" from sP have "sound (\s. P s \ 1)" by(auto intro:tminus_sound) with ft have s2: "sound (t (\s. P s \ 1))" by(auto) from sP ft have "sound (t P)" by(auto) hence s3: "sound (\s. t P s \ 1)" by(auto intro!:tminus_sound) show "wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) P s \ 1 \ wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) (\a. P a \ 1) s" proof(simp add:wp_eval) have "\G\ s * wp body (t P) s + (1 - \G\ s) * P s \ 1 = \G\ s * wp body (t P) s + (1 - \G\ s) * P s \ (\G\ s + (1 - \G\ s))" by(simp) also have "... \ (\G\ s * wp body (t P) s \ \G\ s) + ((1 - \G\ s) * P s \ (1 - \G\ s))" by(rule tminus_add_mono) also have "... = \G\ s * (wp body (t P) s \ 1) + (1 - \G\ s) * (P s \ 1)" by(simp add:tminus_left_distrib) also { from ft sP have "wp body (t P) s \ 1 \ wp body (\s. t P s \ 1) s" by(auto intro:sub_distribD[OF sdb]) also { from IH sP have "\s. t P s \ 1 \ t (\s. P s \ 1)" by(auto) with sP ft s2 s3 have "wp body (\s. t P s \ 1) s \ wp body (t (\s. P s \ 1)) s" by(blast intro:le_funD[OF mono_transD, OF healthy_monoD, OF hb]) } finally have "\G\ s * (wp body (t P) s \ 1) + (1 - \G\ s) * (P s \ 1) \ \G\ s * wp body (t (\s. P s \ 1)) s + (1 - \G\ s) * (P s \ 1)" by(auto intro:add_right_mono mult_left_mono) } finally show "\G\ s * wp body (t P) s + (1 - \G\ s) * P s \ 1 \ \G\ s * wp body (t (\s. P s \ 1)) s + (1 - \G\ s) * (P s \ 1)" . qed next fix t t'::"'s trans" and P::"'s expect" and s::'s assume IH: "\P s. sound P \ t P s \ 1 \ t (\a. P a \ 1) s" and eq: "equiv_trans t t'" and sP: "sound P" from sP have "t' P s \ 1 = t P s \ 1" by(simp add:equiv_transD[OF eq]) also from sP IH have "... \ t (\s. P s \ 1) s" by(auto) also { from sP have "sound (\s. P s \ 1)" by(simp add:tminus_sound) hence "t (\s. P s \ 1) s = t' (\s. P s \ 1) s" by(simp add:equiv_transD[OF eq]) } finally show "t' P s \ 1 \ t' (\s. P s \ 1) s" . qed thus ?thesis by(auto intro!:sub_distribI) qed text \For sub-additivity, we again use the limit-of-iterates characterisation. Firstly, all iterates are sublinear:\ lemma sublinear_iterates: assumes hb: "healthy (wp body)" and sb: "sublinear (wp body)" shows "sublinear (iterates body G i)" by(induct i, auto intro!:sublinear_wp_PC sublinear_wp_Seq sublinear_wp_Skip sublinear_wp_Embed assms healthy_intros iterates_healthy) text \From this, sub-additivity follows for the limit (i.e. the loop), by appealing to the property at all steps.\ lemma sub_add_wp_loop: fixes body::"'s prog" assumes sb: "sublinear (wp body)" and cb: "bd_cts (wp body)" and hwp: "healthy (wp body)" shows "sub_add (wp (do G \ body od))" proof fix P Q::"'s expect" and s::'s assume sP: "sound P" and sQ: "sound Q" from hwp cb sP have "(\i. iterates body G i P s) \ wp do G \ body od P s" by(rule loop_iterates) moreover from hwp cb sQ have "(\i. iterates body G i Q s) \ wp do G \ body od Q s" by(rule loop_iterates) ultimately have "(\i. iterates body G i P s + iterates body G i Q s) \ wp do G \ body od P s + wp do G \ body od Q s" by(rule tendsto_add) moreover { from sublinear_subadd[OF sublinear_iterates, OF hwp sb, OF healthy_feasibleD[OF iterates_healthy, OF hwp]] sP sQ have "\i. iterates body G i P s + iterates body G i Q s \ iterates body G i (\s. P s + Q s) s" by(rule sub_addD) } moreover { from sP sQ have "sound (\s. P s + Q s)" by(blast intro:sound_intros) with hwp cb have "(\i. iterates body G i (\s. P s + Q s) s) \ wp do G \ body od (\s. P s + Q s) s" by(blast intro:loop_iterates) } ultimately show "wp do G \ body od P s + wp do G \ body od Q s \ wp do G \ body od (\s. P s + Q s) s" by(blast intro:LIMSEQ_le) qed lemma sublinear_wp_loop: fixes body::"'s prog" assumes hb: "healthy (wp body)" and nhb: "nearly_healthy (wlp body)" and sb: "sublinear (wp body)" and cb: "bd_cts (wp body)" shows "sublinear (wp (do G \ body od))" using sublinear_sub_distrib[OF sb] sublinear_subadd[OF sb] hb healthy_feasibleD[OF hb] by(iprover intro:sd_sa_sublinear[OF _ _ healthy_wp_loop[OF hb]] sub_distrib_wp_loop sub_add_wp_loop assms) lemmas sublinear_intros = sublinear_wp_Abort sublinear_wp_Skip sublinear_wp_Apply sublinear_wp_Seq sublinear_wp_PC sublinear_wp_DC sublinear_wp_SetPC sublinear_wp_SetDC sublinear_wp_Embed sublinear_wp_repeat sublinear_wp_Bind sublinear_wp_loop end diff --git a/thys/pGCL/WellDefined.thy b/thys/pGCL/WellDefined.thy --- a/thys/pGCL/WellDefined.thy +++ b/thys/pGCL/WellDefined.thy @@ -1,813 +1,813 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section \Well-Defined Programs.\ theory WellDefined imports Healthiness Sublinearity LoopInduction begin text \The definition of a well-defined program collects the various notions of healthiness and well-behavedness that we have so far established: healthiness of the strict and liberal transformers, continuity and sublinearity of the strict transformers, and two new properties. These are that the strict transformer always lies below the liberal one (i.e. that it is at least as \emph{strict}, recalling the standard embedding of a predicate), and that expectation conjunction is distributed between then in a particular manner, which will be crucial in establishing the loop rules.\ subsection \Strict Implies Liberal\ text \This establishes the first connection between the strict and liberal interpretations (@{term wp} and @{term wlp}).\ definition wp_under_wlp :: "'s prog \ bool" where "wp_under_wlp prog \ \P. unitary P \ wp prog P \ wlp prog P" lemma wp_under_wlpI[intro]: "\ \P. unitary P \ wp prog P \ wlp prog P \ \ wp_under_wlp prog" unfolding wp_under_wlp_def by(simp) lemma wp_under_wlpD[dest]: "\ wp_under_wlp prog; unitary P \ \ wp prog P \ wlp prog P" unfolding wp_under_wlp_def by(simp) lemma wp_under_le_trans: "wp_under_wlp a \ le_utrans (wp a) (wlp a)" by(blast) lemma wp_under_wlp_Abort: "wp_under_wlp Abort" by(rule wp_under_wlpI, unfold wp_eval, auto) lemma wp_under_wlp_Skip: "wp_under_wlp Skip" by(rule wp_under_wlpI, unfold wp_eval, blast) lemma wp_under_wlp_Apply: "wp_under_wlp (Apply f)" by(auto simp:wp_eval) lemma wp_under_wlp_Seq: assumes h_wlp_a: "nearly_healthy (wlp a)" and h_wp_b: "healthy (wp b)" and h_wlp_b: "nearly_healthy (wlp b)" and wp_u_a: "wp_under_wlp a" and wp_u_b: "wp_under_wlp b" shows "wp_under_wlp (a ;; b)" proof(rule wp_under_wlpI, unfold wp_eval o_def) fix P::"'a \ real" assume uP: "unitary P" with h_wp_b have "unitary (wp b P)" by(blast) with wp_u_a have "wp a (wp b P) \ wlp a (wp b P)" by(auto) also { from wp_u_b and uP have "wp b P \ wlp b P" by(blast) with h_wlp_a and h_wlp_b and h_wp_b and uP have "wlp a (wp b P) \ wlp a (wlp b P)" by(blast intro:nearly_healthy_monoD[OF h_wlp_a]) } finally show "wp a (wp b P) \ wlp a (wlp b P)" . qed lemma wp_under_wlp_PC: assumes h_wp_a: "healthy (wp a)" and h_wlp_a: "nearly_healthy (wlp a)" and h_wp_b: "healthy (wp b)" and h_wlp_b: "nearly_healthy (wlp b)" and wp_u_a: "wp_under_wlp a" and wp_u_b: "wp_under_wlp b" and uP: "unitary P" shows "wp_under_wlp (a \<^bsub>P\<^esub>\ b)" proof(rule wp_under_wlpI, unfold wp_eval, rule le_funI) fix Q::"'a \ real" and s assume uQ: "unitary Q" from uP have "P s \ 1" by(blast) hence "0 \ 1 - P s" by(simp) moreover from uQ and wp_u_b have "wp b Q s \ wlp b Q s" by(blast) ultimately have "(1 - P s) * wp b Q s \ (1 - P s) * wlp b Q s" by(blast intro:mult_left_mono) moreover { from uQ and wp_u_a have "wp a Q s \ wlp a Q s" by(blast) with uP have "P s * wp a Q s \ P s * wlp a Q s" by(blast intro:mult_left_mono) } ultimately show "P s * wp a Q s + (1 - P s) * wp b Q s \ P s * wlp a Q s + (1 - P s) * wlp b Q s" by(blast intro: add_mono) qed lemma wp_under_wlp_DC: assumes wp_u_a: "wp_under_wlp a" and wp_u_b: "wp_under_wlp b" shows "wp_under_wlp (a \ b)" proof(rule wp_under_wlpI, unfold wp_eval, rule le_funI) fix Q::"'a \ real" and s assume uQ: "unitary Q" from wp_u_a uQ have "wp a Q s \ wlp a Q s" by(blast) moreover from wp_u_b uQ have "wp b Q s \ wlp b Q s" by(blast) ultimately show "min (wp a Q s) (wp b Q s) \ min (wlp a Q s) (wlp b Q s)" by(auto) qed lemma wp_under_wlp_SetPC: assumes wp_u_f: "\s a. a \ supp (P s) \ wp_under_wlp (f a)" and nP: "\s a. a \ supp (P s) \ 0 \ P s a" shows "wp_under_wlp (SetPC f P)" proof(rule wp_under_wlpI, unfold wp_eval, rule le_funI) fix Q::"'a \ real" and s assume uQ: "unitary Q" from wp_u_f uQ nP show "(\a\supp (P s). P s a * wp (f a) Q s) \ (\a\supp (P s). P s a * wlp (f a) Q s)" by(auto intro!:sum_mono mult_left_mono) qed lemma wp_under_wlp_SetDC: assumes wp_u_f: "\s a. a \ S s \ wp_under_wlp (f a)" and hf: "\s a. a \ S s \ healthy (wp (f a))" and nS: "\s. S s \ {}" shows "wp_under_wlp (SetDC f S)" proof(rule wp_under_wlpI, rule le_funI, unfold wp_eval) fix Q::"'a \ real" and s assume uQ: "unitary Q" show "Inf ((\a. wp (f a) Q s) ` S s) \ Inf ((\a. wlp (f a) Q s) ` S s)" proof(rule cInf_mono) from nS show "(\a. wlp (f a) Q s) ` S s \ {}" by(blast) fix x assume xin: "x \ (\a. wlp (f a) Q s) ` S s" then obtain a where ain: "a \ S s" and xrw: "x = wlp (f a) Q s" by(blast) with wp_u_f uQ have "wp (f a) Q s \ wlp (f a) Q s" by(blast) moreover from ain have "wp (f a) Q s \ (\a. wp (f a) Q s) ` S s" by(blast) ultimately show "\y\ (\a. wp (f a) Q s) ` S s. y \ x" by(auto simp:xrw) next fix y assume yin: "y \ (\a. wp (f a) Q s) ` S s" then obtain a where ain: "a \ S s" and yrw: "y = wp (f a) Q s" by(blast) with hf uQ have "unitary (wp (f a) Q)" by(auto) with yrw show "0 \ y" by(auto) qed qed lemma wp_under_wlp_Embed: "wp_under_wlp (Embed t)" by(rule wp_under_wlpI, unfold wp_eval, blast) lemma wp_under_wlp_loop: fixes body::"'s prog" assumes hwp: "healthy (wp body)" and hwlp: "nearly_healthy (wlp body)" and wp_under: "wp_under_wlp body" shows "wp_under_wlp (do G \ body od)" proof(rule wp_under_wlpI) fix P::"'s expect" assume uP: "unitary P" hence sP: "sound P" by(auto) let "?X Q s" = "\G\ s * wp body Q s + \\ G\ s * P s" let "?Y Q s" = "\G\ s * wlp body Q s + \\ G\ s * P s" show "wp (do G \ body od) P \ wlp (do G \ body od) P" proof(simp add:hwp hwlp sP uP wp_Loop1 wlp_Loop1, rule gfp_exp_upperbound) thm lfp_loop_fp from hwp sP have "lfp_exp ?X = ?X (lfp_exp ?X)" by(rule lfp_wp_loop_unfold) hence "lfp_exp ?X \ ?X (lfp_exp ?X)" by(simp) also { from hwp uP have "wp body (lfp_exp ?X) \ wlp body (lfp_exp ?X)" by(auto intro:wp_under_wlpD[OF wp_under] lfp_loop_unitary) hence "?X (lfp_exp ?X) \ ?Y (lfp_exp ?X)" by(auto intro:add_mono mult_left_mono) } finally show "lfp_exp ?X \ ?Y (lfp_exp ?X)" . from hwp uP show "unitary (lfp_exp ?X)" by(auto intro:lfp_loop_unitary) qed qed lemma wp_under_wlp_repeat: "\ healthy (wp a); nearly_healthy (wlp a); wp_under_wlp a \ \ wp_under_wlp (repeat n a)" by(induct n, auto intro!:wp_under_wlp_Skip wp_under_wlp_Seq healthy_intros) lemma wp_under_wlp_Bind: "\ \s. wp_under_wlp (a (f s)) \ \ wp_under_wlp (Bind f a)" unfolding wp_under_wlp_def by(auto simp:wp_eval) lemmas wp_under_wlp_intros = wp_under_wlp_Abort wp_under_wlp_Skip wp_under_wlp_Apply wp_under_wlp_Seq wp_under_wlp_PC wp_under_wlp_DC wp_under_wlp_SetPC wp_under_wlp_SetDC wp_under_wlp_Embed wp_under_wlp_loop wp_under_wlp_repeat wp_under_wlp_Bind subsection \Sub-Distributivity of Conjunction\ definition sub_distrib_pconj :: "'s prog \ bool" where "sub_distrib_pconj prog \ \P Q. unitary P \ unitary Q \ wlp prog P && wp prog Q \ wp prog (P && Q)" lemma sub_distrib_pconjI[intro]: "\\P Q. \ unitary P; unitary Q \ \ wlp prog P && wp prog Q \ wp prog (P && Q) \ \ sub_distrib_pconj prog" unfolding sub_distrib_pconj_def by(simp) lemma sub_distrib_pconjD[dest]: "\P Q. \ sub_distrib_pconj prog; unitary P; unitary Q \ \ wlp prog P && wp prog Q \ wp prog (P && Q)" unfolding sub_distrib_pconj_def by(simp) lemma sdp_Abort: "sub_distrib_pconj Abort" by(rule sub_distrib_pconjI, unfold wp_eval, auto intro:exp_conj_rzero) lemma sdp_Skip: "sub_distrib_pconj Skip" by(rule sub_distrib_pconjI, simp add:wp_eval) lemma sdp_Seq: fixes a and b assumes sdp_a: "sub_distrib_pconj a" and sdp_b: "sub_distrib_pconj b" and h_wp_a: "healthy (wp a)" and h_wp_b: "healthy (wp b)" and h_wlp_b: "nearly_healthy (wlp b)" shows "sub_distrib_pconj (a ;; b)" proof(rule sub_distrib_pconjI, unfold wp_eval o_def) fix P::"'a \ real" and Q::"'a \ real" assume uP: "unitary P" and uQ: "unitary Q" with h_wp_b and h_wlp_b have "wlp a (wlp b P) && wp a (wp b Q) \ wp a (wlp b P && wp b Q)" by(blast intro!:sub_distrib_pconjD[OF sdp_a]) also { from sdp_b and uP and uQ have "wlp b P && wp b Q \ wp b (P && Q)" by(blast) with h_wp_a h_wp_b h_wlp_b uP uQ have "wp a (wlp b P && wp b Q) \ wp a (wp b (P && Q))" by(blast intro!:mono_transD[OF healthy_monoD, OF h_wp_a] unitary_sound unitary_intros sound_intros) } finally show "wlp a (wlp b P) && wp a (wp b Q) \ wp a (wp b (P && Q))" . qed lemma sdp_Apply: "sub_distrib_pconj (Apply f)" by(rule sub_distrib_pconjI, simp add:wp_eval) lemma sdp_DC: fixes a::"'s prog" and b assumes sdp_a: "sub_distrib_pconj a" and sdp_b: "sub_distrib_pconj b" and h_wp_a: "healthy (wp a)" and h_wp_b: "healthy (wp b)" and h_wlp_b: "nearly_healthy (wlp b)" shows "sub_distrib_pconj (a \ b)" proof(rule sub_distrib_pconjI, unfold wp_eval, rule le_funI) fix P::"'s \ real" and Q::"'s \ real" and s::'s assume uP: "unitary P" and uQ: "unitary Q" have "((\s. min (wlp a P s) (wlp b P s)) && (\s. min (wp a Q s) (wp b Q s))) s \ min (wlp a P s .& wp a Q s) (wlp b P s .& wp b Q s)" unfolding exp_conj_def by(rule min_pconj) also { have "(\s. wlp a P s .& wp a Q s) = wlp a P && wp a Q" by(simp add:exp_conj_def) also from sdp_a uP uQ have "... \ wp a (P && Q)" by(blast dest:sub_distrib_pconjD) finally have "wlp a P s .& wp a Q s \ wp a (P && Q) s" by(rule le_funD) moreover { have "(\s. wlp b P s .& wp b Q s) = wlp b P && wp b Q" by(simp add:exp_conj_def) also from sdp_b uP uQ have "... \ wp b (P && Q)" by(blast) finally have "wlp b P s .& wp b Q s \ wp b (P && Q) s" by(rule le_funD) } ultimately have "min (wlp a P s .& wp a Q s) (wlp b P s .& wp b Q s) \ min (wp a (P && Q) s) (wp b (P && Q) s)" by(auto) } finally show "((\s. min (wlp a P s) (wlp b P s)) && (\s. min (wp a Q s) (wp b Q s))) s \ min (wp a (P && Q) s) (wp b (P && Q) s)" . qed lemma sdp_PC: fixes a::"'s prog" and b assumes sdp_a: "sub_distrib_pconj a" and sdp_b: "sub_distrib_pconj b" and h_wp_a: "healthy (wp a)" and h_wp_b: "healthy (wp b)" and h_wlp_b: "nearly_healthy (wlp b)" and uP: "unitary P" shows "sub_distrib_pconj (a \<^bsub>P\<^esub>\ b)" proof(rule sub_distrib_pconjI, unfold wp_eval, rule le_funI) fix Q::"'s \ real" and R::"'s \ real" and s::'s assume uQ: "unitary Q" and uR: "unitary R" have nnA: "0 \ P s" and nnB: "P s \ 1" - using uP by(auto simp:sign_simps) + using uP by auto note nn = nnA nnB have "((\s. P s * wlp a Q s + (1 - P s) * wlp b Q s) && (\s. P s * wp a R s + (1 - P s) * wp b R s)) s = ((P s * wlp a Q s + (1 - P s) * wlp b Q s) + (P s * wp a R s + (1 - P s) * wp b R s)) \ 1" by(simp add:exp_conj_def pconj_def) also have "... = P s * (wlp a Q s + wp a R s) + (1 - P s) * (wlp b Q s + wp b R s) \ 1" by(simp add:field_simps) also have "... = P s * (wlp a Q s + wp a R s) + (1 - P s) * (wlp b Q s + wp b R s) \ (P s + (1 - P s))" by(simp) also have "... \ (P s * (wlp a Q s + wp a R s) \ P s) + ((1 - P s) * (wlp b Q s + wp b R s) \ (1 - P s))" by(rule tminus_add_mono) also have "... = (P s * (wlp a Q s + wp a R s \ 1)) + ((1 - P s) * (wlp b Q s + wp b R s \ 1))" by(simp add:nn tminus_left_distrib) also have "... = P s * ((wlp a Q && wp a R) s) + (1 - P s) * ((wlp b Q && wp b R) s)" by(simp add:exp_conj_def pconj_def) also { from sdp_a sdp_b uQ uR have "P s * (wlp a Q && wp a R) s \ P s * wp a (Q && R) s" and "(1 - P s) * (wlp b Q && wp b R) s \ (1 - P s) * wp b (Q && R) s" by (simp_all add: entailsD mult_left_mono nn sub_distrib_pconjD) hence "P s * ((wlp a Q && wp a R) s) + (1 - P s) * ((wlp b Q && wp b R) s) \ P s * wp a (Q && R) s + (1 - P s) * wp b (Q && R) s" by(auto) } finally show "((\s. P s * wlp a Q s + (1 - P s) * wlp b Q s) && (\s. P s * wp a R s + (1 - P s) * wp b R s)) s \ P s * wp a (Q && R) s + (1 - P s) * wp b (Q && R) s" . qed lemma sdp_Embed: "\ \P Q. \ unitary P; unitary Q \ \ t P && t Q \ t (P && Q) \ \ sub_distrib_pconj (Embed t)" by(auto simp:wp_eval) lemma sdp_repeat: fixes a::"'s prog" assumes sdpa: "sub_distrib_pconj a" and hwp: "healthy (wp a)" and hwlp: "nearly_healthy (wlp a)" shows "sub_distrib_pconj (repeat n a)" (is "?X n") proof(induct n) show "?X 0" by(simp add:sdp_Skip) fix n assume IH: "?X n" show "?X (Suc n)" proof(rule sub_distrib_pconjI, simp add:wp_eval) fix P::"'s \ real" and Q::"'s \ real" assume uP: "unitary P" and uQ: "unitary Q" from assms have hwlpa: "nearly_healthy (wlp (repeat n a))" and hwpa: "healthy (wp (repeat n a))" by(auto intro:healthy_intros) from uP and hwlpa have "unitary (wlp (repeat n a) P)" by(blast) moreover from uQ and hwpa have "unitary (wp (repeat n a) Q)" by(blast) ultimately have "wlp a (wlp (repeat n a) P) && wp a (wp (repeat n a) Q) \ wp a (wlp (repeat n a) P && wp (repeat n a) Q)" using sdpa by(blast) also { from hwlp have "nearly_healthy (wlp (repeat n a))" by(rule healthy_intros) with uP have "sound (wlp (repeat n a) P)" by(auto) moreover from hwp uQ have "sound (wp (repeat n a) Q)" by(auto intro:healthy_intros) ultimately have "sound (wlp (repeat n a) P && wp (repeat n a) Q)" by(rule exp_conj_sound) moreover { from uP uQ have "sound (P && Q)" by(auto intro:exp_conj_sound) with hwp have "sound (wp (repeat n a) (P && Q))" by(auto intro:healthy_intros) } moreover from uP uQ IH have "wlp (repeat n a) P && wp (repeat n a) Q \ wp (repeat n a) (P && Q)" by(blast) ultimately have "wp a (wlp (repeat n a) P && wp (repeat n a) Q) \ wp a (wp (repeat n a) (P && Q))" by(rule mono_transD[OF healthy_monoD, OF hwp]) } finally show "wlp a (wlp (repeat n a) P) && wp a (wp (repeat n a) Q) \ wp a (wp (repeat n a) (P && Q))" . qed qed lemma sdp_SetPC: fixes p::"'a \ 's prog" assumes sdp: "\s a. a \ supp (P s) \ sub_distrib_pconj (p a)" and fin: "\s. finite (supp (P s))" and nnp: "\s a. 0 \ P s a" and sub: "\s. sum (P s) (supp (P s)) \ 1" shows "sub_distrib_pconj (SetPC p P)" proof(rule sub_distrib_pconjI, simp add:wp_eval, rule le_funI) fix Q::"'s \ real" and R::"'s \ real" and s::'s assume uQ: "unitary Q" and uR: "unitary R" have "((\s. \a\supp (P s). P s a * wlp (p a) Q s) && (\s. \a\supp (P s). P s a * wp (p a) R s)) s = (\a\supp (P s). P s a * wlp (p a) Q s) + (\a\supp (P s). P s a * wp (p a) R s) \ 1" by(simp add:exp_conj_def pconj_def) also have "... = (\a\supp (P s). P s a * (wlp (p a) Q s + wp (p a) R s)) \ 1" by(simp add: sum.distrib field_simps) also from sub have "... \ (\a\supp (P s). P s a * (wlp (p a) Q s + wp (p a) R s)) \ (\a\supp (P s). P s a)" by(rule tminus_right_antimono) also from fin have "... \ (\a\supp (P s). P s a * (wlp (p a) Q s + wp (p a) R s) \ P s a)" by(rule tminus_sum_mono) also from nnp have "... = (\a\supp (P s). P s a * (wlp (p a) Q s + wp (p a) R s \ 1))" by(simp add:tminus_left_distrib) also have "... = (\a\supp (P s). P s a * (wlp (p a) Q && wp (p a) R) s)" by(simp add:pconj_def exp_conj_def) also { from sdp uQ uR have "\a. a \ supp (P s) \ wlp (p a) Q && wp (p a) R \ wp (p a) (Q && R)" by(blast intro:sub_distrib_pconjD) with nnp have "(\a\supp (P s). P s a * (wlp (p a) Q && wp (p a) R) s) \ (\a\supp (P s). P s a * (wp (p a) (Q && R)) s)" by(blast intro:sum_mono mult_left_mono) } finally show "((\s. \a\supp (P s). P s a * wlp (p a) Q s) && (\s. \a\supp (P s). P s a * wp (p a) R s)) s \ (\a\supp (P s). P s a * wp (p a) (Q && R) s)" . qed lemma sdp_SetDC: fixes p::"'a \ 's prog" assumes sdp: "\s a. a \ S s \ sub_distrib_pconj (p a)" and hwp: "\s a. a \ S s \ healthy (wp (p a))" and hwlp: "\s a. a \ S s \ nearly_healthy (wlp (p a))" and ne: "\s. S s \ {}" shows "sub_distrib_pconj (SetDC p S)" proof(rule sub_distrib_pconjI, rule le_funI) fix P::"'s \ real" and Q::"'s \ real" and s::'s assume uP: "unitary P" and uQ: "unitary Q" from uP hwlp have "\x. x \ (\a. wlp (p a) P) ` S s \ unitary x" by(auto) hence "\y. y \ (\a. wlp (p a) P s) ` S s \ 0 \ y" by(auto) hence "\a. a \ S s \ wlp (SetDC p S) P s \ wlp (p a) P s" unfolding wp_eval by(intro cInf_lower bdd_belowI, auto) moreover { from uQ hwp have "\a. a \ S s \ 0 \ wp (p a) Q s" by(blast) hence "\a. a \ S s \ wp (SetDC p S) Q s \ wp (p a) Q s" unfolding wp_eval by(intro cInf_lower bdd_belowI, auto) } ultimately have "\a. a \ S s \ wlp (SetDC p S) P s + wp (SetDC p S) Q s \ 1 \ wlp (p a) P s + wp (p a) Q s \ 1" by(auto intro:tminus_left_mono add_mono) also have "\a. wlp (p a) P s + wp (p a) Q s \ 1 = (wlp (p a) P && wp (p a) Q) s" by(simp add:exp_conj_def pconj_def) also from sdp uP uQ have "\a. a \ S s \ ... a \ wp (p a) (P && Q) s" by(blast) also have "\a. ... a = wp (p a) (\s. P s + Q s \ 1) s" by(simp add:exp_conj_def pconj_def) finally show "(wlp (SetDC p S) P && wp (SetDC p S) Q) s \ wp (SetDC p S) (P && Q) s" unfolding exp_conj_def pconj_def wp_eval using ne by(blast intro!:cInf_greatest) qed lemma sdp_Bind: "\ \s. sub_distrib_pconj (p (f s)) \ \ sub_distrib_pconj (Bind f p)" unfolding sub_distrib_pconj_def wp_eval exp_conj_def pconj_def by(blast) text \For loops, we again appeal to our transfinite induction principle, this time taking advantage of the simultaneous treatment of both strict and liberal transformers.\ lemma sdp_loop: fixes body::"'s prog" assumes sdp_body: "sub_distrib_pconj body" and hwlp: "nearly_healthy (wlp body)" and hwp: "healthy (wp body)" shows "sub_distrib_pconj (do G \ body od)" proof(rule sub_distrib_pconjI, rule loop_induct[OF hwp hwlp]) fix P Q::"'s expect" and S::"('s trans \ 's trans) set" assume uP: "unitary P" and uQ: "unitary Q" and ffst: "\x\S. feasible (fst x)" and usnd: "\x\S. \Q. unitary Q \ unitary (snd x Q)" and IH: "\x\S. snd x P && fst x Q \ fst x (P && Q)" show "Inf_utrans (snd ` S) P && Sup_trans (fst ` S) Q \ Sup_trans (fst ` S) (P && Q)" proof(cases) assume "S = {}" thus ?thesis by(simp add:Inf_trans_def Sup_trans_def Inf_utrans_def Inf_exp_def Sup_exp_def exp_conj_def) next assume ne: "S \ {}" let "?f s" = "1 + Sup_trans (fst ` S) (P && Q) s - Inf_utrans (snd ` S) P s" from ne obtain t where tin: "t \ fst ` S" by(auto) from ne obtain u where uin: "u \ snd ` S" by(auto) from tin ffst uP uQ have utPQ: "unitary (t (P && Q))" by(auto intro:exp_conj_unitary) hence "\s. 0 \ t (P && Q) s" by(auto) also { from ffst tin have le: "le_utrans t (Sup_trans (fst ` S))" by(auto intro:Sup_trans_upper) with uP uQ have "\s. t (P && Q) s \ Sup_trans (fst ` S) (P && Q) s" by(auto intro:exp_conj_unitary) } finally have nn_rhs: "\s. 0 \ Sup_trans (fst ` S) (P && Q) s" . have "\R. Inf_utrans (snd ` S) P && R \ Sup_trans (fst ` S) (P && Q) \ R \ ?f" proof(rule contrapos_pp, assumption) fix R assume "\ R \ ?f" then obtain s where "\ R s \ ?f s" by(auto) hence gt: "?f s < R s" by(simp) from nn_rhs have g1: "1 \ 1 + Sup_trans (fst ` S) (P && Q) s" by(auto) hence "Sup_trans (fst ` S) (P && Q) s = Inf_utrans (snd ` S) P s .& ?f s" by(simp add:pconj_def) also from g1 have "... = Inf_utrans (snd ` S) P s + ?f s - 1" by(simp) also from gt have "... < Inf_utrans (snd ` S) P s + R s - 1" by(simp) also { with g1 have "1 \ Inf_utrans (snd ` S) P s + R s" by(simp) hence "Inf_utrans (snd ` S) P s + R s - 1 = Inf_utrans (snd ` S) P s .& R s" by(simp add:pconj_def) } finally have "\ (Inf_utrans (snd ` S) P && R) s \ Sup_trans (fst ` S) (P && Q) s" by(simp add:exp_conj_def) thus "\ Inf_utrans (snd ` S) P && R \ Sup_trans (fst ` S) (P && Q)" by(auto) qed moreover have "\t\fst ` S. Inf_utrans (snd ` S) P && t Q \ Sup_trans (fst ` S) (P && Q)" proof fix t assume tin: "t \ fst ` S" then obtain x where xin: "x \ S" and fx: "t = fst x" by(auto) from xin have "snd x \ snd ` S" by(auto) with uP usnd have "Inf_utrans (snd ` S) P \ snd x P" by(auto intro:le_utransD[OF Inf_utrans_lower]) hence "Inf_utrans (snd ` S) P && fst x Q \ snd x P && fst x Q" by(auto intro:entails_frame) also from xin IH have "... \ fst x (P && Q)" by(auto) also from xin ffst exp_conj_unitary[OF uP uQ] have "... \ Sup_trans (fst ` S) (P && Q)" by(auto intro:le_utransD[OF Sup_trans_upper]) finally show "Inf_utrans (snd ` S) P && t Q \ Sup_trans (fst ` S) (P && Q)" by(simp add:fx) qed ultimately have bt: "\t\fst ` S. t Q \ ?f" by(blast) have "Sup_trans (fst ` S) Q = Sup_exp {t Q |t. t \ fst ` S}" by(simp add:Sup_trans_def) also have "... \ ?f" proof(rule Sup_exp_least) from bt show " \R\{t Q |t. t \ fst ` S}. R \ ?f" by(blast) from ne obtain t where tin: "t \ fst ` S" by(auto) with ffst uQ have "unitary (t Q)" by(auto) hence "\s. 0 \ t Q" by(auto) also from tin bt have "... \ ?f" by(auto) finally show "nneg (\s. 1 + Sup_trans (fst ` S) (P && Q) s - Inf_utrans (snd ` S) P s)" by(auto) qed finally have "Inf_utrans (snd ` S) P && Sup_trans (fst ` S) Q \ Inf_utrans (snd ` S) P && ?f" by(auto intro:entails_frame) also from nn_rhs have "... \ Sup_trans (fst ` S) (P && Q)" by(simp add:exp_conj_def pconj_def) finally show ?thesis . qed next fix P Q::"'s expect" and t u::"'s trans" assume uP: "unitary P" and uQ: "unitary Q" and ft: "feasible t" and uu: "\Q. unitary Q \ unitary (u Q)" and IH: "u P && t Q \ t (P && Q)" show "wlp (body ;; Embed u \<^bsub>\ G \\<^esub>\ Skip) P && wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) Q \ wp (body ;; Embed t \<^bsub>\ G \\<^esub>\ Skip) (P && Q)" proof(rule le_funI, simp add:wp_eval exp_conj_def pconj_def) fix s::'s have "\ G \ s * wlp body (u P) s + (1 - \ G \ s) * P s + (\ G \ s * wp body (t Q) s + (1 - \ G \ s) * Q s) \ 1 = (\ G \ s * wlp body (u P) s + \ G \ s * wp body (t Q) s) + ((1 - \ G \ s) * P s + (1 - \ G \ s) * Q s) \ (\G\ s + (1 - \G\ s))" by(simp add:ac_simps) also have "... \ (\ G \ s * wlp body (u P) s + \ G \ s * wp body (t Q) s \ \G\ s) + ((1 - \ G \ s) * P s + (1 - \ G \ s) * Q s \ (1 - \G\ s))" by(rule tminus_add_mono) also have "... = \ G \ s * (wlp body (u P) s + wp body (t Q) s \ 1) + (1 - \ G \ s) * (P s + Q s \ 1)" by(simp add:tminus_left_distrib distrib_left) also { from uP uQ ft uu have "wlp body (u P) && wp body (t Q) \ wp body (u P && t Q)" by(auto intro:sub_distrib_pconjD[OF sdp_body]) also from IH unitary_sound[OF uP] unitary_sound[OF uQ] ft unitary_sound[OF uu[OF uP]] have "\ \ wp body (t (P && Q))" by(blast intro!:mono_transD[OF healthy_monoD, OF hwp] exp_conj_sound) finally have "wlp body (u P) s + wp body (t Q) s \ 1 \ wp body (t (\s. P s + Q s \ 1)) s" by(auto simp:exp_conj_def pconj_def) hence "\ G \ s * (wlp body (u P) s + wp body (t Q) s \ 1) + (1 - \ G \ s) * (P s + Q s \ 1) \ \ G \ s * wp body (t (\s. P s + Q s \ 1)) s + (1 - \ G \ s) * (P s + Q s \ 1)" by(auto intro:add_right_mono mult_left_mono) } finally show "\ G \ s * wlp body (u P) s + (1 - \ G \ s) * P s + (\ G \ s * wp body (t Q) s + (1 - \ G \ s) * Q s) \ 1 \ \ G \ s * wp body (t (\s. P s + Q s \ 1)) s + (1 - \ G \ s) * (P s + Q s \ 1)" . qed next fix P Q::"'s expect" and t t' u u'::"'s trans" assume "unitary P" "unitary Q" "equiv_trans t t'" "equiv_utrans u u'" "u P && t Q \ t (P && Q)" thus "u' P && t' Q \ t' (P && Q)" by(simp add:equiv_transD unitary_sound equiv_utransD exp_conj_unitary) qed lemmas sdp_intros = sdp_Abort sdp_Skip sdp_Apply sdp_Seq sdp_DC sdp_PC sdp_SetPC sdp_SetDC sdp_Embed sdp_repeat sdp_Bind sdp_loop subsection \The Well-Defined Predicate.\ definition well_def :: "'s prog \ bool" where "well_def prog \ healthy (wp prog) \ nearly_healthy (wlp prog) \ wp_under_wlp prog \ sub_distrib_pconj prog \ sublinear (wp prog) \ bd_cts (wp prog)" lemma well_defI[intro]: "\ healthy (wp prog); nearly_healthy (wlp prog); wp_under_wlp prog; sub_distrib_pconj prog; sublinear (wp prog); bd_cts (wp prog) \ \ well_def prog" unfolding well_def_def by(simp) lemma well_def_wp_healthy[dest]: "well_def prog \ healthy (wp prog)" unfolding well_def_def by(simp) lemma well_def_wlp_nearly_healthy[dest]: "well_def prog \ nearly_healthy (wlp prog)" unfolding well_def_def by(simp) lemma well_def_wp_under[dest]: "well_def prog \ wp_under_wlp prog" unfolding well_def_def by(simp) lemma well_def_sdp[dest]: "well_def prog \ sub_distrib_pconj prog" unfolding well_def_def by(simp) lemma well_def_wp_sublinear[dest]: "well_def prog \ sublinear (wp prog)" unfolding well_def_def by(simp) lemma well_def_wp_cts[dest]: "well_def prog \ bd_cts (wp prog)" unfolding well_def_def by(simp) lemmas wd_dests = well_def_wp_healthy well_def_wlp_nearly_healthy well_def_wp_under well_def_sdp well_def_wp_sublinear well_def_wp_cts lemma wd_Abort: "well_def Abort" by(blast intro:healthy_wp_Abort nearly_healthy_wlp_Abort wp_under_wlp_Abort sdp_Abort sublinear_wp_Abort cts_wp_Abort) lemma wd_Skip: "well_def Skip" by(blast intro:healthy_wp_Skip nearly_healthy_wlp_Skip wp_under_wlp_Skip sdp_Skip sublinear_wp_Skip cts_wp_Skip) lemma wd_Apply: "well_def (Apply f)" by(blast intro:healthy_wp_Apply nearly_healthy_wlp_Apply wp_under_wlp_Apply sdp_Apply sublinear_wp_Apply cts_wp_Apply) lemma wd_Seq: "\ well_def a; well_def b \ \ well_def (a ;; b)" by(blast intro:healthy_wp_Seq nearly_healthy_wlp_Seq wp_under_wlp_Seq sdp_Seq sublinear_wp_Seq cts_wp_Seq) lemma wd_PC: "\ well_def a; well_def b; unitary P \ \ well_def (a \<^bsub>P\<^esub>\ b)" by(blast intro:healthy_wp_PC nearly_healthy_wlp_PC wp_under_wlp_PC sdp_PC sublinear_wp_PC cts_wp_PC) lemma wd_DC: "\ well_def a; well_def b \ \ well_def (a \ b)" by(blast intro:healthy_wp_DC nearly_healthy_wlp_DC wp_under_wlp_DC sdp_DC sublinear_wp_DC cts_wp_DC) lemma wd_SetDC: "\ \x s. x \ S s \ well_def (a x); \s. S s \ {}; \s. finite (S s) \ \ well_def (SetDC a S)" by (simp add: cts_wp_SetDC ex_in_conv healthy_intros(17) healthy_intros(18) sdp_intros(8) sublinear_intros(8) well_def_def wp_under_wlp_intros(8)) lemma wd_SetPC: "\ \x s. x \ (supp (p s)) \ well_def (a x); \s. unitary (p s); \s. finite (supp (p s)); \s. sum (p s) (supp (p s)) \ 1 \ \ well_def (SetPC a p)" by(iprover intro!:well_defI healthy_wp_SetPC nearly_healthy_wlp_SetPC wp_under_wlp_SetPC sdp_SetPC sublinear_wp_SetPC cts_wp_SetPC dest:wd_dests unitary_sound sound_nneg[OF unitary_sound] nnegD) lemma wd_Embed: fixes t::"'s trans" assumes ht: "healthy t" and st: "sublinear t" and ct: "bd_cts t" shows "well_def (Embed t)" proof(intro well_defI) from ht show "healthy (wp (Embed t))" "nearly_healthy (wlp (Embed t))" by(simp add:wp_def wlp_def Embed_def healthy_nearly_healthy)+ from st show "sublinear (wp (Embed t))" by(simp add:wp_def Embed_def) show "wp_under_wlp (Embed t)" by(simp add:wp_under_wlp_def wp_eval) show "sub_distrib_pconj (Embed t)" by(rule sub_distrib_pconjI, auto intro:le_funI[OF sublinearD[OF st, where a=1 and b=1 and c=1, simplified]] simp:exp_conj_def pconj_def wp_def wlp_def Embed_def) from ct show "bd_cts (wp (Embed t))" by(simp add:wp_def Embed_def) qed lemma wd_repeat: "well_def a \ well_def (repeat n a)" by(blast intro:healthy_wp_repeat nearly_healthy_wlp_repeat wp_under_wlp_repeat sdp_repeat sublinear_wp_repeat cts_wp_repeat) lemma wd_Bind: "\ \s. well_def (a (f s)) \ \ well_def (Bind f a)" by(blast intro:healthy_wp_Bind nearly_healthy_wlp_Bind wp_under_wlp_Bind sdp_Bind sublinear_wp_Bind cts_wp_Bind) lemma wd_loop: "well_def body \ well_def (do G \ body od)" by(blast intro:healthy_wp_loop nearly_healthy_wlp_loop wp_under_wlp_loop sdp_loop sublinear_wp_loop cts_wp_loop) lemmas wd_intros = wd_Abort wd_Skip wd_Apply wd_Embed wd_Seq wd_PC wd_DC wd_SetPC wd_SetDC wd_Bind wd_repeat wd_loop end