diff --git a/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy b/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy --- a/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy +++ b/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy @@ -1,955 +1,955 @@ -(* - File: Arith_Prog_Rel_Primes.thy - Author: Jose Manuel Rodriguez Caballero, University of Tartu -*) -section \Problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002)\ -theory Arith_Prog_Rel_Primes - imports - Complex_Main - "HOL-Number_Theory.Number_Theory" -begin - -text \ - Statement of the problem (from ~\cite{putnam}): For which integers $n>1$ does the set of positive - integers less than and relatively prime to $n$ constitute an arithmetic progression? - - The solution of the above problem is theorem @{text arith_prog_rel_primes_solution}. - - First, we will require some auxiliary material before we get started with the actual - solution. -\ - -subsection \Auxiliary results\ - -lemma even_and_odd_parts: - fixes n::nat - assumes \n \ 0\ - shows \\ k q::nat. n = (2::nat)^k*q \ odd q\ -proof- - have \prime (2::nat)\ - by simp - thus ?thesis - using prime_power_canonical[where p = "2" and m = "n"] - assms semiring_normalization_rules(7) by auto -qed - -lemma only_one_odd_div_power2: - fixes n::nat - assumes \n \ 0\ and \\ x. x dvd n \ odd x \ x = 1\ - shows \\ k. n = (2::nat)^k\ - by (metis even_and_odd_parts assms(1) assms(2) dvd_triv_left semiring_normalization_rules(11) - semiring_normalization_rules(7)) - -lemma coprime_power2: - fixes n::nat - assumes \n \ 0\ and \\ x. x < n \ (coprime x n \ odd x)\ - shows \\ k. n = (2::nat)^k\ -proof- - have \x dvd n \ odd x \ x = 1\ - for x - by (metis neq0_conv One_nat_def Suc_1 Suc_lessI assms(1) assms(2) coprime_left_2_iff_odd - dvd_refl linorder_neqE_nat nat_dvd_1_iff_1 nat_dvd_not_less not_coprimeI) - thus ?thesis - using assms(1) only_one_odd_div_power2 - by auto -qed - -subsection \Main result\ - -text \ - The solution to the problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002) -\ - -theorem arith_prog_rel_primes_solution: - fixes n :: nat - assumes \n > 1\ - shows \(prime n \ (\ k. n = 2^k) \ n = 6) \ -(\ a b m. m \ 0 \ {x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ -proof- - have \ (prime n \ (\ k. n = 2^k) \ n = 6) \ - (\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m})\ - proof - show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - if "prime n \ (\k. n = 2 ^ k) \ n = 6" - proof- - have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - if "prime n" - proof- - have \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ - proof- - have \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ - proof - show "{x |x. x < n \ coprime x n} \ {x |x. x \ 0 \ x < n}" - by (smt Collect_mono not_le ord_0_nat ord_eq_0 order_refl prime_gt_1_nat that zero_neq_one) - show "{x |x. x \ 0 \ x < n} \ {x |x. x < n \ coprime x n}" - using coprime_commute prime_nat_iff'' that - by fastforce - qed - obtain m where \m+1 = n\ - using add.commute assms less_imp_add_positive by blast - have \{1+j| j::nat. j < (m::nat)} = {x | x :: nat. x \ 0 \ x < m+1}\ - by (metis Suc_eq_plus1 \m + 1 = n\ gr0_implies_Suc le_simps(3) less_nat_zero_code linorder_not_less nat.simps(3) nat_neq_iff plus_1_eq_Suc ) - hence \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < (m::nat)}\ - using \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ \m+1 = n\ - by auto - from \n > 1\ have \m \ 0\ - using \m + 1 = n\ by linarith - have \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ - using Suc_eq_plus1 \1 < n\ \{x |x. x < n \ coprime x n} = {1 + j |j. j < m}\ - by auto - hence \(\ m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m})\ - using \m \ 0\ - by blast - thus ?thesis by blast - qed - hence \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*1| j::nat. j < m}\ - by auto - thus ?thesis - by blast - qed - moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - if "\k. n = 2 ^ k" - proof- - obtain k where \n = 2 ^ k\ - using \\k. n = 2 ^ k\ by auto - have \k \ 0\ - by (metis \1 < n\ \n = 2 ^ k\ nat_less_le power.simps(1)) - obtain t where \Suc t = k\ - by (metis \k \ 0\ fib.cases) - have \n = 2^(Suc t)\ - by (simp add: \Suc t = k\ \n = 2 ^ k\) - have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < 2^t}\ - proof - show "{x |x. x < n \ coprime x n} \ {1 + j * 2 |j. j < 2^t}" - proof - fix x - assume \x \ {x |x. x < n \ coprime x n}\ - hence \x < n\ - by blast - have \coprime x n\ - using \x \ {x |x. x < n \ coprime x n}\ - by blast - hence \coprime x (2^(Suc k))\ - by (simp add: \k \ 0\ \n = 2 ^ k\) - have \odd x\ - using \coprime x n\ \k \ 0\ \n = 2 ^ k\ - by auto - then obtain j where \x = 1+j*2\ - by (metis add.commute add.left_commute left_add_twice mult_2_right oddE) - have \x < 2^k\ - using \n = 2 ^ k\ \x < n\ \x = 1+j*2\ - by linarith - hence \1+j*2 < 2^k\ - using \x = 1+j*2\ - by blast - hence \j < 2^t\ - using \Suc t = k\ by auto - thus \x \ {1 + j * 2 |j. j < 2^t}\ - using \x = 1+j*2\ - by blast - qed - show "{1 + j * 2 |j. j < 2 ^ t} \ {x |x. x < n \ coprime x n}" - proof - fix x::nat - assume \x \ {1 + j * 2 |j. j < 2 ^ t}\ - then obtain j where \x = 1 + j * 2\ and \j < 2 ^ t\ - by blast - have \x < 2*(2^t)\ - using \x = 1 + j * 2\ \j < 2 ^ t\ - by linarith - hence \x < n\ - by (simp add: \n = 2 ^ Suc t\) - moreover have \coprime x n\ - by (metis (no_types) \\thesis. (\j. \x = 1 + j * 2; j < 2 ^ t\ \ thesis) \ thesis\ \n = 2 ^ k\ coprime_Suc_left_nat coprime_mult_right_iff coprime_power_right_iff plus_1_eq_Suc) - ultimately show \x \ {x |x. x < n \ coprime x n}\ - by blast - qed - qed - have \(2::nat)^(t::nat) \ 0\ - by simp - obtain m where \m = (2::nat)^t\ by blast - have \m \ 0\ - using \m = 2 ^ t\ - by auto - have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ - using \m = 2 ^ t\ \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < 2 ^ t}\ - by auto - from \m \ 0\ \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ - show ?thesis by blast - qed - moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - if "n = 6" - proof- - have \{x | x. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < 2}\ - proof- - have \{x | x::nat. x < 6 \ coprime x 6} = {1, 5}\ - proof - show "{u. \x. u = (x::nat) \ x < 6 \ coprime x 6} \ {1, 5}" - proof - fix u::nat - assume \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ - hence \coprime u 6\ - by blast - have \u < 6\ - using \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ - by blast - moreover have \u \ 0\ - using \coprime u 6\ ord_eq_0 - by fastforce - moreover have \u \ 2\ - using \coprime u 6\ - by auto - moreover have \u \ 3\ - proof- - have \gcd (3::nat) 6 = 3\ - by auto - thus ?thesis - by (metis (no_types) \coprime u 6\ \gcd 3 6 = 3\ coprime_iff_gcd_eq_1 - numeral_eq_one_iff semiring_norm(86)) - qed - moreover have \u \ 4\ - proof- - have \gcd (4::nat) 6 = 2\ - by (metis (no_types, lifting) add_numeral_left gcd_add1 gcd_add2 gcd_nat.idem - numeral_Bit0 numeral_One one_plus_numeral semiring_norm(4) semiring_norm(5)) - thus ?thesis - using \coprime u 6\ coprime_iff_gcd_eq_1 - by auto - qed - ultimately have \u = 1 \ u = 5\ - by auto - thus \u \ {1, 5}\ - by blast - qed - show "{1::nat, 5} \ {x |x. x < 6 \ coprime x 6}" - proof- - have \(1::nat) \ {x |x. x < 6 \ coprime x 6}\ - by simp - moreover have \(5::nat) \ {x |x. x < 6 \ coprime x 6}\ - by (metis Suc_numeral coprime_Suc_right_nat less_add_one mem_Collect_eq - numeral_plus_one semiring_norm(5) semiring_norm(8)) - ultimately show ?thesis - by blast - qed - qed - moreover have \{1+j*4| j::nat. j < 2} = {1, 5}\ - by auto - ultimately show ?thesis by auto - qed - moreover have \(2::nat) \ 0\ - by simp - ultimately have \\ m. m \ 0 \ {x | x :: nat. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < m}\ - by blast - thus ?thesis - using that - by auto - qed - ultimately show ?thesis - using that - by blast - qed - show "prime n \ (\k. n = 2 ^ k) \ n = 6" - if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - proof- - obtain b m where \m \ 0\ and \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - using \\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by auto - show ?thesis - proof(cases \n = 2\) - case True - thus ?thesis - by auto - next - case False - have \b \ 4\ - proof(cases \odd b\) - case True - show ?thesis - proof(rule classical) - assume \\(b \ 4)\ - hence \b > 4\ - using le_less_linear - by blast - obtain m where \m \ 0\ - and \{x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ - using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - have \b \ 0\ - using \4 < b\ - by linarith - have \n = 2 + (m-1)*b\ - proof- - have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1+j*b| j::nat. j < m}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ - \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ - by blast - have \i \ m-1\ - using \i < m\ - by linarith - have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ - using \m \ 0\ - by auto - hence \1 + (m-1)*b \ {x | x::nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \1 + (m-1)*b < n\ - by blast - hence \1 + (m-1)*b \ n-1\ - by linarith - hence \1 + (m-1)*b \ 1+i*b\ - using \n - 1 = 1 + i * b\ - by linarith - hence \(m-1)*b \ i*b\ - by linarith - hence \m-1 \ i\ - using \b \ 0\ - by auto - hence \m-1 = i\ - using \i \ m - 1\ le_antisym - by blast - thus ?thesis - using \m \ 0\ \n - 1 = 1 + i * b\ - by auto - qed - have \m \ 2\ - using \n = 2 + (m - 1)*b\ \n \ 2\ - by auto - hence \1+b \ {1+j*b| j. j < m}\ - by fastforce - hence \1+b \ {x | x::nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \coprime (1+b) n\ - by blast - have \(2::nat) dvd (1+b)\ - using \odd b\ - by simp - hence \coprime (2::nat) n\ - using \coprime (1 + b) n\ coprime_common_divisor coprime_left_2_iff_odd odd_one - by blast - have \(2::nat) < n\ - using \1 < n\ \n \ 2\ - by linarith - have \2 \ {x | x :: nat. x < n \ coprime x n}\ - using \2 < n\ \coprime 2 n\ - by blast - hence \2 \ {1+j*b| j::nat. j < m}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain j::nat where \2 = 1+j*b\ - by blast - have \1 = j*b\ - using \2 = 1+j*b\ - by linarith - thus ?thesis - by simp - qed - next - case False - hence \even b\ - by simp - show ?thesis - proof(rule classical) - assume \\(b \ 4)\ - hence \b > 4\ - using le_less_linear - by blast - obtain m where \ m \ 0\ - and \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ - using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - have \b \ 0\ - using \4 < b\ - by linarith - have \n = 2 + (m-1)*b\ - proof- - have \n-1 \ {x | x::nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1+j*b| j::nat. j < m}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ - \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ - by blast - have \i \ m-1\ - using \i < m\ - by linarith - have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ - using \m \ 0\ - by auto - hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \1 + (m-1)*b < n\ - by blast - hence \1 + (m-1)*b \ n-1\ - by linarith - hence \1 + (m-1)*b \ 1+i*b\ - using \n - 1 = 1 + i * b\ - by linarith - hence \(m-1)*b \ i*b\ - by linarith - hence \m-1 \ i\ - using \b \ 0\ - by auto - hence \m-1 = i\ - using \i \ m - 1\ le_antisym - by blast - thus ?thesis - using \m \ 0\ \n - 1 = 1 + i * b\ - by auto - qed - obtain k :: nat where \b = 2*k\ - using \even b\ - by blast - have \n = 2*(1 + (m-1)*k)\ - using \n = 2 + (m-1)*b\ \b = 2*k\ - by simp - show ?thesis - proof (cases \odd m\) - case True - hence \odd m\ by blast - then obtain t::nat where \m-1 = 2*t\ - by (metis odd_two_times_div_two_nat) - have \n = 2*(1 + b*t)\ - using \m - 1 = 2 * t\ \n = 2 + (m - 1) * b\ - by auto - have \t < m\ - using \m - 1 = 2 * t\ \m \ 0\ - by linarith - have \1 + b*t \ {1+j*b| j::nat. j < m}\ - using \t < m\ - by auto - hence \1 + b*t \ {x | x::nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \coprime (1 + b*t) n\ - by auto - thus ?thesis - by (metis (no_types, lifting) \b = 2 * k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 * (1 + b * t)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_cancel_right_right coprime_mult_right_iff coprime_self mult_cancel_left mult_is_0 nat_dvd_1_iff_1) - next - case False - thus ?thesis - proof(cases \odd k\) - case True - hence \odd k\ - by blast - have \even (1 + (m - 1) * k)\ - by (simp add: False True \m \ 0\) - have \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ - by simp - have \coprime (2 + (m - 1) * k) n\ - using \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ \even (1 + (m - 1) * k)\ - \n = 2 * (1 + (m - 1) * k)\ coprime_common_divisor coprime_mult_right_iff - coprime_right_2_iff_odd odd_one - by blast - have \2 + (m - 1) * k < n\ - by (metis (no_types, lifting) \even (1 + (m - 1) * k)\ \n = 2 * (1 + (m - 1) * k)\ - add_gr_0 add_mono_thms_linordered_semiring(1) dvd_add_left_iff dvd_add_triv_left_iff dvd_imp_le le_add2 le_neq_implies_less less_numeral_extra(1) mult_2 odd_one) - have \2 + (m - 1) * k \ {x | x :: nat. x < n \ coprime x n}\ - using \2 + (m - 1) * k < n\ \coprime (2 + (m - 1) * k) n\ - by blast - hence \2 + (m - 1) * k \ {1 + j * b |j. j < m}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain j::nat where \2 + (m - 1) * k = 1 + j * b\ and \j < m\ - by blast - have \1 + (m - 1) * k = j * b\ - using \2 + (m - 1) * k = 1 + j * b\ - by simp - hence \1 + (m - 1) * k = j * (2*k)\ - using \b = 2 * k\ by blast - thus ?thesis - by (metis \b = 2 * k\ \even b\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ dvd_add_times_triv_right_iff dvd_antisym dvd_imp_le dvd_triv_right even_numeral mult_2 zero_less_numeral) - next - case False - hence \even k\ by auto - have \odd (1 + (m - 1) * k)\ - by (simp add: \even k\ ) - hence \coprime (3 + (m - 1) * k) (1 + (m - 1) * k)\ - by (smt add_numeral_left coprime_common_divisor coprime_right_2_iff_odd dvd_add_left_iff not_coprimeE numeral_Bit1 numeral_One numeral_plus_one one_add_one) - hence \coprime (3 + (m - 1) * k) n\ - by (metis \even k\ \n = 2 * (1 + (m - 1) * k)\ coprime_mult_right_iff coprime_right_2_iff_odd even_add even_mult_iff odd_numeral) - have \3 + (m - 1) * k < n\ - by (smt Groups.add_ac(2) \even k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_Suc_right add_cancel_right_right add_mono_thms_linordered_semiring(1) dvd_imp_le even_add even_mult_iff le_add2 le_neq_implies_less left_add_twice mult_2 neq0_conv numeral_Bit1 numeral_One odd_numeral one_add_one plus_1_eq_Suc) - have \3 + (m - 1) * k \ {x |x. x < n \ coprime x n}\ - using \3 + (m - 1) * k < n\ \coprime (3 + (m - 1) * k) n\ - by blast - hence \3 + (m - 1) * k \ {1 + j * b |j. j < m}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain j::nat where \3 + (m - 1) * k = 1 + j * b\ - by blast - have \2 + (m - 1) * k = j * b\ - using \3 + (m - 1) * k = 1 + j * b\ - by simp - hence \2 + (m - 1) * k = j * 2*k\ - by (simp add: \b = 2 * k\) - thus ?thesis - by (metis \4 < b\ \b = 2 * k\ \even k\ dvd_add_times_triv_right_iff dvd_antisym - dvd_triv_right mult_2 nat_neq_iff numeral_Bit0) - qed - qed - qed - qed - moreover have \b \ 3\ - proof (rule classical) - assume \\ (b \ 3)\ - hence \b = 3\ - by blast - obtain m where \m \ 0\ and - \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ - using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - have \b \ 0\ - by (simp add: \b = 3\) - have \n = 2 + (m-1)*b\ - proof- - have \n-1 \ {x | x::nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1+j*b| j::nat. j < m}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ - \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ - by blast - have \i \ m-1\ - using \i < m\ - by linarith - have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ - using \m \ 0\ - by auto - hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \1 + (m-1)*b < n\ - by blast - hence \1 + (m-1)*b \ n-1\ - by linarith - hence \1 + (m-1)*b \ 1+i*b\ - using \n - 1 = 1 + i * b\ - by linarith - hence \(m-1)*b \ i*b\ - by linarith - hence \m-1 \ i\ - using \b \ 0\ - by auto - hence \m-1 = i\ - using \i \ m - 1\ le_antisym - by blast - thus ?thesis - using \m \ 0\ \n - 1 = 1 + i * b\ - by auto - qed - have \n > 2\ - using \1 < n\ \n \ 2\ - by linarith - hence \ m \ 2 \ using \n = 2 + (m-1)*b\ \b = 3\ - by simp - have \4 \ {1+j*(b::nat)| j::nat. j < m}\ - using \2 \ m\ \b = 3\ - by force - hence \(4::nat) \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by auto - hence \coprime (4::nat) n\ - by blast - have \(2::nat) dvd 4\ - by auto - hence \coprime (2::nat) n\ - using \coprime (4::nat) n\ coprime_divisors dvd_refl - by blast - have \4 < n\ - using \4 \ {x |x. x < n \ coprime x n}\ - by blast - have \2 < (4::nat)\ - by auto - have \2 < n\ - by (simp add: \2 < n\) - hence \2 \ {x | x :: nat. x < n \ coprime x n}\ - using \coprime (2::nat) n\ - by blast - hence \2 \ {1+j*(b::nat)| j::nat. j < m}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain j::nat where \2 = 1+j*3\ - using \b = 3\ - by blast - from \2 = 1+j*3\ - have \1 = j*3\ - by auto - hence \3 dvd 1\ - by auto - thus ?thesis - using nat_dvd_1_iff_1 numeral_eq_one_iff - by blast - qed - ultimately have \b = 0 \ b = 1 \ b = 2 \ b = 4\ - by auto - moreover have \b = 0 \ \k. n = 2^k\ - proof- - assume \b = 0\ - have \{1 + j * b |j. j < m} = {1}\ - using \b = 0\ \m \ 0\ - by auto - hence \{x |x. x < n \ coprime x n} = {1}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \n = 2\ - proof- - have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1}\ - by blast - hence \n-1 = 1\ - by blast - hence \n = 2\ - by simp - thus ?thesis - by blast - qed - hence \n = 2^1\ - by auto - thus ?thesis - by blast - qed - moreover have \b = 1 \ prime n\ - proof- - assume \b = 1\ - have \x < n \ x \ 0 \ coprime x n\ - for x - proof- - assume \x < n\ and \x \ 0\ - have \{1+j| j::nat. j < m} = {x | x::nat. x < m+1 \ x \ 0}\ - by (metis (full_types) Suc_eq_plus1 add_mono1 less_Suc_eq_0_disj nat.simps(3) plus_1_eq_Suc ) - hence \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x < m+1 \ x \ 0}\ - using \b = 1\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by auto - have \coprime (n-1) n\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 < n\ - using \1 < n\ - by auto - have \n-1 \ {x |x. x < n \ coprime x n}\ - using \coprime (n - 1) n\ \n - 1 < n\ - by blast - have \n-1 \ m\ - by (metis (no_types, lifting) CollectD Suc_eq_plus1 Suc_less_eq2 \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ leD le_less_linear not_less_eq_eq ) - have \m \ {x | x :: nat. x < m+1 \ x \ 0}\ - using \m \ 0\ - by auto - have \m \ {x |x. x < n \ coprime x n} \ - using \m \ {x |x. x < m + 1 \ x \ 0}\ - \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ - by blast - have \m < n\ - using \m \ {x |x. x < n \ coprime x n}\ - by blast - have \m+1 = n\ - using \m < n\ \n - 1 \ m\ - by linarith - have \x \ {x | x :: nat. x < m+1 \ x \ 0}\ - using \m + 1 = n\ \x < n\ \x \ 0\ - by blast - hence \x \ {x |x. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ - by blast - thus ?thesis - by blast - qed - thus ?thesis - using assms coprime_commute nat_neq_iff prime_nat_iff'' by auto - qed - moreover have \b = 2 \ \ k. n = 2^k\ - proof- - assume \b = 2\ - have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ - using \b = 2\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by auto - have \x < n \ coprime x n \ odd x\ - for x::nat - proof- - assume \x < n\ - have \coprime x n \ odd x\ - proof- - assume \coprime x n\ - have \x \ {x | x :: nat. x < n \ coprime x n}\ - by (simp add: \coprime x n\ \x < n\) - hence \x \ {1+j*2| j::nat. j < m}\ - using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ - by blast - then obtain j where \x = 1+j*2\ - by blast - thus ?thesis - by simp - qed - moreover have \odd x \ coprime x n\ - proof- - assume \odd x\ - obtain j::nat where \x = 1+j*2\ - by (metis \odd x\ add.commute mult_2_right odd_two_times_div_two_succ one_add_one semiring_normalization_rules(16)) - have \j < (n-1)/2\ - using \x < n\ \x = 1 + j * 2\ - by linarith - have \n = 2*m\ - proof- - have \(2::nat) \ 0\ - by auto - have \n = 2+(m-1)*2\ - proof- - have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1+j*b| j::nat. j < m}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ - \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ - by blast - have \i \ m-1\ - using \i < m\ - by linarith - have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ - using \m \ 0\ by auto - hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \1 + (m-1)*b < n\ - by blast - hence \1 + (m-1)*b \ n-1\ - by linarith - hence \1 + (m-1)*b \ 1+i*b\ - using \n - 1 = 1 + i * b\ - by linarith - hence \(m-1)*b \ i*b\ - by linarith - hence \m-1 \ i\ - proof- - have \b \ 0\ - using \b = 2\ - by simp - thus ?thesis - using \(m - 1) * b \ i * b\ mult_le_cancel2 - by blast - qed - hence \m-1 = i\ - using \i \ m - 1\ le_antisym - by blast - thus ?thesis - using \m \ 0\ \n - 1 = 1 + i * b\ - by (simp add: \b = 2\) - qed - thus ?thesis - by (simp add: \m \ 0\ \n = 2 + (m - 1) * 2\ mult.commute mult_eq_if) - qed - hence \j < m\ - using \x < n\ \x = 1 + j * 2\ - by linarith - hence \x \ {1+j*2| j::nat. j < m}\ - using \x = 1 + j * 2\ - by blast - hence \x \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ - by blast - thus ?thesis - by blast - qed - ultimately show ?thesis - by blast - qed - thus ?thesis - using coprime_power2 assms - by auto - qed - moreover have \b = 4 \ n = 6\ - proof- - assume \b = 4\ - have \n = 2 \ n = 6\ - proof(rule classical) - assume \\ (n = 2 \ n = 6)\ - have \(4::nat) \ 0\ - by auto - have \n = 2+(m-1)*4\ - proof- - have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ - using \1 < n\ coprime_diff_one_left_nat - by auto - have \n-1 \ {1+j*b| j::nat. j < m}\ - using \n - 1 \ {x |x. x < n \ coprime x n}\ - \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ - by blast - have \i \ m-1\ - using \i < m\ - by linarith - have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ - using \m \ 0\ - by auto - hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ - using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - hence \1 + (m-1)*b < n\ - by blast - hence \1 + (m-1)*b \ n-1\ - by linarith - hence \1 + (m-1)*b \ 1+i*b\ - using \n - 1 = 1 + i * b\ - by linarith - hence \(m-1)*b \ i*b\ - by linarith - hence \m-1 \ i\ - proof- - have \b \ 0\ - using \b = 4\ by auto - thus ?thesis - using \(m - 1) * b \ i * b\ mult_le_cancel2 - by blast - qed - hence \m-1 = i\ - using \i \ m - 1\ le_antisym - by blast - thus ?thesis - using \m \ 0\ \n - 1 = 1 + i * b\ - by (simp add: \b = 4\) - qed - hence \n = 4*m - 2\ - by (simp add: \m \ 0\ mult.commute mult_eq_if) - have \m \ 3\ - using \\ (n = 2 \ n = 6)\ \n = 2 + (m - 1) * 4\ - by auto - hence \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ - by force - hence \9 \ {1+j*4| j::nat. j < 3}\ - by force - hence \9 \ {1+j*4| j::nat. j < m}\ - using \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ - by blast - hence \9 \ {x | x :: nat. x < n \ coprime x n}\ - using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by auto - hence \coprime (9::nat) n\ - by blast - have \(3::nat) dvd 9\ - by auto - have \coprime (3::nat) n\ using \coprime (9::nat) n\ \(3::nat) dvd 9\ - by (metis coprime_commute coprime_mult_right_iff dvd_def) - have \(3::nat) < n\ - by (metis One_nat_def Suc_lessI \1 < n\ \\ (n = 2 \ n = 6)\ \coprime 3 n\ - coprime_self numeral_2_eq_2 numeral_3_eq_3 less_numeral_extra(1) nat_dvd_not_less) - have \3 \ {x | x :: nat. x < n \ coprime x n}\ - using \3 < n\ \coprime 3 n\ - by blast - hence \(3::nat) \ {1+j*4| j::nat. j < m}\ - using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ - by blast - then obtain j::nat where \(3::nat) = 1 + j*4\ - by blast - have \2 = j*4\ - using numeral_3_eq_3 \(3::nat) = 1 + j*4\ - by linarith - hence \1 = j*2\ - by linarith - hence \even 1\ - by simp - thus ?thesis - using odd_one - by blast - qed - thus ?thesis - by (simp add: False) - qed - ultimately show ?thesis - by blast - qed - qed - qed - moreover have \(\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}) - \ (\ a b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ - proof - show "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" - if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - using that - by blast - show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" - if "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" - proof- - obtain a b m::nat where \m \ 0\ - and \{x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ - using \\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ - by auto - have \a = 1\ - proof- - have \{x | x :: nat. x < n \ coprime x n} = {(a::nat)+j*(b::nat)| j::nat. j < m} \ a = 1\ - proof- - have \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ - using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ - by auto - have \Min {x | x :: nat. x < n \ coprime x n} = 1\ - proof- - have \finite {x | x :: nat. x < n \ coprime x n}\ - by auto - have \{x | x :: nat. x < n \ coprime x n} \ {}\ - using \1 < n\ by auto - have \1 \ {x | x :: nat. x < n \ coprime x n}\ - using \1 < n\ - by auto - have \\ x. coprime x n \ x \ 1\ - using \1 < n\ le_less_linear - by fastforce - hence \\ x. x < n \ coprime x n \ x \ 1\ - by blast - hence \\ x \ {x | x :: nat. x < n \ coprime x n}. x \ 1\ - by blast - hence \Min {x | x :: nat. x < n \ coprime x n} \ 1\ - using \finite {x | x :: nat. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} \ {}\ - by auto - thus ?thesis - using Min_le \1 \ {x |x. x < n \ coprime x n}\ \finite {x |x. x < n \ coprime x n}\ - antisym by blast - qed - have \Min {a+j*b| j::nat. j < m} = a\ - proof - - have f1: "\n. a = a + n * b \ n < m" - using \m \ 0\ - by auto - have f2: "\n. 1 = a + n * b \ n < m" - using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ assms coprime_1_left - by blast - have f3: "\na. a = na \ na < n \ coprime na n" - using f1 \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by blast - have "n \ 1" - by (metis (lifting) assms less_irrefl_nat) - then have "\ coprime 0 n" - by simp - then show ?thesis - using f3 f2 by (metis \Min {x |x. x < n \ coprime x n} = 1\ \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ less_one linorder_neqE_nat not_add_less1) - qed - hence \Min {a+j*b| j::nat. j < m} = a\ by blast - thus ?thesis - using \Min {x | x :: nat. x < n \ coprime x n} = 1\ \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ - by linarith - qed - thus ?thesis - using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ - by blast - qed - thus ?thesis using \m \ 0\ \{x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ - by auto - qed - qed - ultimately show ?thesis - by simp -qed - +(* + File: Arith_Prog_Rel_Primes.thy + Author: Jose Manuel Rodriguez Caballero, University of Tartu +*) +section \Problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002)\ +theory Arith_Prog_Rel_Primes + imports + Complex_Main + "HOL-Number_Theory.Number_Theory" +begin + +text \ + Statement of the problem (from ~\cite{putnam}): For which integers $n>1$ does the set of positive + integers less than and relatively prime to $n$ constitute an arithmetic progression? + + The solution of the above problem is theorem @{text arith_prog_rel_primes_solution}. + + First, we will require some auxiliary material before we get started with the actual + solution. +\ + +subsection \Auxiliary results\ + +lemma even_and_odd_parts: + fixes n::nat + assumes \n \ 0\ + shows \\ k q::nat. n = (2::nat)^k*q \ odd q\ +proof- + have \prime (2::nat)\ + by simp + thus ?thesis + using prime_power_canonical[where p = "2" and m = "n"] + assms semiring_normalization_rules(7) by auto +qed + +lemma only_one_odd_div_power2: + fixes n::nat + assumes \n \ 0\ and \\ x. x dvd n \ odd x \ x = 1\ + shows \\ k. n = (2::nat)^k\ + by (metis even_and_odd_parts assms(1) assms(2) dvd_triv_left semiring_normalization_rules(11) + semiring_normalization_rules(7)) + +lemma coprime_power2: + fixes n::nat + assumes \n \ 0\ and \\ x. x < n \ (coprime x n \ odd x)\ + shows \\ k. n = (2::nat)^k\ +proof- + have \x dvd n \ odd x \ x = 1\ + for x + by (metis neq0_conv One_nat_def Suc_1 Suc_lessI assms(1) assms(2) coprime_left_2_iff_odd + dvd_refl linorder_neqE_nat nat_dvd_1_iff_1 nat_dvd_not_less not_coprimeI) + thus ?thesis + using assms(1) only_one_odd_div_power2 + by auto +qed + +subsection \Main result\ + +text \ + The solution to the problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002) +\ + +theorem arith_prog_rel_primes_solution: + fixes n :: nat + assumes \n > 1\ + shows \(prime n \ (\ k. n = 2^k) \ n = 6) \ +(\ a b m. m \ 0 \ {x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ +proof- + have \ (prime n \ (\ k. n = 2^k) \ n = 6) \ + (\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m})\ + proof + show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + if "prime n \ (\k. n = 2 ^ k) \ n = 6" + proof- + have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + if "prime n" + proof- + have \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ + proof- + have \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ + proof + show "{x |x. x < n \ coprime x n} \ {x |x. x \ 0 \ x < n}" + by (smt Collect_mono not_le ord_0_nat ord_eq_0 order_refl prime_gt_1_nat that zero_neq_one) + show "{x |x. x \ 0 \ x < n} \ {x |x. x < n \ coprime x n}" + using coprime_commute prime_nat_iff'' that + by fastforce + qed + obtain m where \m+1 = n\ + using add.commute assms less_imp_add_positive by blast + have \{1+j| j::nat. j < (m::nat)} = {x | x :: nat. x \ 0 \ x < m+1}\ + by (metis Suc_eq_plus1 \m + 1 = n\ gr0_implies_Suc le_simps(3) less_nat_zero_code linorder_not_less nat.simps(3) nat_neq_iff plus_1_eq_Suc ) + hence \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < (m::nat)}\ + using \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ \m+1 = n\ + by auto + from \n > 1\ have \m \ 0\ + using \m + 1 = n\ by linarith + have \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ + using Suc_eq_plus1 \1 < n\ \{x |x. x < n \ coprime x n} = {1 + j |j. j < m}\ + by auto + hence \(\ m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m})\ + using \m \ 0\ + by blast + thus ?thesis by blast + qed + hence \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*1| j::nat. j < m}\ + by auto + thus ?thesis + by blast + qed + moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + if "\k. n = 2 ^ k" + proof- + obtain k where \n = 2 ^ k\ + using \\k. n = 2 ^ k\ by auto + have \k \ 0\ + by (metis \1 < n\ \n = 2 ^ k\ nat_less_le power.simps(1)) + obtain t where \Suc t = k\ + by (metis \k \ 0\ fib.cases) + have \n = 2^(Suc t)\ + by (simp add: \Suc t = k\ \n = 2 ^ k\) + have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < 2^t}\ + proof + show "{x |x. x < n \ coprime x n} \ {1 + j * 2 |j. j < 2^t}" + proof + fix x + assume \x \ {x |x. x < n \ coprime x n}\ + hence \x < n\ + by blast + have \coprime x n\ + using \x \ {x |x. x < n \ coprime x n}\ + by blast + hence \coprime x (2^(Suc k))\ + by (simp add: \k \ 0\ \n = 2 ^ k\) + have \odd x\ + using \coprime x n\ \k \ 0\ \n = 2 ^ k\ + by auto + then obtain j where \x = 1+j*2\ + by (metis add.commute add.left_commute left_add_twice mult_2_right oddE) + have \x < 2^k\ + using \n = 2 ^ k\ \x < n\ \x = 1+j*2\ + by linarith + hence \1+j*2 < 2^k\ + using \x = 1+j*2\ + by blast + hence \j < 2^t\ + using \Suc t = k\ by auto + thus \x \ {1 + j * 2 |j. j < 2^t}\ + using \x = 1+j*2\ + by blast + qed + show "{1 + j * 2 |j. j < 2 ^ t} \ {x |x. x < n \ coprime x n}" + proof + fix x::nat + assume \x \ {1 + j * 2 |j. j < 2 ^ t}\ + then obtain j where \x = 1 + j * 2\ and \j < 2 ^ t\ + by blast + have \x < 2*(2^t)\ + using \x = 1 + j * 2\ \j < 2 ^ t\ + by linarith + hence \x < n\ + by (simp add: \n = 2 ^ Suc t\) + moreover have \coprime x n\ + by (metis (no_types) \\thesis. (\j. \x = 1 + j * 2; j < 2 ^ t\ \ thesis) \ thesis\ \n = 2 ^ k\ coprime_Suc_left_nat coprime_mult_right_iff coprime_power_right_iff plus_1_eq_Suc) + ultimately show \x \ {x |x. x < n \ coprime x n}\ + by blast + qed + qed + have \(2::nat)^(t::nat) \ 0\ + by simp + obtain m where \m = (2::nat)^t\ by blast + have \m \ 0\ + using \m = 2 ^ t\ + by auto + have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ + using \m = 2 ^ t\ \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < 2 ^ t}\ + by auto + from \m \ 0\ \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ + show ?thesis by blast + qed + moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + if "n = 6" + proof- + have \{x | x. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < 2}\ + proof- + have \{x | x::nat. x < 6 \ coprime x 6} = {1, 5}\ + proof + show "{u. \x. u = (x::nat) \ x < 6 \ coprime x 6} \ {1, 5}" + proof + fix u::nat + assume \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ + hence \coprime u 6\ + by blast + have \u < 6\ + using \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ + by blast + moreover have \u \ 0\ + using \coprime u 6\ ord_eq_0 + by fastforce + moreover have \u \ 2\ + using \coprime u 6\ + by auto + moreover have \u \ 3\ + proof- + have \gcd (3::nat) 6 = 3\ + by auto + thus ?thesis + by (metis (no_types) \coprime u 6\ \gcd 3 6 = 3\ coprime_iff_gcd_eq_1 + numeral_eq_one_iff semiring_norm(86)) + qed + moreover have \u \ 4\ + proof- + have \gcd (4::nat) 6 = 2\ + by (metis (no_types, lifting) add_numeral_left gcd_add1 gcd_add2 gcd_nat.idem + numeral_Bit0 numeral_One one_plus_numeral semiring_norm(4) semiring_norm(5)) + thus ?thesis + using \coprime u 6\ coprime_iff_gcd_eq_1 + by auto + qed + ultimately have \u = 1 \ u = 5\ + by auto + thus \u \ {1, 5}\ + by blast + qed + show "{1::nat, 5} \ {x |x. x < 6 \ coprime x 6}" + proof- + have \(1::nat) \ {x |x. x < 6 \ coprime x 6}\ + by simp + moreover have \(5::nat) \ {x |x. x < 6 \ coprime x 6}\ + by (metis Suc_numeral coprime_Suc_right_nat less_add_one mem_Collect_eq + numeral_plus_one semiring_norm(5) semiring_norm(8)) + ultimately show ?thesis + by blast + qed + qed + moreover have \{1+j*4| j::nat. j < 2} = {1, 5}\ + by auto + ultimately show ?thesis by auto + qed + moreover have \(2::nat) \ 0\ + by simp + ultimately have \\ m. m \ 0 \ {x | x :: nat. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < m}\ + by blast + thus ?thesis + using that + by auto + qed + ultimately show ?thesis + using that + by blast + qed + show "prime n \ (\k. n = 2 ^ k) \ n = 6" + if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + proof- + obtain b m where \m \ 0\ and \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + using \\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by auto + show ?thesis + proof(cases \n = 2\) + case True + thus ?thesis + by auto + next + case False + have \b \ 4\ + proof(cases \odd b\) + case True + show ?thesis + proof(rule classical) + assume \\(b \ 4)\ + hence \b > 4\ + using le_less_linear + by blast + obtain m where \m \ 0\ + and \{x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ + using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + have \b \ 0\ + using \4 < b\ + by linarith + have \n = 2 + (m-1)*b\ + proof- + have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1+j*b| j::nat. j < m}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ + \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ + by blast + have \i \ m-1\ + using \i < m\ + by linarith + have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ + using \m \ 0\ + by auto + hence \1 + (m-1)*b \ {x | x::nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \1 + (m-1)*b < n\ + by blast + hence \1 + (m-1)*b \ n-1\ + by linarith + hence \1 + (m-1)*b \ 1+i*b\ + using \n - 1 = 1 + i * b\ + by linarith + hence \(m-1)*b \ i*b\ + by linarith + hence \m-1 \ i\ + using \b \ 0\ + by auto + hence \m-1 = i\ + using \i \ m - 1\ le_antisym + by blast + thus ?thesis + using \m \ 0\ \n - 1 = 1 + i * b\ + by auto + qed + have \m \ 2\ + using \n = 2 + (m - 1)*b\ \n \ 2\ + by auto + hence \1+b \ {1+j*b| j. j < m}\ + by fastforce + hence \1+b \ {x | x::nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \coprime (1+b) n\ + by blast + have \(2::nat) dvd (1+b)\ + using \odd b\ + by simp + hence \coprime (2::nat) n\ + using \coprime (1 + b) n\ coprime_common_divisor coprime_left_2_iff_odd odd_one + by blast + have \(2::nat) < n\ + using \1 < n\ \n \ 2\ + by linarith + have \2 \ {x | x :: nat. x < n \ coprime x n}\ + using \2 < n\ \coprime 2 n\ + by blast + hence \2 \ {1+j*b| j::nat. j < m}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain j::nat where \2 = 1+j*b\ + by blast + have \1 = j*b\ + using \2 = 1+j*b\ + by linarith + thus ?thesis + by simp + qed + next + case False + hence \even b\ + by simp + show ?thesis + proof(rule classical) + assume \\(b \ 4)\ + hence \b > 4\ + using le_less_linear + by blast + obtain m where \ m \ 0\ + and \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ + using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + have \b \ 0\ + using \4 < b\ + by linarith + have \n = 2 + (m-1)*b\ + proof- + have \n-1 \ {x | x::nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1+j*b| j::nat. j < m}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ + \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ + by blast + have \i \ m-1\ + using \i < m\ + by linarith + have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ + using \m \ 0\ + by auto + hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \1 + (m-1)*b < n\ + by blast + hence \1 + (m-1)*b \ n-1\ + by linarith + hence \1 + (m-1)*b \ 1+i*b\ + using \n - 1 = 1 + i * b\ + by linarith + hence \(m-1)*b \ i*b\ + by linarith + hence \m-1 \ i\ + using \b \ 0\ + by auto + hence \m-1 = i\ + using \i \ m - 1\ le_antisym + by blast + thus ?thesis + using \m \ 0\ \n - 1 = 1 + i * b\ + by auto + qed + obtain k :: nat where \b = 2*k\ + using \even b\ + by blast + have \n = 2*(1 + (m-1)*k)\ + using \n = 2 + (m-1)*b\ \b = 2*k\ + by simp + show ?thesis + proof (cases \odd m\) + case True + hence \odd m\ by blast + then obtain t::nat where \m-1 = 2*t\ + by (metis odd_two_times_div_two_nat) + have \n = 2*(1 + b*t)\ + using \m - 1 = 2 * t\ \n = 2 + (m - 1) * b\ + by auto + have \t < m\ + using \m - 1 = 2 * t\ \m \ 0\ + by linarith + have \1 + b*t \ {1+j*b| j::nat. j < m}\ + using \t < m\ + by auto + hence \1 + b*t \ {x | x::nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \coprime (1 + b*t) n\ + by auto + thus ?thesis + by (metis (no_types, lifting) \b = 2 * k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 * (1 + b * t)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_cancel_right_right coprime_mult_right_iff coprime_self mult_cancel_left mult_is_0 nat_dvd_1_iff_1) + next + case False + thus ?thesis + proof(cases \odd k\) + case True + hence \odd k\ + by blast + have \even (1 + (m - 1) * k)\ + by (simp add: False True \m \ 0\) + have \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ + by simp + have \coprime (2 + (m - 1) * k) n\ + using \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ \even (1 + (m - 1) * k)\ + \n = 2 * (1 + (m - 1) * k)\ coprime_common_divisor coprime_mult_right_iff + coprime_right_2_iff_odd odd_one + by blast + have \2 + (m - 1) * k < n\ + by (metis (no_types, lifting) \even (1 + (m - 1) * k)\ \n = 2 * (1 + (m - 1) * k)\ + add_gr_0 add_mono_thms_linordered_semiring(1) dvd_add_left_iff dvd_add_triv_left_iff dvd_imp_le le_add2 le_neq_implies_less less_numeral_extra(1) mult_2 odd_one) + have \2 + (m - 1) * k \ {x | x :: nat. x < n \ coprime x n}\ + using \2 + (m - 1) * k < n\ \coprime (2 + (m - 1) * k) n\ + by blast + hence \2 + (m - 1) * k \ {1 + j * b |j. j < m}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain j::nat where \2 + (m - 1) * k = 1 + j * b\ and \j < m\ + by blast + have \1 + (m - 1) * k = j * b\ + using \2 + (m - 1) * k = 1 + j * b\ + by simp + hence \1 + (m - 1) * k = j * (2*k)\ + using \b = 2 * k\ by blast + thus ?thesis + by (metis \b = 2 * k\ \even b\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ dvd_add_times_triv_right_iff dvd_antisym dvd_imp_le dvd_triv_right even_numeral mult_2 zero_less_numeral) + next + case False + hence \even k\ by auto + have \odd (1 + (m - 1) * k)\ + by (simp add: \even k\ ) + hence \coprime (3 + (m - 1) * k) (1 + (m - 1) * k)\ + by (smt add_numeral_left coprime_common_divisor coprime_right_2_iff_odd dvd_add_left_iff not_coprimeE numeral_Bit1 numeral_One numeral_plus_one one_add_one) + hence \coprime (3 + (m - 1) * k) n\ + by (metis \even k\ \n = 2 * (1 + (m - 1) * k)\ coprime_mult_right_iff coprime_right_2_iff_odd even_add even_mult_iff odd_numeral) + have \3 + (m - 1) * k < n\ + by (smt Groups.add_ac(2) \even k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_Suc_right add_cancel_right_right add_mono_thms_linordered_semiring(1) dvd_imp_le even_add even_mult_iff le_add2 le_neq_implies_less left_add_twice mult_2 neq0_conv numeral_Bit1 numeral_One odd_numeral one_add_one plus_1_eq_Suc) + have \3 + (m - 1) * k \ {x |x. x < n \ coprime x n}\ + using \3 + (m - 1) * k < n\ \coprime (3 + (m - 1) * k) n\ + by blast + hence \3 + (m - 1) * k \ {1 + j * b |j. j < m}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain j::nat where \3 + (m - 1) * k = 1 + j * b\ + by blast + have \2 + (m - 1) * k = j * b\ + using \3 + (m - 1) * k = 1 + j * b\ + by simp + hence \2 + (m - 1) * k = j * 2*k\ + by (simp add: \b = 2 * k\) + thus ?thesis + by (metis \4 < b\ \b = 2 * k\ \even k\ dvd_add_times_triv_right_iff dvd_antisym + dvd_triv_right mult_2 nat_neq_iff numeral_Bit0) + qed + qed + qed + qed + moreover have \b \ 3\ + proof (rule classical) + assume \\ (b \ 3)\ + hence \b = 3\ + by blast + obtain m where \m \ 0\ and + \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ + using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + have \b \ 0\ + by (simp add: \b = 3\) + have \n = 2 + (m-1)*b\ + proof- + have \n-1 \ {x | x::nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1+j*b| j::nat. j < m}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ + \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ + by blast + have \i \ m-1\ + using \i < m\ + by linarith + have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ + using \m \ 0\ + by auto + hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \1 + (m-1)*b < n\ + by blast + hence \1 + (m-1)*b \ n-1\ + by linarith + hence \1 + (m-1)*b \ 1+i*b\ + using \n - 1 = 1 + i * b\ + by linarith + hence \(m-1)*b \ i*b\ + by linarith + hence \m-1 \ i\ + using \b \ 0\ + by auto + hence \m-1 = i\ + using \i \ m - 1\ le_antisym + by blast + thus ?thesis + using \m \ 0\ \n - 1 = 1 + i * b\ + by auto + qed + have \n > 2\ + using \1 < n\ \n \ 2\ + by linarith + hence \ m \ 2 \ using \n = 2 + (m-1)*b\ \b = 3\ + by simp + have \4 \ {1+j*(b::nat)| j::nat. j < m}\ + using \2 \ m\ \b = 3\ + by force + hence \(4::nat) \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by auto + hence \coprime (4::nat) n\ + by blast + have \(2::nat) dvd 4\ + by auto + hence \coprime (2::nat) n\ + using \coprime (4::nat) n\ coprime_divisors dvd_refl + by blast + have \4 < n\ + using \4 \ {x |x. x < n \ coprime x n}\ + by blast + have \2 < (4::nat)\ + by auto + have \2 < n\ + by (simp add: \2 < n\) + hence \2 \ {x | x :: nat. x < n \ coprime x n}\ + using \coprime (2::nat) n\ + by blast + hence \2 \ {1+j*(b::nat)| j::nat. j < m}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain j::nat where \2 = 1+j*3\ + using \b = 3\ + by blast + from \2 = 1+j*3\ + have \1 = j*3\ + by auto + hence \3 dvd 1\ + by auto + thus ?thesis + using nat_dvd_1_iff_1 numeral_eq_one_iff + by blast + qed + ultimately have \b = 0 \ b = 1 \ b = 2 \ b = 4\ + by auto + moreover have \b = 0 \ \k. n = 2^k\ + proof- + assume \b = 0\ + have \{1 + j * b |j. j < m} = {1}\ + using \b = 0\ \m \ 0\ + by auto + hence \{x |x. x < n \ coprime x n} = {1}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \n = 2\ + proof- + have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1}\ + by blast + hence \n-1 = 1\ + by blast + hence \n = 2\ + by simp + thus ?thesis + by blast + qed + hence \n = 2^1\ + by auto + thus ?thesis + by blast + qed + moreover have \b = 1 \ prime n\ + proof- + assume \b = 1\ + have \x < n \ x \ 0 \ coprime x n\ + for x + proof- + assume \x < n\ and \x \ 0\ + have \{1+j| j::nat. j < m} = {x | x::nat. x < m+1 \ x \ 0}\ + by (metis (full_types) Suc_eq_plus1 add_mono1 less_Suc_eq_0_disj nat.simps(3) plus_1_eq_Suc ) + hence \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x < m+1 \ x \ 0}\ + using \b = 1\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by auto + have \coprime (n-1) n\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 < n\ + using \1 < n\ + by auto + have \n-1 \ {x |x. x < n \ coprime x n}\ + using \coprime (n - 1) n\ \n - 1 < n\ + by blast + have \n-1 \ m\ + by (metis (no_types, lifting) CollectD Suc_eq_plus1 Suc_less_eq2 \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ leD le_less_linear not_less_eq_eq ) + have \m \ {x | x :: nat. x < m+1 \ x \ 0}\ + using \m \ 0\ + by auto + have \m \ {x |x. x < n \ coprime x n} \ + using \m \ {x |x. x < m + 1 \ x \ 0}\ + \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ + by blast + have \m < n\ + using \m \ {x |x. x < n \ coprime x n}\ + by blast + have \m+1 = n\ + using \m < n\ \n - 1 \ m\ + by linarith + have \x \ {x | x :: nat. x < m+1 \ x \ 0}\ + using \m + 1 = n\ \x < n\ \x \ 0\ + by blast + hence \x \ {x |x. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ + by blast + thus ?thesis + by blast + qed + thus ?thesis + using assms coprime_commute nat_neq_iff prime_nat_iff'' by auto + qed + moreover have \b = 2 \ \ k. n = 2^k\ + proof- + assume \b = 2\ + have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ + using \b = 2\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by auto + have \x < n \ coprime x n \ odd x\ + for x::nat + proof- + assume \x < n\ + have \coprime x n \ odd x\ + proof- + assume \coprime x n\ + have \x \ {x | x :: nat. x < n \ coprime x n}\ + by (simp add: \coprime x n\ \x < n\) + hence \x \ {1+j*2| j::nat. j < m}\ + using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ + by blast + then obtain j where \x = 1+j*2\ + by blast + thus ?thesis + by simp + qed + moreover have \odd x \ coprime x n\ + proof- + assume \odd x\ + obtain j::nat where \x = 1+j*2\ + by (metis \odd x\ add.commute mult_2_right odd_two_times_div_two_succ one_add_one semiring_normalization_rules(16)) + have \j < (n-1)/2\ + using \x < n\ \x = 1 + j * 2\ + by linarith + have \n = 2*m\ + proof- + have \(2::nat) \ 0\ + by auto + have \n = 2+(m-1)*2\ + proof- + have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1+j*b| j::nat. j < m}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ + \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ + by blast + have \i \ m-1\ + using \i < m\ + by linarith + have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ + using \m \ 0\ by auto + hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \1 + (m-1)*b < n\ + by blast + hence \1 + (m-1)*b \ n-1\ + by linarith + hence \1 + (m-1)*b \ 1+i*b\ + using \n - 1 = 1 + i * b\ + by linarith + hence \(m-1)*b \ i*b\ + by linarith + hence \m-1 \ i\ + proof- + have \b \ 0\ + using \b = 2\ + by simp + thus ?thesis + using \(m - 1) * b \ i * b\ mult_le_cancel2 + by blast + qed + hence \m-1 = i\ + using \i \ m - 1\ le_antisym + by blast + thus ?thesis + using \m \ 0\ \n - 1 = 1 + i * b\ + by (simp add: \b = 2\) + qed + thus ?thesis + by (simp add: \m \ 0\ \n = 2 + (m - 1) * 2\ mult.commute mult_eq_if) + qed + hence \j < m\ + using \x < n\ \x = 1 + j * 2\ + by linarith + hence \x \ {1+j*2| j::nat. j < m}\ + using \x = 1 + j * 2\ + by blast + hence \x \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ + by blast + thus ?thesis + by blast + qed + ultimately show ?thesis + by blast + qed + thus ?thesis + using coprime_power2 assms + by auto + qed + moreover have \b = 4 \ n = 6\ + proof- + assume \b = 4\ + have \n = 2 \ n = 6\ + proof(rule classical) + assume \\ (n = 2 \ n = 6)\ + have \(4::nat) \ 0\ + by auto + have \n = 2+(m-1)*4\ + proof- + have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ + using \1 < n\ coprime_diff_one_left_nat + by auto + have \n-1 \ {1+j*b| j::nat. j < m}\ + using \n - 1 \ {x |x. x < n \ coprime x n}\ + \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ + by blast + have \i \ m-1\ + using \i < m\ + by linarith + have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ + using \m \ 0\ + by auto + hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ + using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + hence \1 + (m-1)*b < n\ + by blast + hence \1 + (m-1)*b \ n-1\ + by linarith + hence \1 + (m-1)*b \ 1+i*b\ + using \n - 1 = 1 + i * b\ + by linarith + hence \(m-1)*b \ i*b\ + by linarith + hence \m-1 \ i\ + proof- + have \b \ 0\ + using \b = 4\ by auto + thus ?thesis + using \(m - 1) * b \ i * b\ mult_le_cancel2 + by blast + qed + hence \m-1 = i\ + using \i \ m - 1\ le_antisym + by blast + thus ?thesis + using \m \ 0\ \n - 1 = 1 + i * b\ + by (simp add: \b = 4\) + qed + hence \n = 4*m - 2\ + by (simp add: \m \ 0\ mult.commute mult_eq_if) + have \m \ 3\ + using \\ (n = 2 \ n = 6)\ \n = 2 + (m - 1) * 4\ + by auto + hence \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ + by force + hence \9 \ {1+j*4| j::nat. j < 3}\ + by force + hence \9 \ {1+j*4| j::nat. j < m}\ + using \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ + by blast + hence \9 \ {x | x :: nat. x < n \ coprime x n}\ + using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by auto + hence \coprime (9::nat) n\ + by blast + have \(3::nat) dvd 9\ + by auto + have \coprime (3::nat) n\ using \coprime (9::nat) n\ \(3::nat) dvd 9\ + by (metis coprime_commute coprime_mult_right_iff dvd_def) + have \(3::nat) < n\ + by (metis One_nat_def Suc_lessI \1 < n\ \\ (n = 2 \ n = 6)\ \coprime 3 n\ + coprime_self numeral_2_eq_2 numeral_3_eq_3 less_numeral_extra(1) nat_dvd_not_less) + have \3 \ {x | x :: nat. x < n \ coprime x n}\ + using \3 < n\ \coprime 3 n\ + by blast + hence \(3::nat) \ {1+j*4| j::nat. j < m}\ + using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ + by blast + then obtain j::nat where \(3::nat) = 1 + j*4\ + by blast + have \2 = j*4\ + using numeral_3_eq_3 \(3::nat) = 1 + j*4\ + by linarith + hence \1 = j*2\ + by linarith + hence \even 1\ + by simp + thus ?thesis + using odd_one + by blast + qed + thus ?thesis + by (simp add: False) + qed + ultimately show ?thesis + by blast + qed + qed + qed + moreover have \(\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}) + \ (\ a b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ + proof + show "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" + if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + using that + by blast + show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" + if "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" + proof- + obtain a b m::nat where \m \ 0\ + and \{x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ + using \\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ + by auto + have \a = 1\ + proof- + have \{x | x :: nat. x < n \ coprime x n} = {(a::nat)+j*(b::nat)| j::nat. j < m} \ a = 1\ + proof- + have \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ + using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ + by auto + have \Min {x | x :: nat. x < n \ coprime x n} = 1\ + proof- + have \finite {x | x :: nat. x < n \ coprime x n}\ + by auto + have \{x | x :: nat. x < n \ coprime x n} \ {}\ + using \1 < n\ by auto + have \1 \ {x | x :: nat. x < n \ coprime x n}\ + using \1 < n\ + by auto + have \\ x. coprime x n \ x \ 1\ + using \1 < n\ le_less_linear + by fastforce + hence \\ x. x < n \ coprime x n \ x \ 1\ + by blast + hence \\ x \ {x | x :: nat. x < n \ coprime x n}. x \ 1\ + by blast + hence \Min {x | x :: nat. x < n \ coprime x n} \ 1\ + using \finite {x | x :: nat. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} \ {}\ + by auto + thus ?thesis + using Min_le \1 \ {x |x. x < n \ coprime x n}\ \finite {x |x. x < n \ coprime x n}\ + antisym by blast + qed + have \Min {a+j*b| j::nat. j < m} = a\ + proof - + have f1: "\n. a = a + n * b \ n < m" + using \m \ 0\ + by auto + have f2: "\n. 1 = a + n * b \ n < m" + using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ assms coprime_1_left + by blast + have f3: "\na. a = na \ na < n \ coprime na n" + using f1 \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by blast + have "n \ 1" + by (metis (lifting) assms less_irrefl_nat) + then have "\ coprime 0 n" + by simp + then show ?thesis + using f3 f2 by (metis \Min {x |x. x < n \ coprime x n} = 1\ \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ less_one linorder_neqE_nat not_add_less1) + qed + hence \Min {a+j*b| j::nat. j < m} = a\ by blast + thus ?thesis + using \Min {x | x :: nat. x < n \ coprime x n} = 1\ \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ + by linarith + qed + thus ?thesis + using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ + by blast + qed + thus ?thesis using \m \ 0\ \{x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ + by auto + qed + qed + ultimately show ?thesis + by simp +qed + end \ No newline at end of file diff --git a/thys/Arith_Prog_Rel_Primes/document/root.bib b/thys/Arith_Prog_Rel_Primes/document/root.bib --- a/thys/Arith_Prog_Rel_Primes/document/root.bib +++ b/thys/Arith_Prog_Rel_Primes/document/root.bib @@ -1,5 +1,5 @@ - -@misc{putnam, - title="Problem {``ARITHMETIC PROGRESSIONS"}, from {P}utnam exam problems 2002, https://www.ocf.berkeley.edu/~wwu/riddles/putnam.shtml" -} - + +@misc{putnam, + title="Problem {``ARITHMETIC PROGRESSIONS"}, from {P}utnam exam problems 2002, https://www.ocf.berkeley.edu/~wwu/riddles/putnam.shtml" +} + diff --git a/thys/Arith_Prog_Rel_Primes/document/root.tex b/thys/Arith_Prog_Rel_Primes/document/root.tex --- a/thys/Arith_Prog_Rel_Primes/document/root.tex +++ b/thys/Arith_Prog_Rel_Primes/document/root.tex @@ -1,31 +1,31 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym} - - -%this should be the last package used -\usepackage{pdfsetup} - -% urls in roman style, theory text in math-similar italics -\urlstyle{rm} -\isabellestyle{it} - -\begin{document} - -\title{Arithmetic progressions and relative primes} -\author{Jos\'e Manuel Rodr\'iguez Caballero} -\maketitle - -\begin{abstract} -This article provides a formalization of the solution obtained by the author of the Problem ``ARITHMETIC PROGRESSIONS" from the -Putnam exam problems \cite{putnam} of $2002$. The statement of the problem is as follows: For which integers $n>1$ does the set of positive integers less than and relatively prime to $n$ constitute an arithmetic progression? -\end{abstract} - -\tableofcontents - -\input{session} - - -\bibliographystyle{abbrv} -\bibliography{root} - -\end{document} +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + + +%this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Arithmetic progressions and relative primes} +\author{Jos\'e Manuel Rodr\'iguez Caballero} +\maketitle + +\begin{abstract} +This article provides a formalization of the solution obtained by the author of the Problem ``ARITHMETIC PROGRESSIONS" from the +Putnam exam problems \cite{putnam} of $2002$. The statement of the problem is as follows: For which integers $n>1$ does the set of positive integers less than and relatively prime to $n$ constitute an arithmetic progression? +\end{abstract} + +\tableofcontents + +\input{session} + + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/Banach_Steinhaus/Banach_Steinhaus.thy b/thys/Banach_Steinhaus/Banach_Steinhaus.thy --- a/thys/Banach_Steinhaus/Banach_Steinhaus.thy +++ b/thys/Banach_Steinhaus/Banach_Steinhaus.thy @@ -1,485 +1,485 @@ -(* - File: Banach_Steinhaus.thy - Author: Dominique Unruh, University of Tartu - Author: Jose Manuel Rodriguez Caballero, University of Tartu -*) -section \Banach-Steinhaus theorem\ - -theory Banach_Steinhaus - imports Banach_Steinhaus_Missing -begin - -text \ - We formalize Banach-Steinhaus theorem as theorem @{text banach_steinhaus}. This theorem was - originally proved in Banach-Steinhaus's paper~\cite{banach1927principe}. For the proof, we follow - Sokal's approach~\cite{sokal2011really}. Furthermore, we prove as a corollary a result about - pointwise convergent sequences of bounded operators whose domain is a Banach space. -\ - -subsection \Preliminaries for Sokal's proof of Banach-Steinhaus theorem\ - -lemma linear_plus_norm: - includes notation_norm - assumes \linear f\ - shows \\f \\ \ max \f (x + \)\ \f (x - \)\\ - text \ - Explanation: For arbitrary \<^term>\x\ and a linear operator \<^term>\f\, - \<^term>\norm (f \)\ is upper bounded by the maximum of the norms - of the shifts of \<^term>\f\ (i.e., \<^term>\f (x + \)\ and \<^term>\f (x - \)\). -\ -proof- - have \norm (f \) = norm ( (inverse (of_nat 2)) *\<^sub>R (f (x + \) - f (x - \)) )\ - by (smt add_diff_cancel_left' assms diff_add_cancel diff_diff_add linear_diff midpoint_def - midpoint_plus_self of_nat_1 of_nat_add one_add_one scaleR_half_double) - also have \\ = inverse (of_nat 2) * norm (f (x + \) - f (x - \))\ - using Real_Vector_Spaces.real_normed_vector_class.norm_scaleR by simp - also have \\ \ inverse (of_nat 2) * (norm (f (x + \)) + norm (f (x - \)))\ - by (simp add: norm_triangle_ineq4) - also have \\ \ max (norm (f (x + \))) (norm (f (x - \)))\ - by auto - finally show ?thesis by blast -qed - -lemma onorm_Sup_on_ball: - includes notation_norm - assumes \r > 0\ - shows "\f\ \ Sup ( (\x. \f *\<^sub>v x\) ` (ball x r) ) / r" - text \ - Explanation: Let \<^term>\f\ be a bounded operator and let \<^term>\x\ be a point. For any \<^term>\r > 0\, - the operator norm of \<^term>\f\ is bounded above by the supremum of $f$ applied to the open ball of - radius \<^term>\r\ around \<^term>\x\, divided by \<^term>\r\. -\ -proof- - have bdd_above_3: \bdd_above ((\x. \f *\<^sub>v x\) ` (ball 0 r))\ - proof - - obtain M where \\ \. \f *\<^sub>v \\ \ M * norm \\ and \M \ 0\ - using norm_blinfun norm_ge_zero by blast - hence \\ \. \ \ ball 0 r \ \f *\<^sub>v \\ \ M * r\ - using \r > 0\ by (smt mem_ball_0 mult_left_mono) - thus ?thesis by (meson bdd_aboveI2) - qed - have bdd_above_2: \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ - proof- - have \bdd_above ((\ \. \f *\<^sub>v x\) ` (ball 0 r))\ - by auto - moreover have \bdd_above ((\ \. \f *\<^sub>v \\) ` (ball 0 r))\ - using bdd_above_3 by blast - ultimately have \bdd_above ((\ \. \f *\<^sub>v x\ + \f *\<^sub>v \\) ` (ball 0 r))\ - by (rule bdd_above_plus) - then obtain M where \\ \. \ \ ball 0 r \ \f *\<^sub>v x\ + \f *\<^sub>v \\ \ M\ - unfolding bdd_above_def by (meson image_eqI) - moreover have \\f *\<^sub>v (x + \)\ \ \f *\<^sub>v x\ + \f *\<^sub>v \\\ for \ - by (simp add: blinfun.add_right norm_triangle_ineq) - ultimately have \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ M\ - by (simp add: blinfun.add_right norm_triangle_le) - thus ?thesis by (meson bdd_aboveI2) - qed - have bdd_above_4: \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - proof- - obtain K where K_def: \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ K\ - using \bdd_above ((\ \. norm (f (x + \))) ` (ball 0 r))\ unfolding bdd_above_def - by (meson image_eqI) - have \\ \ ball (0::'a) r \ -\ \ ball 0 r\ for \ - by auto - thus ?thesis by (metis K_def ab_group_add_class.ab_diff_conv_add_uminus bdd_aboveI2) - qed - have bdd_above_1: \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - proof- - have \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ - using bdd_above_2 by blast - moreover have \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - using bdd_above_4 by blast - ultimately show ?thesis - unfolding max_def apply auto apply (meson bdd_above_Int1 bdd_above_mono image_Int_subset) - by (meson bdd_above_Int1 bdd_above_mono image_Int_subset) - qed - have bdd_above_6: \bdd_above ((\t. \f *\<^sub>v t\) ` ball x r)\ - proof- - have \bounded (ball x r)\ - by simp - hence \bounded ((\t. \f *\<^sub>v t\) ` ball x r)\ - by (metis (no_types) add.left_neutral bdd_above_2 bdd_above_norm bounded_norm_comp - image_add_ball image_image) - thus ?thesis - by (simp add: bounded_imp_bdd_above) - qed - have norm_1: \(\\. \f *\<^sub>v (x + \)\) ` ball 0 r = (\t. \f *\<^sub>v t\) ` ball x r\ - by (metis add.right_neutral ball_translation image_image) - have bdd_above_5: \bdd_above ((\\. norm (f (x + \))) ` ball 0 r)\ - by (simp add: bdd_above_2) - have norm_2: \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \ - proof- - assume \\\\ < r\ - hence \\ \ ball (0::'a) r\ - by auto - hence \-\ \ ball (0::'a) r\ - by auto - thus ?thesis - by (metis (no_types, lifting) ab_group_add_class.ab_diff_conv_add_uminus image_iff) - qed - have norm_2': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \ - proof- - assume \norm \ < r\ - hence \\ \ ball (0::'a) r\ - by auto - hence \-\ \ ball (0::'a) r\ - by auto - thus ?thesis - by (metis (no_types, lifting) diff_minus_eq_add image_iff) - qed - have bdd_above_6: \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ - by (simp add: bdd_above_4) - have Sup_2: \(SUP \\ball 0 r. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) = - max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\)\ - proof- - have \ball (0::'a) r \ {}\ - using \r > 0\ by auto - moreover have \bdd_above ((\\. \f *\<^sub>v (x + \)\) ` ball 0 r)\ - using bdd_above_5 by blast - moreover have \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ - using bdd_above_6 by blast - ultimately show ?thesis - using max_Sup - by (metis (mono_tags, lifting) Banach_Steinhaus_Missing.pointwise_max_def image_cong) - qed - have Sup_3': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \::'a - by (simp add: norm_2') - have Sup_3'': \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \::'a - by (simp add: norm_2) - have Sup_3: \max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\) = - (SUP \\ball 0 r. \f *\<^sub>v (x + \)\)\ - proof- - have \(\\. \f *\<^sub>v (x + \)\) ` (ball 0 r) = (\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)\ - apply auto using Sup_3' apply auto using Sup_3'' by blast - hence \Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))=Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - by simp - thus ?thesis by simp - qed - have Sup_1: \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ Sup ( (\\. \f *\<^sub>v \\) ` (ball x r) )\ - proof- - have \(\t. \f *\<^sub>v t\) \ \ max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\\ for \ - apply(rule linear_plus_norm) apply (rule bounded_linear.linear) - by (simp add: blinfun.bounded_linear_right) - moreover have \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - using bdd_above_1 by blast - moreover have \ball (0::'a) r \ {}\ - using \r > 0\ by auto - ultimately have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ - Sup ((\\. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - using cSUP_mono by smt - also have \\ = max (Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))) - (Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)))\ - using Sup_2 by blast - also have \\ = Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ - using Sup_3 by blast - also have \\ = Sup ((\\. \f *\<^sub>v \\) ` (ball x r))\ - by (metis add.right_neutral ball_translation image_image) - finally show ?thesis by blast - qed - have \\f\ = (SUP x\ball 0 r. \f *\<^sub>v x\) / r\ - using \0 < r\ onorm_r by blast - moreover have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) / r \ Sup ((\\. \f *\<^sub>v \\) ` (ball x r)) / r\ - using Sup_1 \0 < r\ divide_right_mono by fastforce - ultimately have \\f\ \ Sup ((\t. \f *\<^sub>v t\) ` ball x r) / r\ - by simp - thus ?thesis by simp -qed - -lemma onorm_Sup_on_ball': - includes notation_norm - assumes \r > 0\ and \\ < 1\ - shows \\\\ball x r. \ * r * \f\ \ \f *\<^sub>v \\\ - text \ - In the proof of Banach-Steinhaus theorem, we will use this variation of the - lemma @{text onorm_Sup_on_ball}. - - Explanation: Let \<^term>\f\ be a bounded operator, let \<^term>\x\ be a point and let \<^term>\r\ be a - positive real number. For any real number \<^term>\\ < 1\, there is a point \<^term>\\\ in the open ball - of radius \<^term>\r\ around \<^term>\x\ such that \<^term>\\ * r * \f\ \ \f *\<^sub>v \\\. -\ -proof(cases \f = 0\) - case True - thus ?thesis by (metis assms(1) centre_in_ball mult_zero_right norm_zero order_refl - zero_blinfun.rep_eq) -next - case False - have bdd_above_1: \bdd_above ((\t. \(*\<^sub>v) f t\) ` ball x r)\ for f::\'a \\<^sub>L 'b\ - using assms(1) bounded_linear_image by (simp add: bounded_linear_image - blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) - have \norm f > 0\ - using \f \ 0\ by auto - have \norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) ) / r\ - using \r > 0\ by (simp add: onorm_Sup_on_ball) - hence \r * norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) )\ - using \0 < r\ by (smt divide_strict_right_mono nonzero_mult_div_cancel_left) - moreover have \\ * r * norm f < r * norm f\ - using \\ < 1\ using \0 < norm f\ \0 < r\ by auto - ultimately have \\ * r * norm f < Sup ( (norm \ ((*\<^sub>v) f)) ` (ball x r) )\ - by simp - moreover have \(norm \ ( (*\<^sub>v) f)) ` (ball x r) \ {}\ - using \0 < r\ by auto - moreover have \bdd_above ((norm \ ( (*\<^sub>v) f)) ` (ball x r))\ - using bdd_above_1 apply transfer by simp - ultimately have \\t \ (norm \ ( (*\<^sub>v) f)) ` (ball x r). \ * r * norm f < t\ - by (simp add: less_cSup_iff) - thus ?thesis by (smt comp_def image_iff) -qed - -subsection \Banach-Steinhaus theorem\ - -theorem banach_steinhaus: - fixes f::\'c \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ - assumes \\x. bounded (range (\n. (f n) *\<^sub>v x))\ - shows \bounded (range f)\ - text\ - This is Banach-Steinhaus Theorem. - - Explanation: If a family of bounded operators on a Banach space - is pointwise bounded, then it is uniformly bounded. -\ -proof(rule classical) - assume \\(bounded (range f))\ - have sum_1: \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ - proof- - have \summable (\n. (inverse (real_of_nat 3))^n)\ - using Series.summable_geometric_iff [where c = "inverse (real_of_nat 3)"] by auto - moreover have \(inverse (real_of_nat 3))^n = inverse (real_of_nat 3^n)\ for n::nat - using power_inverse by blast - ultimately have \summable (\n. inverse (real_of_nat 3^n))\ - by auto - hence \bounded (range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. - using summable_imp_sums_bounded[where f = "(\n. inverse (real_of_nat 3^n))"] - lessThan_atLeast0 by auto - hence \\M. \h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. M\ - using bounded_iff by blast - then obtain M where \h\range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ - for h - by blast - have sum_2: \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n - proof- - have \norm (sum (\ k. inverse (real 3 ^ k)) {0..< Suc n}) \ M\ - using \\h. h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ - by blast - hence \norm (sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ - by (simp add: atLeastLessThanSuc_atLeastAtMost) - hence \(sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ - by auto - thus ?thesis by blast - qed - have \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n - using sum_2 by blast - thus ?thesis by blast - qed - have \of_rat 2/3 < (1::real)\ - by auto - hence \\g::'a \\<^sub>L 'b. \x. \r. \\. g \ 0 \ r > 0 - \ (\\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g \))\ - using onorm_Sup_on_ball' by blast - hence \\\. \g::'a \\<^sub>L 'b. \x. \r. g \ 0 \ r > 0 - \ ((\ g x r)\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r)))\ - by metis - then obtain \ where f1: \\g \ 0; r > 0\ \ - \ g x r \ ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r))\ - for g::\'a \\<^sub>L 'b\ and x and r - by blast - have \\n. \k. norm (f k) \ 4^n\ - using \\(bounded (range f))\ by (metis (mono_tags, hide_lams) boundedI image_iff linear) - hence \\k. \n. norm (f (k n)) \ 4^n\ - by metis - hence \\k. \n. norm ((f \ k) n) \ 4^n\ - by simp - then obtain k where \norm ((f \ k) n) \ 4^n\ for n - by blast - define T where \T = f \ k\ - have \T n \ range f\ for n - unfolding T_def by simp - have \norm (T n) \ of_nat (4^n)\ for n - unfolding T_def using \\ n. norm ((f \ k) n) \ 4^n\ by auto - hence \T n \ 0\ for n - by (smt T_def \\n. 4 ^ n \ norm ((f \ k) n)\ norm_zero power_not_zero zero_le_power) - have \inverse (of_nat 3^n) > (0::real)\ for n - by auto - define y::\nat \ 'a\ where \y = rec_nat 0 (\n x. \ (T n) x (inverse (of_nat 3^n)))\ - have \y (Suc n) \ ball (y n) (inverse (of_nat 3^n))\ for n - using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto - hence \norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ for n - unfolding ball_def apply auto using dist_norm by (smt norm_minus_commute) - moreover have \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ - using sum_1 by blast - moreover have \Cauchy y\ - using convergent_series_Cauchy[where a = "\n. inverse (of_nat 3^n)" and \ = y] dist_norm - by (metis calculation(1) calculation(2)) - hence \\ x. y \ x\ - by (simp add: convergent_eq_Cauchy) - then obtain x where \y \ x\ - by blast - have norm_2: \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n - proof- - have \inverse (real_of_nat 3) < 1\ - by simp - moreover have \y 0 = 0\ - using y_def by auto - ultimately have \norm (x - y (Suc n)) - \ (inverse (of_nat 3)) * inverse (1 - (inverse (of_nat 3))) * ((inverse (of_nat 3)) ^ n)\ - using bound_Cauchy_to_lim[where c = "inverse (of_nat 3)" and y = y and x = x] - power_inverse semiring_norm(77) \y \ x\ - \\ n. norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ by (metis divide_inverse) - moreover have \inverse (real_of_nat 3) * inverse (1 - (inverse (of_nat 3))) - = inverse (of_nat 2)\ - by auto - ultimately show ?thesis - by (metis power_inverse) - qed - have \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n - using norm_2 by blast - have \\ M. \ n. norm ((*\<^sub>v) (T n) x) \ M\ - unfolding T_def apply auto - by (metis \\x. bounded (range (\n. (*\<^sub>v) (f n) x))\ bounded_iff rangeI) - then obtain M where \norm ((*\<^sub>v) (T n) x) \ M\ for n - by blast - have norm_1: \norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x) - \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + norm ((*\<^sub>v) (T n) x)\ for n - proof- - have \norm (y (Suc n) - x) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ - using \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ - by (simp add: norm_minus_commute) - moreover have \norm (T n) \ 0\ - by auto - ultimately have \norm (T n) * norm (y (Suc n) - x) - \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n)\ - by (simp add: \\n. T n \ 0\) - thus ?thesis by simp - qed - have inverse_2: \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) - \ norm ((*\<^sub>v) (T n) x)\ for n - proof- - have \(of_rat 2/3)*(inverse (of_nat 3^n))*norm (T n) \ norm ((*\<^sub>v) (T n) (y (Suc n)))\ - using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto - also have \\ = norm ((*\<^sub>v) (T n) ((y (Suc n) - x) + x))\ - by auto - also have \\ = norm ((*\<^sub>v) (T n) (y (Suc n) - x) + (*\<^sub>v) (T n) x)\ - apply transfer apply auto by (metis diff_add_cancel linear_simps(1)) - also have \\ \ norm ((*\<^sub>v) (T n) (y (Suc n) - x)) + norm ((*\<^sub>v) (T n) x)\ - by (simp add: norm_triangle_ineq) - also have \\ \ norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x)\ - apply transfer apply auto using onorm by auto - also have \\ \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n) - + norm ((*\<^sub>v) (T n) x)\ - using norm_1 by blast - finally have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) - \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) - + norm ((*\<^sub>v) (T n) x)\ - by blast - hence \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) - - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ - by linarith - moreover have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) - - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) - = (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ - by fastforce - ultimately show \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ - by linarith - qed - have inverse_3: \(inverse (of_nat 6)) * (of_rat (4/3)^n) - \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ for n - proof- - have \of_rat (4/3)^n = inverse (real 3 ^ n) * (of_nat 4^n)\ - apply auto by (metis divide_inverse_commute of_rat_divide power_divide of_rat_numeral_eq) - also have \\ \ inverse (real 3 ^ n) * norm (T n)\ - using \\n. norm (T n) \ of_nat (4^n)\ by simp - finally have \of_rat (4/3)^n \ inverse (real 3 ^ n) * norm (T n)\ - by blast - moreover have \inverse (of_nat 6) > (0::real)\ - by auto - ultimately show ?thesis by auto - qed - have inverse_1: \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n - proof- - have \(inverse (of_nat 6)) * (of_rat (4/3)^n) - \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ - using inverse_3 by blast - also have \\ \ norm ((*\<^sub>v) (T n) x)\ - using inverse_2 by blast - finally have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ norm ((*\<^sub>v) (T n) x)\ - by auto - thus ?thesis using \\ n. norm ((*\<^sub>v) (T n) x) \ M\ by smt - qed - have \\n. M < (inverse (of_nat 6)) * (of_rat (4/3)^n)\ - using Real.real_arch_pow by auto - moreover have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n - using inverse_1 by blast - ultimately show ?thesis by smt -qed - -subsection \A consequence of Banach-Steinhaus theorem\ - -corollary bounded_linear_limit_bounded_linear: - fixes f::\nat \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ - assumes \\x. convergent (\n. (f n) *\<^sub>v x)\ - shows \\g. (\n. (*\<^sub>v) (f n)) \pointwise\ (*\<^sub>v) g\ - text\ - Explanation: If a sequence of bounded operators on a Banach space converges - pointwise, then the limit is also a bounded operator. -\ -proof- - have \\l. (\n. (*\<^sub>v) (f n) x) \ l\ for x - by (simp add: \\x. convergent (\n. (*\<^sub>v) (f n) x)\ convergentD) - hence \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ - unfolding pointwise_convergent_to_def by metis - obtain F where \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ - using \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto - have \\x. (\ n. (*\<^sub>v) (f n) x) \ F x\ - using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer - by (simp add: pointwise_convergent_to_def) - have \bounded (range f)\ - using \\x. convergent (\n. (*\<^sub>v) (f n) x)\ banach_steinhaus - \\x. \l. (\n. (*\<^sub>v) (f n) x) \ l\ convergent_imp_bounded by blast - have norm_f_n: \\ M. \ n. norm (f n) \ M\ - unfolding bounded_def - by (meson UNIV_I \bounded (range f)\ bounded_iff image_eqI) - have \isCont (\ t::'b. norm t) y\ for y::'b - using Limits.isCont_norm by simp - hence \(\ n. norm ((*\<^sub>v) (f n) x)) \ (norm (F x))\ for x - using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ by (simp add: tendsto_norm) - hence norm_f_n_x: \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x - using Elementary_Metric_Spaces.convergent_imp_bounded - by (metis UNIV_I \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ bounded_iff image_eqI) - have norm_f: \\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K\ - proof- - have \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x - using norm_f_n_x \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by blast - hence \\ M. \ n. norm (f n) \ M\ - using norm_f_n by simp - then obtain M::real where \\ M. \ n. norm (f n) \ M\ - by blast - have \\ n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * norm (f n)\ - apply transfer apply auto by (metis mult.commute onorm) - thus ?thesis using \\ M. \ n. norm (f n) \ M\ - by (metis (no_types, hide_lams) dual_order.trans norm_eq_zero order_refl - real_mult_le_cancel_iff2 vector_space_over_itself.scale_zero_left zero_less_norm_iff) - qed - have norm_F_x: \\K. \x. norm (F x) \ norm x * K\ - proof- - have "\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K" - using norm_f \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by auto - thus ?thesis - using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ apply transfer - by (metis Lim_bounded tendsto_norm) - qed - have \linear F\ - proof(rule linear_limit_linear) - show \linear ((*\<^sub>v) (f n))\ for n - apply transfer apply auto by (simp add: bounded_linear.linear) - show \f \pointwise\ F\ - using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto - qed - moreover have \bounded_linear_axioms F\ - using norm_F_x by (simp add: \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ bounded_linear_axioms_def) - ultimately have \bounded_linear F\ - unfolding bounded_linear_def by blast - hence \\g. (*\<^sub>v) g = F\ - using bounded_linear_Blinfun_apply by auto - thus ?thesis using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer by auto -qed - -end +(* + File: Banach_Steinhaus.thy + Author: Dominique Unruh, University of Tartu + Author: Jose Manuel Rodriguez Caballero, University of Tartu +*) +section \Banach-Steinhaus theorem\ + +theory Banach_Steinhaus + imports Banach_Steinhaus_Missing +begin + +text \ + We formalize Banach-Steinhaus theorem as theorem @{text banach_steinhaus}. This theorem was + originally proved in Banach-Steinhaus's paper~\cite{banach1927principe}. For the proof, we follow + Sokal's approach~\cite{sokal2011really}. Furthermore, we prove as a corollary a result about + pointwise convergent sequences of bounded operators whose domain is a Banach space. +\ + +subsection \Preliminaries for Sokal's proof of Banach-Steinhaus theorem\ + +lemma linear_plus_norm: + includes notation_norm + assumes \linear f\ + shows \\f \\ \ max \f (x + \)\ \f (x - \)\\ + text \ + Explanation: For arbitrary \<^term>\x\ and a linear operator \<^term>\f\, + \<^term>\norm (f \)\ is upper bounded by the maximum of the norms + of the shifts of \<^term>\f\ (i.e., \<^term>\f (x + \)\ and \<^term>\f (x - \)\). +\ +proof- + have \norm (f \) = norm ( (inverse (of_nat 2)) *\<^sub>R (f (x + \) - f (x - \)) )\ + by (smt add_diff_cancel_left' assms diff_add_cancel diff_diff_add linear_diff midpoint_def + midpoint_plus_self of_nat_1 of_nat_add one_add_one scaleR_half_double) + also have \\ = inverse (of_nat 2) * norm (f (x + \) - f (x - \))\ + using Real_Vector_Spaces.real_normed_vector_class.norm_scaleR by simp + also have \\ \ inverse (of_nat 2) * (norm (f (x + \)) + norm (f (x - \)))\ + by (simp add: norm_triangle_ineq4) + also have \\ \ max (norm (f (x + \))) (norm (f (x - \)))\ + by auto + finally show ?thesis by blast +qed + +lemma onorm_Sup_on_ball: + includes notation_norm + assumes \r > 0\ + shows "\f\ \ Sup ( (\x. \f *\<^sub>v x\) ` (ball x r) ) / r" + text \ + Explanation: Let \<^term>\f\ be a bounded operator and let \<^term>\x\ be a point. For any \<^term>\r > 0\, + the operator norm of \<^term>\f\ is bounded above by the supremum of $f$ applied to the open ball of + radius \<^term>\r\ around \<^term>\x\, divided by \<^term>\r\. +\ +proof- + have bdd_above_3: \bdd_above ((\x. \f *\<^sub>v x\) ` (ball 0 r))\ + proof - + obtain M where \\ \. \f *\<^sub>v \\ \ M * norm \\ and \M \ 0\ + using norm_blinfun norm_ge_zero by blast + hence \\ \. \ \ ball 0 r \ \f *\<^sub>v \\ \ M * r\ + using \r > 0\ by (smt mem_ball_0 mult_left_mono) + thus ?thesis by (meson bdd_aboveI2) + qed + have bdd_above_2: \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ + proof- + have \bdd_above ((\ \. \f *\<^sub>v x\) ` (ball 0 r))\ + by auto + moreover have \bdd_above ((\ \. \f *\<^sub>v \\) ` (ball 0 r))\ + using bdd_above_3 by blast + ultimately have \bdd_above ((\ \. \f *\<^sub>v x\ + \f *\<^sub>v \\) ` (ball 0 r))\ + by (rule bdd_above_plus) + then obtain M where \\ \. \ \ ball 0 r \ \f *\<^sub>v x\ + \f *\<^sub>v \\ \ M\ + unfolding bdd_above_def by (meson image_eqI) + moreover have \\f *\<^sub>v (x + \)\ \ \f *\<^sub>v x\ + \f *\<^sub>v \\\ for \ + by (simp add: blinfun.add_right norm_triangle_ineq) + ultimately have \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ M\ + by (simp add: blinfun.add_right norm_triangle_le) + thus ?thesis by (meson bdd_aboveI2) + qed + have bdd_above_4: \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + proof- + obtain K where K_def: \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ K\ + using \bdd_above ((\ \. norm (f (x + \))) ` (ball 0 r))\ unfolding bdd_above_def + by (meson image_eqI) + have \\ \ ball (0::'a) r \ -\ \ ball 0 r\ for \ + by auto + thus ?thesis by (metis K_def ab_group_add_class.ab_diff_conv_add_uminus bdd_aboveI2) + qed + have bdd_above_1: \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + proof- + have \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ + using bdd_above_2 by blast + moreover have \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + using bdd_above_4 by blast + ultimately show ?thesis + unfolding max_def apply auto apply (meson bdd_above_Int1 bdd_above_mono image_Int_subset) + by (meson bdd_above_Int1 bdd_above_mono image_Int_subset) + qed + have bdd_above_6: \bdd_above ((\t. \f *\<^sub>v t\) ` ball x r)\ + proof- + have \bounded (ball x r)\ + by simp + hence \bounded ((\t. \f *\<^sub>v t\) ` ball x r)\ + by (metis (no_types) add.left_neutral bdd_above_2 bdd_above_norm bounded_norm_comp + image_add_ball image_image) + thus ?thesis + by (simp add: bounded_imp_bdd_above) + qed + have norm_1: \(\\. \f *\<^sub>v (x + \)\) ` ball 0 r = (\t. \f *\<^sub>v t\) ` ball x r\ + by (metis add.right_neutral ball_translation image_image) + have bdd_above_5: \bdd_above ((\\. norm (f (x + \))) ` ball 0 r)\ + by (simp add: bdd_above_2) + have norm_2: \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \ + proof- + assume \\\\ < r\ + hence \\ \ ball (0::'a) r\ + by auto + hence \-\ \ ball (0::'a) r\ + by auto + thus ?thesis + by (metis (no_types, lifting) ab_group_add_class.ab_diff_conv_add_uminus image_iff) + qed + have norm_2': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \ + proof- + assume \norm \ < r\ + hence \\ \ ball (0::'a) r\ + by auto + hence \-\ \ ball (0::'a) r\ + by auto + thus ?thesis + by (metis (no_types, lifting) diff_minus_eq_add image_iff) + qed + have bdd_above_6: \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ + by (simp add: bdd_above_4) + have Sup_2: \(SUP \\ball 0 r. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) = + max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\)\ + proof- + have \ball (0::'a) r \ {}\ + using \r > 0\ by auto + moreover have \bdd_above ((\\. \f *\<^sub>v (x + \)\) ` ball 0 r)\ + using bdd_above_5 by blast + moreover have \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ + using bdd_above_6 by blast + ultimately show ?thesis + using max_Sup + by (metis (mono_tags, lifting) Banach_Steinhaus_Missing.pointwise_max_def image_cong) + qed + have Sup_3': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \::'a + by (simp add: norm_2') + have Sup_3'': \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \::'a + by (simp add: norm_2) + have Sup_3: \max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\) = + (SUP \\ball 0 r. \f *\<^sub>v (x + \)\)\ + proof- + have \(\\. \f *\<^sub>v (x + \)\) ` (ball 0 r) = (\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)\ + apply auto using Sup_3' apply auto using Sup_3'' by blast + hence \Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))=Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + by simp + thus ?thesis by simp + qed + have Sup_1: \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ Sup ( (\\. \f *\<^sub>v \\) ` (ball x r) )\ + proof- + have \(\t. \f *\<^sub>v t\) \ \ max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\\ for \ + apply(rule linear_plus_norm) apply (rule bounded_linear.linear) + by (simp add: blinfun.bounded_linear_right) + moreover have \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + using bdd_above_1 by blast + moreover have \ball (0::'a) r \ {}\ + using \r > 0\ by auto + ultimately have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ + Sup ((\\. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ + using cSUP_mono by smt + also have \\ = max (Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))) + (Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)))\ + using Sup_2 by blast + also have \\ = Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ + using Sup_3 by blast + also have \\ = Sup ((\\. \f *\<^sub>v \\) ` (ball x r))\ + by (metis add.right_neutral ball_translation image_image) + finally show ?thesis by blast + qed + have \\f\ = (SUP x\ball 0 r. \f *\<^sub>v x\) / r\ + using \0 < r\ onorm_r by blast + moreover have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) / r \ Sup ((\\. \f *\<^sub>v \\) ` (ball x r)) / r\ + using Sup_1 \0 < r\ divide_right_mono by fastforce + ultimately have \\f\ \ Sup ((\t. \f *\<^sub>v t\) ` ball x r) / r\ + by simp + thus ?thesis by simp +qed + +lemma onorm_Sup_on_ball': + includes notation_norm + assumes \r > 0\ and \\ < 1\ + shows \\\\ball x r. \ * r * \f\ \ \f *\<^sub>v \\\ + text \ + In the proof of Banach-Steinhaus theorem, we will use this variation of the + lemma @{text onorm_Sup_on_ball}. + + Explanation: Let \<^term>\f\ be a bounded operator, let \<^term>\x\ be a point and let \<^term>\r\ be a + positive real number. For any real number \<^term>\\ < 1\, there is a point \<^term>\\\ in the open ball + of radius \<^term>\r\ around \<^term>\x\ such that \<^term>\\ * r * \f\ \ \f *\<^sub>v \\\. +\ +proof(cases \f = 0\) + case True + thus ?thesis by (metis assms(1) centre_in_ball mult_zero_right norm_zero order_refl + zero_blinfun.rep_eq) +next + case False + have bdd_above_1: \bdd_above ((\t. \(*\<^sub>v) f t\) ` ball x r)\ for f::\'a \\<^sub>L 'b\ + using assms(1) bounded_linear_image by (simp add: bounded_linear_image + blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) + have \norm f > 0\ + using \f \ 0\ by auto + have \norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) ) / r\ + using \r > 0\ by (simp add: onorm_Sup_on_ball) + hence \r * norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) )\ + using \0 < r\ by (smt divide_strict_right_mono nonzero_mult_div_cancel_left) + moreover have \\ * r * norm f < r * norm f\ + using \\ < 1\ using \0 < norm f\ \0 < r\ by auto + ultimately have \\ * r * norm f < Sup ( (norm \ ((*\<^sub>v) f)) ` (ball x r) )\ + by simp + moreover have \(norm \ ( (*\<^sub>v) f)) ` (ball x r) \ {}\ + using \0 < r\ by auto + moreover have \bdd_above ((norm \ ( (*\<^sub>v) f)) ` (ball x r))\ + using bdd_above_1 apply transfer by simp + ultimately have \\t \ (norm \ ( (*\<^sub>v) f)) ` (ball x r). \ * r * norm f < t\ + by (simp add: less_cSup_iff) + thus ?thesis by (smt comp_def image_iff) +qed + +subsection \Banach-Steinhaus theorem\ + +theorem banach_steinhaus: + fixes f::\'c \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ + assumes \\x. bounded (range (\n. (f n) *\<^sub>v x))\ + shows \bounded (range f)\ + text\ + This is Banach-Steinhaus Theorem. + + Explanation: If a family of bounded operators on a Banach space + is pointwise bounded, then it is uniformly bounded. +\ +proof(rule classical) + assume \\(bounded (range f))\ + have sum_1: \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ + proof- + have \summable (\n. (inverse (real_of_nat 3))^n)\ + using Series.summable_geometric_iff [where c = "inverse (real_of_nat 3)"] by auto + moreover have \(inverse (real_of_nat 3))^n = inverse (real_of_nat 3^n)\ for n::nat + using power_inverse by blast + ultimately have \summable (\n. inverse (real_of_nat 3^n))\ + by auto + hence \bounded (range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. + using summable_imp_sums_bounded[where f = "(\n. inverse (real_of_nat 3^n))"] + lessThan_atLeast0 by auto + hence \\M. \h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. M\ + using bounded_iff by blast + then obtain M where \h\range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ + for h + by blast + have sum_2: \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n + proof- + have \norm (sum (\ k. inverse (real 3 ^ k)) {0..< Suc n}) \ M\ + using \\h. h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ + by blast + hence \norm (sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ + by (simp add: atLeastLessThanSuc_atLeastAtMost) + hence \(sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ + by auto + thus ?thesis by blast + qed + have \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n + using sum_2 by blast + thus ?thesis by blast + qed + have \of_rat 2/3 < (1::real)\ + by auto + hence \\g::'a \\<^sub>L 'b. \x. \r. \\. g \ 0 \ r > 0 + \ (\\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g \))\ + using onorm_Sup_on_ball' by blast + hence \\\. \g::'a \\<^sub>L 'b. \x. \r. g \ 0 \ r > 0 + \ ((\ g x r)\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r)))\ + by metis + then obtain \ where f1: \\g \ 0; r > 0\ \ + \ g x r \ ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r))\ + for g::\'a \\<^sub>L 'b\ and x and r + by blast + have \\n. \k. norm (f k) \ 4^n\ + using \\(bounded (range f))\ by (metis (mono_tags, hide_lams) boundedI image_iff linear) + hence \\k. \n. norm (f (k n)) \ 4^n\ + by metis + hence \\k. \n. norm ((f \ k) n) \ 4^n\ + by simp + then obtain k where \norm ((f \ k) n) \ 4^n\ for n + by blast + define T where \T = f \ k\ + have \T n \ range f\ for n + unfolding T_def by simp + have \norm (T n) \ of_nat (4^n)\ for n + unfolding T_def using \\ n. norm ((f \ k) n) \ 4^n\ by auto + hence \T n \ 0\ for n + by (smt T_def \\n. 4 ^ n \ norm ((f \ k) n)\ norm_zero power_not_zero zero_le_power) + have \inverse (of_nat 3^n) > (0::real)\ for n + by auto + define y::\nat \ 'a\ where \y = rec_nat 0 (\n x. \ (T n) x (inverse (of_nat 3^n)))\ + have \y (Suc n) \ ball (y n) (inverse (of_nat 3^n))\ for n + using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto + hence \norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ for n + unfolding ball_def apply auto using dist_norm by (smt norm_minus_commute) + moreover have \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ + using sum_1 by blast + moreover have \Cauchy y\ + using convergent_series_Cauchy[where a = "\n. inverse (of_nat 3^n)" and \ = y] dist_norm + by (metis calculation(1) calculation(2)) + hence \\ x. y \ x\ + by (simp add: convergent_eq_Cauchy) + then obtain x where \y \ x\ + by blast + have norm_2: \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n + proof- + have \inverse (real_of_nat 3) < 1\ + by simp + moreover have \y 0 = 0\ + using y_def by auto + ultimately have \norm (x - y (Suc n)) + \ (inverse (of_nat 3)) * inverse (1 - (inverse (of_nat 3))) * ((inverse (of_nat 3)) ^ n)\ + using bound_Cauchy_to_lim[where c = "inverse (of_nat 3)" and y = y and x = x] + power_inverse semiring_norm(77) \y \ x\ + \\ n. norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ by (metis divide_inverse) + moreover have \inverse (real_of_nat 3) * inverse (1 - (inverse (of_nat 3))) + = inverse (of_nat 2)\ + by auto + ultimately show ?thesis + by (metis power_inverse) + qed + have \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n + using norm_2 by blast + have \\ M. \ n. norm ((*\<^sub>v) (T n) x) \ M\ + unfolding T_def apply auto + by (metis \\x. bounded (range (\n. (*\<^sub>v) (f n) x))\ bounded_iff rangeI) + then obtain M where \norm ((*\<^sub>v) (T n) x) \ M\ for n + by blast + have norm_1: \norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x) + \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + norm ((*\<^sub>v) (T n) x)\ for n + proof- + have \norm (y (Suc n) - x) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ + using \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ + by (simp add: norm_minus_commute) + moreover have \norm (T n) \ 0\ + by auto + ultimately have \norm (T n) * norm (y (Suc n) - x) + \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n)\ + by (simp add: \\n. T n \ 0\) + thus ?thesis by simp + qed + have inverse_2: \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) + \ norm ((*\<^sub>v) (T n) x)\ for n + proof- + have \(of_rat 2/3)*(inverse (of_nat 3^n))*norm (T n) \ norm ((*\<^sub>v) (T n) (y (Suc n)))\ + using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto + also have \\ = norm ((*\<^sub>v) (T n) ((y (Suc n) - x) + x))\ + by auto + also have \\ = norm ((*\<^sub>v) (T n) (y (Suc n) - x) + (*\<^sub>v) (T n) x)\ + apply transfer apply auto by (metis diff_add_cancel linear_simps(1)) + also have \\ \ norm ((*\<^sub>v) (T n) (y (Suc n) - x)) + norm ((*\<^sub>v) (T n) x)\ + by (simp add: norm_triangle_ineq) + also have \\ \ norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x)\ + apply transfer apply auto using onorm by auto + also have \\ \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n) + + norm ((*\<^sub>v) (T n) x)\ + using norm_1 by blast + finally have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) + \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + + norm ((*\<^sub>v) (T n) x)\ + by blast + hence \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) + - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ + by linarith + moreover have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) + - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + = (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ + by fastforce + ultimately show \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ + by linarith + qed + have inverse_3: \(inverse (of_nat 6)) * (of_rat (4/3)^n) + \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ for n + proof- + have \of_rat (4/3)^n = inverse (real 3 ^ n) * (of_nat 4^n)\ + apply auto by (metis divide_inverse_commute of_rat_divide power_divide of_rat_numeral_eq) + also have \\ \ inverse (real 3 ^ n) * norm (T n)\ + using \\n. norm (T n) \ of_nat (4^n)\ by simp + finally have \of_rat (4/3)^n \ inverse (real 3 ^ n) * norm (T n)\ + by blast + moreover have \inverse (of_nat 6) > (0::real)\ + by auto + ultimately show ?thesis by auto + qed + have inverse_1: \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n + proof- + have \(inverse (of_nat 6)) * (of_rat (4/3)^n) + \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ + using inverse_3 by blast + also have \\ \ norm ((*\<^sub>v) (T n) x)\ + using inverse_2 by blast + finally have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ norm ((*\<^sub>v) (T n) x)\ + by auto + thus ?thesis using \\ n. norm ((*\<^sub>v) (T n) x) \ M\ by smt + qed + have \\n. M < (inverse (of_nat 6)) * (of_rat (4/3)^n)\ + using Real.real_arch_pow by auto + moreover have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n + using inverse_1 by blast + ultimately show ?thesis by smt +qed + +subsection \A consequence of Banach-Steinhaus theorem\ + +corollary bounded_linear_limit_bounded_linear: + fixes f::\nat \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ + assumes \\x. convergent (\n. (f n) *\<^sub>v x)\ + shows \\g. (\n. (*\<^sub>v) (f n)) \pointwise\ (*\<^sub>v) g\ + text\ + Explanation: If a sequence of bounded operators on a Banach space converges + pointwise, then the limit is also a bounded operator. +\ +proof- + have \\l. (\n. (*\<^sub>v) (f n) x) \ l\ for x + by (simp add: \\x. convergent (\n. (*\<^sub>v) (f n) x)\ convergentD) + hence \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ + unfolding pointwise_convergent_to_def by metis + obtain F where \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ + using \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto + have \\x. (\ n. (*\<^sub>v) (f n) x) \ F x\ + using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer + by (simp add: pointwise_convergent_to_def) + have \bounded (range f)\ + using \\x. convergent (\n. (*\<^sub>v) (f n) x)\ banach_steinhaus + \\x. \l. (\n. (*\<^sub>v) (f n) x) \ l\ convergent_imp_bounded by blast + have norm_f_n: \\ M. \ n. norm (f n) \ M\ + unfolding bounded_def + by (meson UNIV_I \bounded (range f)\ bounded_iff image_eqI) + have \isCont (\ t::'b. norm t) y\ for y::'b + using Limits.isCont_norm by simp + hence \(\ n. norm ((*\<^sub>v) (f n) x)) \ (norm (F x))\ for x + using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ by (simp add: tendsto_norm) + hence norm_f_n_x: \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x + using Elementary_Metric_Spaces.convergent_imp_bounded + by (metis UNIV_I \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ bounded_iff image_eqI) + have norm_f: \\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K\ + proof- + have \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x + using norm_f_n_x \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by blast + hence \\ M. \ n. norm (f n) \ M\ + using norm_f_n by simp + then obtain M::real where \\ M. \ n. norm (f n) \ M\ + by blast + have \\ n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * norm (f n)\ + apply transfer apply auto by (metis mult.commute onorm) + thus ?thesis using \\ M. \ n. norm (f n) \ M\ + by (metis (no_types, hide_lams) dual_order.trans norm_eq_zero order_refl + real_mult_le_cancel_iff2 vector_space_over_itself.scale_zero_left zero_less_norm_iff) + qed + have norm_F_x: \\K. \x. norm (F x) \ norm x * K\ + proof- + have "\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K" + using norm_f \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by auto + thus ?thesis + using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ apply transfer + by (metis Lim_bounded tendsto_norm) + qed + have \linear F\ + proof(rule linear_limit_linear) + show \linear ((*\<^sub>v) (f n))\ for n + apply transfer apply auto by (simp add: bounded_linear.linear) + show \f \pointwise\ F\ + using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto + qed + moreover have \bounded_linear_axioms F\ + using norm_F_x by (simp add: \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ bounded_linear_axioms_def) + ultimately have \bounded_linear F\ + unfolding bounded_linear_def by blast + hence \\g. (*\<^sub>v) g = F\ + using bounded_linear_Blinfun_apply by auto + thus ?thesis using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer by auto +qed + +end diff --git a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy --- a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy +++ b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy @@ -1,898 +1,898 @@ -(* - File: Banach_Steinhaus_Missing.thy - Author: Dominique Unruh, University of Tartu - Author: Jose Manuel Rodriguez Caballero, University of Tartu -*) -section \Missing results for the proof of Banach-Steinhaus theorem\ - -theory Banach_Steinhaus_Missing - imports - "HOL-Analysis.Infinite_Set_Sum" - -begin -subsection \Results missing for the proof of Banach-Steinhaus theorem\ -text \ - The results proved here are preliminaries for the proof of Banach-Steinhaus theorem using Sokal's - approach, but they do not explicitly appear in Sokal's paper ~\cite{sokal2011reall}. -\ - -text\Notation for the norm\ -bundle notation_norm begin -notation norm ("\_\") -end - -bundle no_notation_norm begin -no_notation norm ("\_\") -end - -unbundle notation_norm - -text\Notation for apply bilinear function\ -bundle notation_blinfun_apply begin -notation blinfun_apply (infixr "*\<^sub>v" 70) -end - -bundle no_notation_blinfun_apply begin -no_notation blinfun_apply (infixr "*\<^sub>v" 70) -end - -unbundle notation_blinfun_apply - -lemma bdd_above_plus: - fixes f::\'a \ real\ - assumes \bdd_above (f ` S)\ and \bdd_above (g ` S)\ - shows \bdd_above ((\ x. f x + g x) ` S)\ - text \ - Explanation: If the images of two real-valued functions \<^term>\f\,\<^term>\g\ are bounded above on a - set \<^term>\S\, then the image of their sum is bounded on \<^term>\S\. -\ -proof- - obtain M where \\ x. x\S \ f x \ M\ - using \bdd_above (f ` S)\ unfolding bdd_above_def by blast - obtain N where \\ x. x\S \ g x \ N\ - using \bdd_above (g ` S)\ unfolding bdd_above_def by blast - have \\ x. x\S \ f x + g x \ M + N\ - using \\x. x \ S \ f x \ M\ \\x. x \ S \ g x \ N\ by fastforce - thus ?thesis unfolding bdd_above_def by blast -qed - -text\The maximum of two functions\ -definition pointwise_max:: "('a \ 'b::ord) \ ('a \ 'b) \ ('a \ 'b)" where - \pointwise_max f g = (\x. max (f x) (g x))\ - -lemma max_Sup_absorb_left: - fixes f g::\'a \ real\ - assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ - shows \Sup ((pointwise_max f g) ` X) = Sup (f ` X)\ - - text \Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\, if the supremum of \<^term>\f\ is - greater-equal the supremum of \<^term>\g\, then the supremum of \<^term>\max f g\ equals the supremum of - \<^term>\f\. (Under some technical conditions.)\ - -proof- - have y_Sup: \y \ ((\ x. max (f x) (g x)) ` X) \ y \ Sup (f ` X)\ for y - proof- - assume \y \ ((\ x. max (f x) (g x)) ` X)\ - then obtain x where \y = max (f x) (g x)\ and \x \ X\ - by blast - have \f x \ Sup (f ` X)\ - by (simp add: \x \ X\ \bdd_above (f ` X)\ cSUP_upper) - moreover have \g x \ Sup (g ` X)\ - by (simp add: \x \ X\ \bdd_above (g ` X)\ cSUP_upper) - ultimately have \max (f x) (g x) \ Sup (f ` X)\ - using \Sup (f ` X) \ Sup (g ` X)\ by auto - thus ?thesis by (simp add: \y = max (f x) (g x)\) - qed - have y_f_X: \y \ f ` X \ y \ Sup ((\ x. max (f x) (g x)) ` X)\ for y - proof- - assume \y \ f ` X\ - then obtain x where \x \ X\ and \y = f x\ - by blast - have \bdd_above ((\ \. max (f \) (g \)) ` X)\ - by (metis (no_types) \bdd_above (f ` X)\ \bdd_above (g ` X)\ bdd_above_image_sup sup_max) - moreover have \e > 0 \ \ k \ (\ \. max (f \) (g \)) ` X. y \ k + e\ - for e::real - using \Sup (f ` X) \ Sup (g ` X)\ by (smt \x \ X\ \y = f x\ image_eqI) - ultimately show ?thesis - using \x \ X\ \y = f x\ cSUP_upper by fastforce - qed - have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ - using y_Sup by (simp add: \X \ {}\ cSup_least) - moreover have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ - using y_f_X by (metis (mono_tags) cSup_least calculation empty_is_image) - ultimately show ?thesis unfolding pointwise_max_def by simp -qed - -lemma max_Sup_absorb_right: - fixes f g::\'a \ real\ - assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ - shows \Sup ((pointwise_max f g) ` X) = Sup (g ` X)\ - text \ - Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\ and a nonempty set \<^term>\X\, such that - the \<^term>\f\ and \<^term>\g\ are bounded above on \<^term>\X\, if the supremum of \<^term>\f\ on \<^term>\X\ is - lower-equal the supremum of \<^term>\g\ on \<^term>\X\, then the supremum of \<^term>\pointwise_max f g\ on \<^term>\X\ - equals the supremum of \<^term>\g\. This is the right analog of @{text max_Sup_absorb_left}. -\ -proof- - have \Sup ((pointwise_max g f) ` X) = Sup (g ` X)\ - using assms by (simp add: max_Sup_absorb_left) - moreover have \pointwise_max g f = pointwise_max f g\ - unfolding pointwise_max_def by auto - ultimately show ?thesis by simp -qed - -lemma max_Sup: - fixes f g::\'a \ real\ - assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ - shows \Sup ((pointwise_max f g) ` X) = max (Sup (f ` X)) (Sup (g ` X))\ - text \ - Explanation: Let \<^term>\X\ be a nonempty set. Two supremum over \<^term>\X\ of the maximum of two - real-value functions is equal to the maximum of their suprema over \<^term>\X\, provided that the - functions are bounded above on \<^term>\X\. -\ -proof(cases \Sup (f ` X) \ Sup (g ` X)\) - case True thus ?thesis by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_left) -next - case False - have f1: "\ 0 \ Sup (f ` X) + - 1 * Sup (g ` X)" - using False by linarith - hence "Sup (Banach_Steinhaus_Missing.pointwise_max f g ` X) = Sup (g ` X)" - by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_right) - thus ?thesis - using f1 by linarith -qed - - -lemma identity_telescopic: - fixes x :: \_ \ 'a::real_normed_vector\ - assumes \x \ l\ - shows \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) \ l - x n\ - text\ - Expression of a limit as a telescopic series. - Explanation: If \<^term>\x\ converges to \<^term>\l\ then the sum \<^term>\sum (\ k. x (Suc k) - x k) {n..N}\ - converges to \<^term>\l - x n\ as \<^term>\N\ goes to infinity. -\ -proof- - have \(\ p. x (p + Suc n)) \ l\ - using \x \ l\ by (rule LIMSEQ_ignore_initial_segment) - hence \(\ p. x (Suc n + p)) \ l\ - by (simp add: add.commute) - hence \(\ p. x (Suc (n + p))) \ l\ - by simp - hence \(\ t. (- (x n)) + (\ p. x (Suc (n + p))) t ) \ (- (x n)) + l\ - using tendsto_add_const_iff by metis - hence f1: \(\ p. x (Suc (n + p)) - x n)\ l - x n\ - by simp - have \sum (\ k. x (Suc k) - x k) {n..n+p} = x (Suc (n+p)) - x n\ for p - by (simp add: sum_Suc_diff) - moreover have \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + t) - = (\ p. sum (\ k. x (Suc k) - x k) {n..n+p}) t\ for t - by blast - ultimately have \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + p)) \ l - x n\ - using f1 by simp - hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (p + n)) \ l - x n\ - by (simp add: add.commute) - hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) p) \ l - x n\ - using Topological_Spaces.LIMSEQ_offset[where f = "(\ N. sum (\ k. x (Suc k) - x k) {n..N})" - and a = "l - x n" and k = n] by blast - hence \(\ M. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) M) \ l - x n\ - by simp - thus ?thesis by blast -qed - -lemma bound_Cauchy_to_lim: - assumes \y \ x\ and \\n. \y (Suc n) - y n\ \ c^n\ and \y 0 = 0\ and \c < 1\ - shows \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ - text\ - Inequality about a sequence of approximations assuming that the sequence of differences is bounded - by a geometric progression. - Explanation: Let \<^term>\y\ be a sequence converging to \<^term>\x\. - If \<^term>\y\ satisfies the inequality \\y (Suc n) - y n\ \ c ^ n\ for some \<^term>\c < 1\ and - assuming \<^term>\y 0 = 0\ then the inequality \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ holds. -\ -proof- - have \c \ 0\ - using \\ n. \y (Suc n) - y n\ \ c^n\ by (smt norm_imp_pos_and_ge power_Suc0_right) - have norm_1: \norm (\k = Suc n..N. y (Suc k) - y k) \ (c ^ Suc n)/(1 - c)\ for N - proof(cases \N < Suc n\) - case True - hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ = 0\ - by auto - thus ?thesis using \c \ 0\ \c < 1\ by auto - next - case False - hence \N \ Suc n\ - by auto - have \c^(Suc N) \ 0\ - using \c \ 0\ by auto - have \1 - c > 0\ - by (simp add: \c < 1\) - hence \(1 - c)/(1 - c) = 1\ - by auto - have \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (\k. \y (Suc k) - y k\) {Suc n .. N})\ - by (simp add: sum_norm_le) - hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (power c) {Suc n .. N})\ - by (simp add: assms(2) sum_norm_le) - hence \(1 - c) * \sum (\k. y (Suc k) - y k) {Suc n .. N}\ - \ (1 - c) * (sum (power c) {Suc n .. N})\ - using \0 < 1 - c\ real_mult_le_cancel_iff2 by blast - also have \\ = c^(Suc n) - c^(Suc N)\ - using Set_Interval.sum_gp_multiplied \Suc n \ N\ by blast - also have \\ \ c^(Suc n)\ - using \c^(Suc N) \ 0\ by auto - finally have \(1 - c) * \\k = Suc n..N. y (Suc k) - y k\ \ c ^ Suc n\ - by blast - hence \((1 - c) * \\k = Suc n..N. y (Suc k) - y k\)/(1 - c) - \ (c ^ Suc n)/(1 - c)\ - using \0 < 1 - c\ by (smt divide_right_mono) - thus \\\k = Suc n..N. y (Suc k) - y k\ \ (c ^ Suc n)/(1 - c)\ - using \0 < 1 - c\ by auto - qed - have \(\ N. (sum (\k. y (Suc k) - y k) {Suc n .. N})) \ x - y (Suc n)\ - by (metis (no_types) \y \ x\ identity_telescopic) - hence \(\ N. \sum (\k. y (Suc k) - y k) {Suc n .. N}\) \ \x - y (Suc n)\\ - using tendsto_norm by blast - hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ - using norm_1 Lim_bounded by blast - hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ - by auto - moreover have \(c ^ Suc n)/(1 - c) = (c / (1 - c)) * (c ^ n)\ - by (simp add: divide_inverse_commute) - ultimately show \\x - y (Suc n)\ \ (c / (1 - c)) * (c ^ n)\ by linarith -qed - -lemma onorm_open_ball: - includes notation_norm - shows \\f\ = Sup { \f *\<^sub>v x\ | x. \x\ < 1 }\ - text \ - Explanation: Let \<^term>\f\ be a bounded linear operator. The operator norm of \<^term>\f\ is the - supremum of \<^term>\norm (f x)\ for \<^term>\x\ such that \<^term>\norm x < 1\. -\ -proof(cases \(UNIV::'a set) = 0\) - case True - hence \x = 0\ for x::'a - by auto - hence \f *\<^sub>v x = 0\ for x - by (metis (full_types) blinfun.zero_right) - hence \\f\ = 0\ - by (simp add: blinfun_eqI zero_blinfun.rep_eq) - have \{ \f *\<^sub>v x\ | x. \x\ < 1} = {0}\ - by (smt Collect_cong \\x. f *\<^sub>v x = 0\ norm_zero singleton_conv) - hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} = 0\ - by simp - thus ?thesis using \\f\ = 0\ by auto -next - case False - hence \(UNIV::'a set) \ 0\ - by simp - have nonnegative: \\f *\<^sub>v x\ \ 0\ for x - by simp - have \\ x::'a. x \ 0\ - using \UNIV \ 0\ by auto - then obtain x::'a where \x \ 0\ - by blast - hence \\x\ \ 0\ - by auto - define y where \y = x /\<^sub>R \x\\ - have \norm y = \ x /\<^sub>R \x\ \\ - unfolding y_def by auto - also have \\ = \x\ /\<^sub>R \x\\ - by auto - also have \\ = 1\ - using \\x\ \ 0\ by auto - finally have \\y\ = 1\ - by blast - hence norm_1_non_empty: \{ \f *\<^sub>v x\ | x. \x\ = 1} \ {}\ - by blast - have norm_1_bounded: \bdd_above { \f *\<^sub>v x\ | x. \x\ = 1}\ - unfolding bdd_above_def apply auto - by (metis norm_blinfun) - have norm_less_1_non_empty: \{\f *\<^sub>v x\ | x. \x\ < 1} \ {}\ - by (metis (mono_tags, lifting) Collect_empty_eq_bot bot_empty_eq empty_iff norm_zero - zero_less_one) - have norm_less_1_bounded: \bdd_above {\f *\<^sub>v x\ | x. \x\ < 1}\ - proof- - have \\r. \a r\ < 1 \ \f *\<^sub>v (a r)\ \ r\ for a :: "real \ 'a" - proof- - obtain r :: "('a \\<^sub>L 'b) \ real" where - "\f x. 0 \ r f \ (bounded_linear f \ \f *\<^sub>v x\ \ \x\ * r f)" - using bounded_linear.nonneg_bounded by moura - have \\ \f\ < 0\ - by simp - hence "(\r. \f\ * \a r\ \ r) \ (\r. \a r\ < 1 \ \f *\<^sub>v a r\ \ r)" - by (meson less_eq_real_def mult_le_cancel_left2) - thus ?thesis using dual_order.trans norm_blinfun by blast - qed - hence \\ M. \ x. \x\ < 1 \ \f *\<^sub>v x\ \ M\ - by metis - thus ?thesis by auto - qed - have Sup_non_neg: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} \ 0\ - by (smt Collect_empty_eq cSup_upper mem_Collect_eq nonnegative norm_1_bounded norm_1_non_empty) - have \{0::real} \ {}\ - by simp - have \bdd_above {0::real}\ - by simp - show \\f\ = Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ - proof(cases \\x. f *\<^sub>v x = 0\) - case True - have \\f *\<^sub>v x\ = 0\ for x - by (simp add: True) - hence \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ - by blast - moreover have \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ - using calculation norm_less_1_non_empty by fastforce - ultimately have \{\f *\<^sub>v x\ | x. \x\ < 1 } = {0}\ - by blast - hence Sup1: \Sup {\f *\<^sub>v x\ | x. \x\ < 1 } = 0\ - by simp - have \\f\ = 0\ - by (simp add: True blinfun_eqI) - moreover have \Sup {\f *\<^sub>v x\ | x. \x\ < 1} = 0\ - using Sup1 by blast - ultimately show ?thesis by simp - next - case False - have norm_f_eq_leq: \y \ {\f *\<^sub>v x\ | x. \x\ = 1} \ - y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for y - proof- - assume \y \ {\f *\<^sub>v x\ | x. \x\ = 1}\ - hence \\ x. y = \f *\<^sub>v x\ \ \x\ = 1\ - by blast - then obtain x where \y = \f *\<^sub>v x\\ and \\x\ = 1\ - by auto - define y' where \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R y\ for n - have \y' n \ {\f *\<^sub>v x\ | x. \x\ < 1}\ for n - proof- - have \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R \f *\<^sub>v x\\ - using y'_def \y = \f *\<^sub>v x\\ by blast - also have \... = \(1 - (inverse (real (Suc n))))\ *\<^sub>R \f *\<^sub>v x\\ - by (metis (mono_tags, hide_lams) \y = \f *\<^sub>v x\\ abs_1 abs_le_self_iff abs_of_nat - abs_of_nonneg add_diff_cancel_left' add_eq_if cancel_comm_monoid_add_class.diff_cancel - diff_ge_0_iff_ge eq_iff_diff_eq_0 inverse_1 inverse_le_iff_le nat.distinct(1) of_nat_0 - of_nat_Suc of_nat_le_0_iff zero_less_abs_iff zero_neq_one) - also have \... = \f *\<^sub>v ((1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ - by (simp add: blinfun.scaleR_right) - finally have y'_1: \y' n = \f *\<^sub>v ( (1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ - by blast - have \\(1 - (inverse (Suc n))) *\<^sub>R x\ = (1 - (inverse (real (Suc n)))) * \x\\ - by (simp add: linordered_field_class.inverse_le_1_iff) - hence \\(1 - (inverse (Suc n))) *\<^sub>R x\ < 1\ - by (simp add: \\x\ = 1\) - thus ?thesis using y'_1 by blast - qed - have \(\n. (1 - (inverse (real (Suc n)))) ) \ 1\ - using Limits.LIMSEQ_inverse_real_of_nat_add_minus by simp - hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ 1 *\<^sub>R y\ - using Limits.tendsto_scaleR by blast - hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ y\ - by simp - hence \(\n. y' n) \ y\ - using y'_def by simp - hence \y' \ y\ - by simp - have \y' n \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for n - using cSup_upper \\n. y' n \ {\f *\<^sub>v x\ |x. \x\ < 1}\ norm_less_1_bounded by blast - hence \y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ - using \y' \ y\ Topological_Spaces.Sup_lim by (meson LIMSEQ_le_const2) - thus ?thesis by blast - qed - hence \Sup {\f *\<^sub>v x\ | x. \x\ = 1} \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ - by (metis (lifting) cSup_least norm_1_non_empty) - have \y \ {\f *\<^sub>v x\ | x. \x\ < 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ for y - proof(cases \y = 0\) - case True thus ?thesis by (simp add: Sup_non_neg) - next - case False - hence \y \ 0\ by blast - assume \y \ {\f *\<^sub>v x\ | x. \x\ < 1}\ - hence \\ x. y = \f *\<^sub>v x\ \ \x\ < 1\ - by blast - then obtain x where \y = \f *\<^sub>v x\\ and \\x\ < 1\ - by blast - have \(1/\x\) * y = (1/\x\) * \f x\\ - by (simp add: \y = \f *\<^sub>v x\\) - also have \... = \1/\x\\ * \f *\<^sub>v x\\ - by simp - also have \... = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ - by simp - also have \... = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ - by (simp add: blinfun.scaleR_right) - finally have \(1/\x\) * y = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ - by blast - have \x \ 0\ - using \y \ 0\ \y = \f *\<^sub>v x\\ blinfun.zero_right by auto - have \\ (1/\x\) *\<^sub>R x \ = \ (1/\x\) \ * \x\\ - by simp - also have \... = (1/\x\) * \x\\ - by simp - finally have \\(1/\x\) *\<^sub>R x\ = 1\ - using \x \ 0\ by simp - hence \(1/\x\) * y \ { \f *\<^sub>v x\ | x. \x\ = 1}\ - using \1 / \x\ * y = \f *\<^sub>v (1 / \x\) *\<^sub>R x\\ by blast - hence \(1/\x\) * y \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ - by (simp add: cSup_upper norm_1_bounded) - moreover have \y \ (1/\x\) * y\ - by (metis \\x\ < 1\ \y = \f *\<^sub>v x\\ mult_le_cancel_right1 norm_not_less_zero - order.strict_implies_order \x \ 0\ less_divide_eq_1_pos zero_less_norm_iff) - ultimately show ?thesis by linarith - qed - hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ - by (smt cSup_least norm_less_1_non_empty) - hence \Sup { \f *\<^sub>v x\ | x. \x\ = 1} = Sup { \f *\<^sub>v x\ | x. \x\ < 1}\ - using \Sup {\f *\<^sub>v x\ |x. norm x = 1} \ Sup { \f *\<^sub>v x\ |x. \x\ < 1}\ by linarith - have f1: \(SUP x. \f *\<^sub>v x\ / \x\) = Sup { \f *\<^sub>v x\ / \x\ | x. True}\ - by (simp add: full_SetCompr_eq) - have \y \ { \f *\<^sub>v x\ / \x\ |x. True} \ y \ { \f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ - for y - proof- - assume \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ show ?thesis - proof(cases \y = 0\) - case True thus ?thesis by simp - next - case False - have \\ x. y = \f *\<^sub>v x\ / \x\\ - using \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ by auto - then obtain x where \y = \f *\<^sub>v x\ / \x\\ - by blast - hence \y = \(1/\x\)\ * \ f *\<^sub>v x \\ - by simp - hence \y = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ - by simp - hence \y = \f ((1/\x\) *\<^sub>R x)\\ - by (simp add: blinfun.scaleR_right) - moreover have \\ (1/\x\) *\<^sub>R x \ = 1\ - using False \y = \f *\<^sub>v x\ / \x\\ by auto - ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ - by blast - thus ?thesis by blast - qed - qed - moreover have \y \ {\f x\ |x. \x\ = 1} \ {0} \ y \ {\f *\<^sub>v x\ / \x\ |x. True}\ - for y - proof(cases \y = 0\) - case True thus ?thesis by auto - next - case False - hence \y \ {0}\ - by simp - moreover assume \y \ {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ - ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ - by simp - then obtain x where \\x\ = 1\ and \y = \f *\<^sub>v x\\ - by auto - have \y = \f *\<^sub>v x\ / \x\\ using \\x\ = 1\ \y = \f *\<^sub>v x\\ - by simp - thus ?thesis by auto - qed - ultimately have \{\f *\<^sub>v x\ / \x\ |x. True} = {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ - by blast - hence \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ - by simp - have "\r s. \ (r::real) \ s \ sup r s = s" - by (metis (lifting) sup.absorb_iff1 sup_commute) - hence \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {(0::real)}) - = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0::real})\ - using \0 \ Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ \bdd_above {0}\ \{0} \ {}\ cSup_singleton - cSup_union_distrib max.absorb_iff1 sup_commute norm_1_bounded norm_1_non_empty - by (metis (no_types, lifting) ) - moreover have \Sup {(0::real)} = (0::real)\ - by simp - ultimately have \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ - using Sup_non_neg by linarith - moreover have \Sup ( {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) - = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0}) \ - using Sup_non_neg \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) - = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0})\ - by auto - ultimately have f2: \Sup {\f *\<^sub>v x\ / \x\ | x. True} = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ - using \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by linarith - have \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ - using f1 f2 by linarith - hence \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ < 1 }\ - by (simp add: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\) - thus ?thesis apply transfer by (simp add: onorm_def) - qed -qed - -lemma onorm_r: - includes notation_norm - assumes \r > 0\ - shows \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ - text \ - Explanation: The norm of \<^term>\f\ is \<^term>\1/r\ of the supremum of the norm of \<^term>\f *\<^sub>v x\ for - \<^term>\x\ in the ball of radius \<^term>\r\ centered at the origin. -\ -proof- - have \\f\ = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\ - using onorm_open_ball by blast - moreover have \{\f *\<^sub>v x\ |x. \x\ < 1} = (\x. \f *\<^sub>v x\) ` (ball 0 1)\ - unfolding ball_def by auto - ultimately have onorm_f: \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 1))\ - by simp - have s2: \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1 \ x \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1)\ for x - proof- - assume \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ - hence \\ t. x = r *\<^sub>R \f *\<^sub>v t\ \ \t\ < 1\ - by auto - then obtain t where \x = r *\<^sub>R \f *\<^sub>v t\\ and \\t\ < 1\ - by blast - define y where \y = x /\<^sub>R r\ - have \x = r * (inverse r * x)\ - using \x = r *\<^sub>R norm (f t)\ by auto - hence \x - (r * (inverse r * x)) \ 0\ - by linarith - hence \x \ r * (x /\<^sub>R r)\ - by auto - have \y \ (\k. \f *\<^sub>v k\) ` ball 0 1\ - unfolding y_def by (smt \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ assms image_iff - inverse_inverse_eq pos_le_divideR_eq positive_imp_inverse_positive) - moreover have \x \ r * y\ - using \x \ r * (x /\<^sub>R r)\ y_def by blast - ultimately have y_norm_f: \y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ - by blast - have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ - by simp - moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ - by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above - bounded_norm_comp) - moreover have \\ y. y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ - using y_norm_f by blast - ultimately show ?thesis - by (smt \0 < r\ cSup_upper ordered_comm_semiring_class.comm_mult_left_mono) - qed - have s3: \(\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y) \ - r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y\ for y - proof- - assume \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ - have x_leq: \x \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ y / r\ for x - proof- - assume \x \ (\t. \f *\<^sub>v t\) ` ball 0 1\ - then obtain t where \t \ ball (0::'a) 1\ and \x = \f *\<^sub>v t\\ - by auto - define x' where \x' = r *\<^sub>R x\ - have \x' = r * \f *\<^sub>v t\\ - by (simp add: \x = \f *\<^sub>v t\\ x'_def) - hence \x' \ (\t. r * \f *\<^sub>v t\) ` ball 0 1\ - using \t \ ball (0::'a) 1\ by auto - hence \x' \ y\ - using \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ by blast - thus \x \ y / r\ - unfolding x'_def using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) - qed - have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ - by simp - moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ - by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above - bounded_norm_comp) - ultimately have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y/r\ - using x_leq by (simp add: \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ cSup_least) - thus ?thesis using \r > 0\ - by (smt divide_strict_right_mono nonzero_mult_div_cancel_left) - qed - have norm_scaleR: \norm \ ((*\<^sub>R) r) = ((*\<^sub>R) \r\) \ (norm::'a \ real)\ - by auto - have f_x1: \f (r *\<^sub>R x) = r *\<^sub>R f x\ for x - by (simp add: blinfun.scaleR_right) - have \ball (0::'a) r = ((*\<^sub>R) r) ` (ball 0 1)\ - by (smt assms ball_scale nonzero_mult_div_cancel_left right_inverse_eq scale_zero_right) - hence \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) = Sup ((\t. \f *\<^sub>v t\) ` (((*\<^sub>R) r) ` (ball 0 1)))\ - by simp - also have \\ = Sup (((\t. \f *\<^sub>v t\) \ ((*\<^sub>R) r)) ` (ball 0 1))\ - using Sup.SUP_image by auto - also have \\ = Sup ((\t. \f *\<^sub>v (r *\<^sub>R t)\) ` (ball 0 1))\ - using f_x1 by (simp add: comp_assoc) - also have \\ = Sup ((\t. \r\ *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ - using norm_scaleR f_x1 by auto - also have \\ = Sup ((\t. r *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ - using \r > 0\ by auto - also have \\ = r * Sup ((\t. \f *\<^sub>v t\) ` (ball 0 1))\ - apply (rule cSup_eq_non_empty) apply simp using s2 apply auto using s3 by auto - also have \\ = r * \f\\ - using onorm_f by auto - finally have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 r) = r * \f\\ - by blast - thus \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ using \r > 0\ by simp -qed - -text\Pointwise convergence\ -definition pointwise_convergent_to :: - \( nat \ ('a \ 'b::topological_space) ) \ ('a \ 'b) \ bool\ - (\((_)/ \pointwise\ (_))\ [60, 60] 60) where - \pointwise_convergent_to x l = (\ t::'a. (\ n. (x n) t) \ l t)\ - -lemma linear_limit_linear: - fixes f :: \_ \ ('a::real_vector \ 'b::real_normed_vector)\ - assumes \\n. linear (f n)\ and \f \pointwise\ F\ - shows \linear F\ - text\ - Explanation: If a family of linear operators converges pointwise, then the limit is also a linear - operator. -\ -proof - show "F (x + y) = F x + F y" for x y - proof- - have "\a. F a = lim (\n. f n a)" - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (full_types) limI) - moreover have "\f b c g. (lim (\n. g n + f n) = (b::'b) + c \ \ f \ c) \ \ g \ b" - by (metis (no_types) limI tendsto_add) - moreover have "\a. (\n. f n a) \ F a" - using assms(2) pointwise_convergent_to_def by force - ultimately have - lim_sum: \lim (\ n. (f n) x + (f n) y) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ - by metis - have \(f n) (x + y) = (f n) x + (f n) y\ for n - using \\ n. linear (f n)\ unfolding linear_def using Real_Vector_Spaces.linear_iff assms(1) - by auto - hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x + (f n) y)\ - by simp - hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ - using lim_sum by simp - moreover have \(\ n. (f n) (x + y)) \ F (x + y)\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast - moreover have \(\ n. (f n) x) \ F x\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast - moreover have \(\ n. (f n) y) \ F y\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast - ultimately show ?thesis - by (metis limI) - qed - show "F (r *\<^sub>R x) = r *\<^sub>R F x" for r and x - proof- - have \(f n) (r *\<^sub>R x) = r *\<^sub>R (f n) x\ for n - using \\ n. linear (f n)\ - by (simp add: Real_Vector_Spaces.linear_def real_vector.linear_scale) - hence \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ - by simp - have \convergent (\ n. (f n) x)\ - by (metis assms(2) convergentI pointwise_convergent_to_def) - moreover have \isCont (\ t::'b. r *\<^sub>R t) tt\ for tt - by (simp add: bounded_linear_scaleR_right) - ultimately have \lim (\ n. r *\<^sub>R ((f n) x)) = r *\<^sub>R lim (\ n. (f n) x)\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def - by (metis (mono_tags) isCont_tendsto_compose limI) - hence \lim (\ n. (f n) (r *\<^sub>R x)) = r *\<^sub>R lim (\ n. (f n) x)\ - using \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp - moreover have \(\ n. (f n) x) \ F x\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast - moreover have \(\ n. (f n) (r *\<^sub>R x)) \ F (r *\<^sub>R x)\ - using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast - ultimately show ?thesis - by (metis limI) - qed -qed - - -lemma non_Cauchy_unbounded: - fixes a ::\_ \ real\ - assumes \\n. a n \ 0\ and \e > 0\ - and \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ - shows \(\n. (sum a {0..n})) \ \\ - text\ - Explanation: If the sequence of partial sums of nonnegative terms is not Cauchy, then it converges - to infinite. -\ -proof- - define S::"ereal set" where \S = range (\n. sum a {0..n})\ - have \\s\S. k*e \ s\ for k::nat - proof(induction k) - case 0 - from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ - obtain m n where \m \ 0\ and \n \ 0\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast - have \n < Suc n\ - by simp - hence \{0..n} \ {Suc n..m} = {0..m}\ - using Set_Interval.ivl_disj_un(7) \n < m\ by auto - moreover have \finite {0..n}\ - by simp - moreover have \finite {Suc n..m}\ - by simp - moreover have \{0..n} \ {Suc n..m} = {}\ - by simp - ultimately have \sum a {0..n} + sum a {Suc n..m} = sum a {0..m}\ - by (metis sum.union_disjoint) - moreover have \sum a {Suc n..m} > 0\ - using \e > 0\ \sum a {Suc n..m} \ e\ by linarith - moreover have \sum a {0..n} \ 0\ - by (simp add: assms(1) sum_nonneg) - ultimately have \sum a {0..m} > 0\ - by linarith - moreover have \sum a {0..m} \ S\ - unfolding S_def by blast - ultimately have \\s\S. 0 \ s\ - using ereal_less_eq(5) by fastforce - thus ?case - by (simp add: zero_ereal_def) - next - case (Suc k) - assume \\s\S. k*e \ s\ - then obtain s where \s\S\ and \ereal (k * e) \ s\ - by blast - have \\N. s = sum a {0..N}\ - using \s\S\ unfolding S_def by blast - then obtain N where \s = sum a {0..N}\ - by blast - from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ - obtain m n where \m \ Suc N\ and \n \ Suc N\ and \m > n\ and \sum a {Suc n..m} \ e\ - by blast - have \finite {Suc N..n}\ - by simp - moreover have \finite {Suc n..m}\ - by simp - moreover have \{Suc N..n} \ {Suc n..m} = {Suc N..m}\ - using Set_Interval.ivl_disj_un - by (smt \Suc N \ n\ \n < m\ atLeastSucAtMost_greaterThanAtMost less_imp_le_nat) - moreover have \{} = {Suc N..n} \ {Suc n..m}\ - by simp - ultimately have \sum a {Suc N..m} = sum a {Suc N..n} + sum a {Suc n..m}\ - by (metis sum.union_disjoint) - moreover have \sum a {Suc N..n} \ 0\ - using \\n. a n \ 0\ by (simp add: sum_nonneg) - ultimately have \sum a {Suc N..m} \ e\ - using \e \ sum a {Suc n..m}\ by linarith - have \finite {0..N}\ - by simp - have \finite {Suc N..m}\ - by simp - moreover have \{0..N} \ {Suc N..m} = {0..m}\ - using Set_Interval.ivl_disj_un(7) \Suc N \ m\ by auto - moreover have \{0..N} \ {Suc N..m} = {}\ - by simp - ultimately have \sum a {0..N} + sum a {Suc N..m} = sum a {0..m}\ - by (metis \finite {0..N}\ sum.union_disjoint) - hence \e + k * e \ sum a {0..m}\ - using \ereal (real k * e) \ s\ \s = ereal (sum a {0..N})\ \e \ sum a {Suc N..m}\ by auto - moreover have \e + k * e = (Suc k) * e\ - by (simp add: semiring_normalization_rules(3)) - ultimately have \(Suc k) * e \ sum a {0..m}\ - by linarith - hence \ereal ((Suc k) * e) \ sum a {0..m}\ - by auto - moreover have \sum a {0..m}\S\ - unfolding S_def by blast - ultimately show ?case by blast - qed - hence \\s\S. (real n) \ s\ for n - by (meson assms(2) ereal_le_le ex_less_of_nat_mult less_le_not_le) - hence \Sup S = \\ - using Sup_le_iff Sup_subset_mono dual_order.strict_trans1 leD less_PInf_Ex_of_nat subsetI - by metis - hence Sup: \Sup ((range (\ n. (sum a {0..n})))::ereal set) = \\ using S_def - by blast - have \incseq (\n. (sum a {.. - using \\n. a n \ 0\ using Extended_Real.incseq_sumI by auto - hence \incseq (\n. (sum a {..< Suc n}))\ - by (meson incseq_Suc_iff) - hence \incseq (\n. (sum a {0..n})::ereal)\ - using incseq_ereal by (simp add: atLeast0AtMost lessThan_Suc_atMost) - hence \(\n. sum a {0..n}) \ Sup (range (\n. (sum a {0..n})::ereal))\ - using LIMSEQ_SUP by auto - thus ?thesis using Sup PInfty_neq_ereal by auto -qed - -lemma sum_Cauchy_positive: - fixes a ::\_ \ real\ - assumes \\n. a n \ 0\ and \\K. \n. (sum a {0..n}) \ K\ - shows \Cauchy (\n. sum a {0..n})\ - text\ - Explanation: If a series of nonnegative reals is bounded, then the series is - Cauchy. -\ -proof (unfold Cauchy_altdef2, rule, rule) - fix e::real - assume \e>0\ - have \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ - proof(rule classical) - assume \\(\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e)\ - hence \\M. \m. \n. m \ M \ n \ M \ m > n \ \(sum a {Suc n..m} < e)\ - by blast - hence \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ - by fastforce - hence \(\n. (sum a {0..n}) ) \ \\ - using non_Cauchy_unbounded \0 < e\ assms(1) by blast - from \\K. \n. sum a {0..n} \ K\ - obtain K where \\n. sum a {0..n} \ K\ - by blast - from \(\n. sum a {0..n}) \ \\ - have \\B. \N. \n\N. (\ n. (sum a {0..n}) ) n \ B\ - using Lim_PInfty by simp - hence \\n. (sum a {0..n}) \ K+1\ - using ereal_less_eq(3) by blast - thus ?thesis using \\n. (sum a {0..n}) \ K\ by smt - qed - have \sum a {Suc n..m} = sum a {0..m} - sum a {0..n}\ - if "m > n" for m n - apply (simp add: that atLeast0AtMost) using sum_up_index_split - by (smt less_imp_add_positive that) - hence \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ - using \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by smt - from \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ - obtain M where \\m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ - by blast - moreover have \m > n \ sum a {0..m} \ sum a {0..n}\ for m n - using \\ n. a n \ 0\ by (simp add: sum_mono2) - ultimately have \\M. \m\M. \n\M. m > n \ \sum a {0..m} - sum a {0..n}\ < e\ - by auto - hence \\M. \m\M. \n\M. m \ n \ \sum a {0..m} - sum a {0..n}\ < e\ - by (metis \0 < e\ abs_zero cancel_comm_monoid_add_class.diff_cancel diff_is_0_eq' - less_irrefl_nat linorder_neqE_nat zero_less_diff) - hence \\M. \m\M. \n\M. \sum a {0..m} - sum a {0..n}\ < e\ - by (metis abs_minus_commute nat_le_linear) - hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ - by (simp add: dist_real_def) - hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by blast - thus \\N. \n\N. dist (sum a {0..n}) (sum a {0..N}) < e\ by auto -qed - -lemma convergent_series_Cauchy: - fixes a::\nat \ real\ and \::\nat \ 'a::metric_space\ - assumes \\M. \n. sum a {0..n} \ M\ and \\n. dist (\ (Suc n)) (\ n) \ a n\ - shows \Cauchy \\ - text\ - Explanation: Let \<^term>\a\ be a real-valued sequence and let \<^term>\\\ be sequence in a metric space. - If the partial sums of \<^term>\a\ are uniformly bounded and the distance between consecutive terms of \<^term>\\\ - are bounded by the sequence \<^term>\a\, then \<^term>\\\ is Cauchy.\ -proof (unfold Cauchy_altdef2, rule, rule) - fix e::real - assume \e > 0\ - have \\k. a k \ 0\ - using \\n. dist (\ (Suc n)) (\ n) \ a n\ dual_order.trans zero_le_dist by blast - hence \Cauchy (\k. sum a {0..k})\ - using \\M. \n. sum a {0..n} \ M\ sum_Cauchy_positive by blast - hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ - unfolding Cauchy_def using \e > 0\ by blast - hence \\M. \m\M. \n\M. m > n \ dist (sum a {0..m}) (sum a {0..n}) < e\ - by blast - have \dist (sum a {0..m}) (sum a {0..n}) = sum a {Suc n..m}\ if \n for m n - proof - - have \n < Suc n\ - by simp - have \finite {0..n}\ - by simp - moreover have \finite {Suc n..m}\ - by simp - moreover have \{0..n} \ {Suc n..m} = {0..m}\ - using \n < Suc n\ \n < m\ by auto - moreover have \{0..n} \ {Suc n..m} = {}\ - by simp - ultimately have sum_plus: \(sum a {0..n}) + sum a {Suc n..m} = (sum a {0..m})\ - by (metis sum.union_disjoint) - have \dist (sum a {0..m}) (sum a {0..n}) = \(sum a {0..m}) - (sum a {0..n})\\ - using dist_real_def by blast - moreover have \(sum a {0..m}) - (sum a {0..n}) = sum a {Suc n..m}\ - using sum_plus by linarith - ultimately show ?thesis - by (simp add: \\k. 0 \ a k\ sum_nonneg) - qed - hence sum_a: \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ - by (metis \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\) - obtain M where \\m\M. \n\M. m > n \ sum a {Suc n..m} < e\ - using sum_a \e > 0\ by blast - hence \\m. \n. Suc m \ Suc M \ Suc n \ Suc M \ Suc m > Suc n \ sum a {Suc n..Suc m - 1} < e\ - by simp - hence \\m\1. \n\1. m \ Suc M \ n \ Suc M \ m > n \ sum a {n..m - 1} < e\ - by (metis Suc_le_D) - hence sum_a2: \\M. \m\M. \n\M. m > n \ sum a {n..m-1} < e\ - by (meson add_leE) - have \dist (\ (n+p+1)) (\ n) \ sum a {n..n+p}\ for p n :: nat - proof(induction p) - case 0 thus ?case by (simp add: assms(2)) - next - case (Suc p) thus ?case - by (smt Suc_eq_plus1 add_Suc_right add_less_same_cancel1 assms(2) dist_self dist_triangle2 - gr_implies_not0 sum.cl_ivl_Suc) - qed - hence \m > n \ dist (\ m) (\ n) \ sum a {n..m-1}\ for m n :: nat - by (metis Suc_eq_plus1 Suc_le_D diff_Suc_1 gr0_implies_Suc less_eq_Suc_le less_imp_Suc_add - zero_less_Suc) - hence \\M. \m\M. \n\M. m > n \ dist (\ m) (\ n) < e\ - using sum_a2 \e > 0\ by smt - thus "\N. \n\N. dist (\ n) (\ N) < e" - using \0 < e\ by fastforce -qed - -unbundle notation_blinfun_apply - -unbundle no_notation_norm - -end +(* + File: Banach_Steinhaus_Missing.thy + Author: Dominique Unruh, University of Tartu + Author: Jose Manuel Rodriguez Caballero, University of Tartu +*) +section \Missing results for the proof of Banach-Steinhaus theorem\ + +theory Banach_Steinhaus_Missing + imports + "HOL-Analysis.Infinite_Set_Sum" + +begin +subsection \Results missing for the proof of Banach-Steinhaus theorem\ +text \ + The results proved here are preliminaries for the proof of Banach-Steinhaus theorem using Sokal's + approach, but they do not explicitly appear in Sokal's paper ~\cite{sokal2011reall}. +\ + +text\Notation for the norm\ +bundle notation_norm begin +notation norm ("\_\") +end + +bundle no_notation_norm begin +no_notation norm ("\_\") +end + +unbundle notation_norm + +text\Notation for apply bilinear function\ +bundle notation_blinfun_apply begin +notation blinfun_apply (infixr "*\<^sub>v" 70) +end + +bundle no_notation_blinfun_apply begin +no_notation blinfun_apply (infixr "*\<^sub>v" 70) +end + +unbundle notation_blinfun_apply + +lemma bdd_above_plus: + fixes f::\'a \ real\ + assumes \bdd_above (f ` S)\ and \bdd_above (g ` S)\ + shows \bdd_above ((\ x. f x + g x) ` S)\ + text \ + Explanation: If the images of two real-valued functions \<^term>\f\,\<^term>\g\ are bounded above on a + set \<^term>\S\, then the image of their sum is bounded on \<^term>\S\. +\ +proof- + obtain M where \\ x. x\S \ f x \ M\ + using \bdd_above (f ` S)\ unfolding bdd_above_def by blast + obtain N where \\ x. x\S \ g x \ N\ + using \bdd_above (g ` S)\ unfolding bdd_above_def by blast + have \\ x. x\S \ f x + g x \ M + N\ + using \\x. x \ S \ f x \ M\ \\x. x \ S \ g x \ N\ by fastforce + thus ?thesis unfolding bdd_above_def by blast +qed + +text\The maximum of two functions\ +definition pointwise_max:: "('a \ 'b::ord) \ ('a \ 'b) \ ('a \ 'b)" where + \pointwise_max f g = (\x. max (f x) (g x))\ + +lemma max_Sup_absorb_left: + fixes f g::\'a \ real\ + assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ + shows \Sup ((pointwise_max f g) ` X) = Sup (f ` X)\ + + text \Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\, if the supremum of \<^term>\f\ is + greater-equal the supremum of \<^term>\g\, then the supremum of \<^term>\max f g\ equals the supremum of + \<^term>\f\. (Under some technical conditions.)\ + +proof- + have y_Sup: \y \ ((\ x. max (f x) (g x)) ` X) \ y \ Sup (f ` X)\ for y + proof- + assume \y \ ((\ x. max (f x) (g x)) ` X)\ + then obtain x where \y = max (f x) (g x)\ and \x \ X\ + by blast + have \f x \ Sup (f ` X)\ + by (simp add: \x \ X\ \bdd_above (f ` X)\ cSUP_upper) + moreover have \g x \ Sup (g ` X)\ + by (simp add: \x \ X\ \bdd_above (g ` X)\ cSUP_upper) + ultimately have \max (f x) (g x) \ Sup (f ` X)\ + using \Sup (f ` X) \ Sup (g ` X)\ by auto + thus ?thesis by (simp add: \y = max (f x) (g x)\) + qed + have y_f_X: \y \ f ` X \ y \ Sup ((\ x. max (f x) (g x)) ` X)\ for y + proof- + assume \y \ f ` X\ + then obtain x where \x \ X\ and \y = f x\ + by blast + have \bdd_above ((\ \. max (f \) (g \)) ` X)\ + by (metis (no_types) \bdd_above (f ` X)\ \bdd_above (g ` X)\ bdd_above_image_sup sup_max) + moreover have \e > 0 \ \ k \ (\ \. max (f \) (g \)) ` X. y \ k + e\ + for e::real + using \Sup (f ` X) \ Sup (g ` X)\ by (smt \x \ X\ \y = f x\ image_eqI) + ultimately show ?thesis + using \x \ X\ \y = f x\ cSUP_upper by fastforce + qed + have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ + using y_Sup by (simp add: \X \ {}\ cSup_least) + moreover have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ + using y_f_X by (metis (mono_tags) cSup_least calculation empty_is_image) + ultimately show ?thesis unfolding pointwise_max_def by simp +qed + +lemma max_Sup_absorb_right: + fixes f g::\'a \ real\ + assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ + shows \Sup ((pointwise_max f g) ` X) = Sup (g ` X)\ + text \ + Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\ and a nonempty set \<^term>\X\, such that + the \<^term>\f\ and \<^term>\g\ are bounded above on \<^term>\X\, if the supremum of \<^term>\f\ on \<^term>\X\ is + lower-equal the supremum of \<^term>\g\ on \<^term>\X\, then the supremum of \<^term>\pointwise_max f g\ on \<^term>\X\ + equals the supremum of \<^term>\g\. This is the right analog of @{text max_Sup_absorb_left}. +\ +proof- + have \Sup ((pointwise_max g f) ` X) = Sup (g ` X)\ + using assms by (simp add: max_Sup_absorb_left) + moreover have \pointwise_max g f = pointwise_max f g\ + unfolding pointwise_max_def by auto + ultimately show ?thesis by simp +qed + +lemma max_Sup: + fixes f g::\'a \ real\ + assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ + shows \Sup ((pointwise_max f g) ` X) = max (Sup (f ` X)) (Sup (g ` X))\ + text \ + Explanation: Let \<^term>\X\ be a nonempty set. Two supremum over \<^term>\X\ of the maximum of two + real-value functions is equal to the maximum of their suprema over \<^term>\X\, provided that the + functions are bounded above on \<^term>\X\. +\ +proof(cases \Sup (f ` X) \ Sup (g ` X)\) + case True thus ?thesis by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_left) +next + case False + have f1: "\ 0 \ Sup (f ` X) + - 1 * Sup (g ` X)" + using False by linarith + hence "Sup (Banach_Steinhaus_Missing.pointwise_max f g ` X) = Sup (g ` X)" + by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_right) + thus ?thesis + using f1 by linarith +qed + + +lemma identity_telescopic: + fixes x :: \_ \ 'a::real_normed_vector\ + assumes \x \ l\ + shows \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) \ l - x n\ + text\ + Expression of a limit as a telescopic series. + Explanation: If \<^term>\x\ converges to \<^term>\l\ then the sum \<^term>\sum (\ k. x (Suc k) - x k) {n..N}\ + converges to \<^term>\l - x n\ as \<^term>\N\ goes to infinity. +\ +proof- + have \(\ p. x (p + Suc n)) \ l\ + using \x \ l\ by (rule LIMSEQ_ignore_initial_segment) + hence \(\ p. x (Suc n + p)) \ l\ + by (simp add: add.commute) + hence \(\ p. x (Suc (n + p))) \ l\ + by simp + hence \(\ t. (- (x n)) + (\ p. x (Suc (n + p))) t ) \ (- (x n)) + l\ + using tendsto_add_const_iff by metis + hence f1: \(\ p. x (Suc (n + p)) - x n)\ l - x n\ + by simp + have \sum (\ k. x (Suc k) - x k) {n..n+p} = x (Suc (n+p)) - x n\ for p + by (simp add: sum_Suc_diff) + moreover have \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + t) + = (\ p. sum (\ k. x (Suc k) - x k) {n..n+p}) t\ for t + by blast + ultimately have \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + p)) \ l - x n\ + using f1 by simp + hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (p + n)) \ l - x n\ + by (simp add: add.commute) + hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) p) \ l - x n\ + using Topological_Spaces.LIMSEQ_offset[where f = "(\ N. sum (\ k. x (Suc k) - x k) {n..N})" + and a = "l - x n" and k = n] by blast + hence \(\ M. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) M) \ l - x n\ + by simp + thus ?thesis by blast +qed + +lemma bound_Cauchy_to_lim: + assumes \y \ x\ and \\n. \y (Suc n) - y n\ \ c^n\ and \y 0 = 0\ and \c < 1\ + shows \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ + text\ + Inequality about a sequence of approximations assuming that the sequence of differences is bounded + by a geometric progression. + Explanation: Let \<^term>\y\ be a sequence converging to \<^term>\x\. + If \<^term>\y\ satisfies the inequality \\y (Suc n) - y n\ \ c ^ n\ for some \<^term>\c < 1\ and + assuming \<^term>\y 0 = 0\ then the inequality \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ holds. +\ +proof- + have \c \ 0\ + using \\ n. \y (Suc n) - y n\ \ c^n\ by (smt norm_imp_pos_and_ge power_Suc0_right) + have norm_1: \norm (\k = Suc n..N. y (Suc k) - y k) \ (c ^ Suc n)/(1 - c)\ for N + proof(cases \N < Suc n\) + case True + hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ = 0\ + by auto + thus ?thesis using \c \ 0\ \c < 1\ by auto + next + case False + hence \N \ Suc n\ + by auto + have \c^(Suc N) \ 0\ + using \c \ 0\ by auto + have \1 - c > 0\ + by (simp add: \c < 1\) + hence \(1 - c)/(1 - c) = 1\ + by auto + have \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (\k. \y (Suc k) - y k\) {Suc n .. N})\ + by (simp add: sum_norm_le) + hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (power c) {Suc n .. N})\ + by (simp add: assms(2) sum_norm_le) + hence \(1 - c) * \sum (\k. y (Suc k) - y k) {Suc n .. N}\ + \ (1 - c) * (sum (power c) {Suc n .. N})\ + using \0 < 1 - c\ real_mult_le_cancel_iff2 by blast + also have \\ = c^(Suc n) - c^(Suc N)\ + using Set_Interval.sum_gp_multiplied \Suc n \ N\ by blast + also have \\ \ c^(Suc n)\ + using \c^(Suc N) \ 0\ by auto + finally have \(1 - c) * \\k = Suc n..N. y (Suc k) - y k\ \ c ^ Suc n\ + by blast + hence \((1 - c) * \\k = Suc n..N. y (Suc k) - y k\)/(1 - c) + \ (c ^ Suc n)/(1 - c)\ + using \0 < 1 - c\ by (smt divide_right_mono) + thus \\\k = Suc n..N. y (Suc k) - y k\ \ (c ^ Suc n)/(1 - c)\ + using \0 < 1 - c\ by auto + qed + have \(\ N. (sum (\k. y (Suc k) - y k) {Suc n .. N})) \ x - y (Suc n)\ + by (metis (no_types) \y \ x\ identity_telescopic) + hence \(\ N. \sum (\k. y (Suc k) - y k) {Suc n .. N}\) \ \x - y (Suc n)\\ + using tendsto_norm by blast + hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ + using norm_1 Lim_bounded by blast + hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ + by auto + moreover have \(c ^ Suc n)/(1 - c) = (c / (1 - c)) * (c ^ n)\ + by (simp add: divide_inverse_commute) + ultimately show \\x - y (Suc n)\ \ (c / (1 - c)) * (c ^ n)\ by linarith +qed + +lemma onorm_open_ball: + includes notation_norm + shows \\f\ = Sup { \f *\<^sub>v x\ | x. \x\ < 1 }\ + text \ + Explanation: Let \<^term>\f\ be a bounded linear operator. The operator norm of \<^term>\f\ is the + supremum of \<^term>\norm (f x)\ for \<^term>\x\ such that \<^term>\norm x < 1\. +\ +proof(cases \(UNIV::'a set) = 0\) + case True + hence \x = 0\ for x::'a + by auto + hence \f *\<^sub>v x = 0\ for x + by (metis (full_types) blinfun.zero_right) + hence \\f\ = 0\ + by (simp add: blinfun_eqI zero_blinfun.rep_eq) + have \{ \f *\<^sub>v x\ | x. \x\ < 1} = {0}\ + by (smt Collect_cong \\x. f *\<^sub>v x = 0\ norm_zero singleton_conv) + hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} = 0\ + by simp + thus ?thesis using \\f\ = 0\ by auto +next + case False + hence \(UNIV::'a set) \ 0\ + by simp + have nonnegative: \\f *\<^sub>v x\ \ 0\ for x + by simp + have \\ x::'a. x \ 0\ + using \UNIV \ 0\ by auto + then obtain x::'a where \x \ 0\ + by blast + hence \\x\ \ 0\ + by auto + define y where \y = x /\<^sub>R \x\\ + have \norm y = \ x /\<^sub>R \x\ \\ + unfolding y_def by auto + also have \\ = \x\ /\<^sub>R \x\\ + by auto + also have \\ = 1\ + using \\x\ \ 0\ by auto + finally have \\y\ = 1\ + by blast + hence norm_1_non_empty: \{ \f *\<^sub>v x\ | x. \x\ = 1} \ {}\ + by blast + have norm_1_bounded: \bdd_above { \f *\<^sub>v x\ | x. \x\ = 1}\ + unfolding bdd_above_def apply auto + by (metis norm_blinfun) + have norm_less_1_non_empty: \{\f *\<^sub>v x\ | x. \x\ < 1} \ {}\ + by (metis (mono_tags, lifting) Collect_empty_eq_bot bot_empty_eq empty_iff norm_zero + zero_less_one) + have norm_less_1_bounded: \bdd_above {\f *\<^sub>v x\ | x. \x\ < 1}\ + proof- + have \\r. \a r\ < 1 \ \f *\<^sub>v (a r)\ \ r\ for a :: "real \ 'a" + proof- + obtain r :: "('a \\<^sub>L 'b) \ real" where + "\f x. 0 \ r f \ (bounded_linear f \ \f *\<^sub>v x\ \ \x\ * r f)" + using bounded_linear.nonneg_bounded by moura + have \\ \f\ < 0\ + by simp + hence "(\r. \f\ * \a r\ \ r) \ (\r. \a r\ < 1 \ \f *\<^sub>v a r\ \ r)" + by (meson less_eq_real_def mult_le_cancel_left2) + thus ?thesis using dual_order.trans norm_blinfun by blast + qed + hence \\ M. \ x. \x\ < 1 \ \f *\<^sub>v x\ \ M\ + by metis + thus ?thesis by auto + qed + have Sup_non_neg: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} \ 0\ + by (smt Collect_empty_eq cSup_upper mem_Collect_eq nonnegative norm_1_bounded norm_1_non_empty) + have \{0::real} \ {}\ + by simp + have \bdd_above {0::real}\ + by simp + show \\f\ = Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ + proof(cases \\x. f *\<^sub>v x = 0\) + case True + have \\f *\<^sub>v x\ = 0\ for x + by (simp add: True) + hence \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ + by blast + moreover have \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ + using calculation norm_less_1_non_empty by fastforce + ultimately have \{\f *\<^sub>v x\ | x. \x\ < 1 } = {0}\ + by blast + hence Sup1: \Sup {\f *\<^sub>v x\ | x. \x\ < 1 } = 0\ + by simp + have \\f\ = 0\ + by (simp add: True blinfun_eqI) + moreover have \Sup {\f *\<^sub>v x\ | x. \x\ < 1} = 0\ + using Sup1 by blast + ultimately show ?thesis by simp + next + case False + have norm_f_eq_leq: \y \ {\f *\<^sub>v x\ | x. \x\ = 1} \ + y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for y + proof- + assume \y \ {\f *\<^sub>v x\ | x. \x\ = 1}\ + hence \\ x. y = \f *\<^sub>v x\ \ \x\ = 1\ + by blast + then obtain x where \y = \f *\<^sub>v x\\ and \\x\ = 1\ + by auto + define y' where \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R y\ for n + have \y' n \ {\f *\<^sub>v x\ | x. \x\ < 1}\ for n + proof- + have \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R \f *\<^sub>v x\\ + using y'_def \y = \f *\<^sub>v x\\ by blast + also have \... = \(1 - (inverse (real (Suc n))))\ *\<^sub>R \f *\<^sub>v x\\ + by (metis (mono_tags, hide_lams) \y = \f *\<^sub>v x\\ abs_1 abs_le_self_iff abs_of_nat + abs_of_nonneg add_diff_cancel_left' add_eq_if cancel_comm_monoid_add_class.diff_cancel + diff_ge_0_iff_ge eq_iff_diff_eq_0 inverse_1 inverse_le_iff_le nat.distinct(1) of_nat_0 + of_nat_Suc of_nat_le_0_iff zero_less_abs_iff zero_neq_one) + also have \... = \f *\<^sub>v ((1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ + by (simp add: blinfun.scaleR_right) + finally have y'_1: \y' n = \f *\<^sub>v ( (1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ + by blast + have \\(1 - (inverse (Suc n))) *\<^sub>R x\ = (1 - (inverse (real (Suc n)))) * \x\\ + by (simp add: linordered_field_class.inverse_le_1_iff) + hence \\(1 - (inverse (Suc n))) *\<^sub>R x\ < 1\ + by (simp add: \\x\ = 1\) + thus ?thesis using y'_1 by blast + qed + have \(\n. (1 - (inverse (real (Suc n)))) ) \ 1\ + using Limits.LIMSEQ_inverse_real_of_nat_add_minus by simp + hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ 1 *\<^sub>R y\ + using Limits.tendsto_scaleR by blast + hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ y\ + by simp + hence \(\n. y' n) \ y\ + using y'_def by simp + hence \y' \ y\ + by simp + have \y' n \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for n + using cSup_upper \\n. y' n \ {\f *\<^sub>v x\ |x. \x\ < 1}\ norm_less_1_bounded by blast + hence \y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ + using \y' \ y\ Topological_Spaces.Sup_lim by (meson LIMSEQ_le_const2) + thus ?thesis by blast + qed + hence \Sup {\f *\<^sub>v x\ | x. \x\ = 1} \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ + by (metis (lifting) cSup_least norm_1_non_empty) + have \y \ {\f *\<^sub>v x\ | x. \x\ < 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ for y + proof(cases \y = 0\) + case True thus ?thesis by (simp add: Sup_non_neg) + next + case False + hence \y \ 0\ by blast + assume \y \ {\f *\<^sub>v x\ | x. \x\ < 1}\ + hence \\ x. y = \f *\<^sub>v x\ \ \x\ < 1\ + by blast + then obtain x where \y = \f *\<^sub>v x\\ and \\x\ < 1\ + by blast + have \(1/\x\) * y = (1/\x\) * \f x\\ + by (simp add: \y = \f *\<^sub>v x\\) + also have \... = \1/\x\\ * \f *\<^sub>v x\\ + by simp + also have \... = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ + by simp + also have \... = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ + by (simp add: blinfun.scaleR_right) + finally have \(1/\x\) * y = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ + by blast + have \x \ 0\ + using \y \ 0\ \y = \f *\<^sub>v x\\ blinfun.zero_right by auto + have \\ (1/\x\) *\<^sub>R x \ = \ (1/\x\) \ * \x\\ + by simp + also have \... = (1/\x\) * \x\\ + by simp + finally have \\(1/\x\) *\<^sub>R x\ = 1\ + using \x \ 0\ by simp + hence \(1/\x\) * y \ { \f *\<^sub>v x\ | x. \x\ = 1}\ + using \1 / \x\ * y = \f *\<^sub>v (1 / \x\) *\<^sub>R x\\ by blast + hence \(1/\x\) * y \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ + by (simp add: cSup_upper norm_1_bounded) + moreover have \y \ (1/\x\) * y\ + by (metis \\x\ < 1\ \y = \f *\<^sub>v x\\ mult_le_cancel_right1 norm_not_less_zero + order.strict_implies_order \x \ 0\ less_divide_eq_1_pos zero_less_norm_iff) + ultimately show ?thesis by linarith + qed + hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ + by (smt cSup_least norm_less_1_non_empty) + hence \Sup { \f *\<^sub>v x\ | x. \x\ = 1} = Sup { \f *\<^sub>v x\ | x. \x\ < 1}\ + using \Sup {\f *\<^sub>v x\ |x. norm x = 1} \ Sup { \f *\<^sub>v x\ |x. \x\ < 1}\ by linarith + have f1: \(SUP x. \f *\<^sub>v x\ / \x\) = Sup { \f *\<^sub>v x\ / \x\ | x. True}\ + by (simp add: full_SetCompr_eq) + have \y \ { \f *\<^sub>v x\ / \x\ |x. True} \ y \ { \f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ + for y + proof- + assume \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ show ?thesis + proof(cases \y = 0\) + case True thus ?thesis by simp + next + case False + have \\ x. y = \f *\<^sub>v x\ / \x\\ + using \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ by auto + then obtain x where \y = \f *\<^sub>v x\ / \x\\ + by blast + hence \y = \(1/\x\)\ * \ f *\<^sub>v x \\ + by simp + hence \y = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ + by simp + hence \y = \f ((1/\x\) *\<^sub>R x)\\ + by (simp add: blinfun.scaleR_right) + moreover have \\ (1/\x\) *\<^sub>R x \ = 1\ + using False \y = \f *\<^sub>v x\ / \x\\ by auto + ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ + by blast + thus ?thesis by blast + qed + qed + moreover have \y \ {\f x\ |x. \x\ = 1} \ {0} \ y \ {\f *\<^sub>v x\ / \x\ |x. True}\ + for y + proof(cases \y = 0\) + case True thus ?thesis by auto + next + case False + hence \y \ {0}\ + by simp + moreover assume \y \ {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ + ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ + by simp + then obtain x where \\x\ = 1\ and \y = \f *\<^sub>v x\\ + by auto + have \y = \f *\<^sub>v x\ / \x\\ using \\x\ = 1\ \y = \f *\<^sub>v x\\ + by simp + thus ?thesis by auto + qed + ultimately have \{\f *\<^sub>v x\ / \x\ |x. True} = {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ + by blast + hence \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ + by simp + have "\r s. \ (r::real) \ s \ sup r s = s" + by (metis (lifting) sup.absorb_iff1 sup_commute) + hence \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {(0::real)}) + = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0::real})\ + using \0 \ Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ \bdd_above {0}\ \{0} \ {}\ cSup_singleton + cSup_union_distrib max.absorb_iff1 sup_commute norm_1_bounded norm_1_non_empty + by (metis (no_types, lifting) ) + moreover have \Sup {(0::real)} = (0::real)\ + by simp + ultimately have \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ + using Sup_non_neg by linarith + moreover have \Sup ( {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) + = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0}) \ + using Sup_non_neg \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) + = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0})\ + by auto + ultimately have f2: \Sup {\f *\<^sub>v x\ / \x\ | x. True} = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ + using \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by linarith + have \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ + using f1 f2 by linarith + hence \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ < 1 }\ + by (simp add: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\) + thus ?thesis apply transfer by (simp add: onorm_def) + qed +qed + +lemma onorm_r: + includes notation_norm + assumes \r > 0\ + shows \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ + text \ + Explanation: The norm of \<^term>\f\ is \<^term>\1/r\ of the supremum of the norm of \<^term>\f *\<^sub>v x\ for + \<^term>\x\ in the ball of radius \<^term>\r\ centered at the origin. +\ +proof- + have \\f\ = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\ + using onorm_open_ball by blast + moreover have \{\f *\<^sub>v x\ |x. \x\ < 1} = (\x. \f *\<^sub>v x\) ` (ball 0 1)\ + unfolding ball_def by auto + ultimately have onorm_f: \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 1))\ + by simp + have s2: \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1 \ x \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1)\ for x + proof- + assume \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ + hence \\ t. x = r *\<^sub>R \f *\<^sub>v t\ \ \t\ < 1\ + by auto + then obtain t where \x = r *\<^sub>R \f *\<^sub>v t\\ and \\t\ < 1\ + by blast + define y where \y = x /\<^sub>R r\ + have \x = r * (inverse r * x)\ + using \x = r *\<^sub>R norm (f t)\ by auto + hence \x - (r * (inverse r * x)) \ 0\ + by linarith + hence \x \ r * (x /\<^sub>R r)\ + by auto + have \y \ (\k. \f *\<^sub>v k\) ` ball 0 1\ + unfolding y_def by (smt \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ assms image_iff + inverse_inverse_eq pos_le_divideR_eq positive_imp_inverse_positive) + moreover have \x \ r * y\ + using \x \ r * (x /\<^sub>R r)\ y_def by blast + ultimately have y_norm_f: \y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ + by blast + have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ + by simp + moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ + by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above + bounded_norm_comp) + moreover have \\ y. y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ + using y_norm_f by blast + ultimately show ?thesis + by (smt \0 < r\ cSup_upper ordered_comm_semiring_class.comm_mult_left_mono) + qed + have s3: \(\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y) \ + r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y\ for y + proof- + assume \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ + have x_leq: \x \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ y / r\ for x + proof- + assume \x \ (\t. \f *\<^sub>v t\) ` ball 0 1\ + then obtain t where \t \ ball (0::'a) 1\ and \x = \f *\<^sub>v t\\ + by auto + define x' where \x' = r *\<^sub>R x\ + have \x' = r * \f *\<^sub>v t\\ + by (simp add: \x = \f *\<^sub>v t\\ x'_def) + hence \x' \ (\t. r * \f *\<^sub>v t\) ` ball 0 1\ + using \t \ ball (0::'a) 1\ by auto + hence \x' \ y\ + using \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ by blast + thus \x \ y / r\ + unfolding x'_def using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) + qed + have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ + by simp + moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ + by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above + bounded_norm_comp) + ultimately have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y/r\ + using x_leq by (simp add: \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ cSup_least) + thus ?thesis using \r > 0\ + by (smt divide_strict_right_mono nonzero_mult_div_cancel_left) + qed + have norm_scaleR: \norm \ ((*\<^sub>R) r) = ((*\<^sub>R) \r\) \ (norm::'a \ real)\ + by auto + have f_x1: \f (r *\<^sub>R x) = r *\<^sub>R f x\ for x + by (simp add: blinfun.scaleR_right) + have \ball (0::'a) r = ((*\<^sub>R) r) ` (ball 0 1)\ + by (smt assms ball_scale nonzero_mult_div_cancel_left right_inverse_eq scale_zero_right) + hence \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) = Sup ((\t. \f *\<^sub>v t\) ` (((*\<^sub>R) r) ` (ball 0 1)))\ + by simp + also have \\ = Sup (((\t. \f *\<^sub>v t\) \ ((*\<^sub>R) r)) ` (ball 0 1))\ + using Sup.SUP_image by auto + also have \\ = Sup ((\t. \f *\<^sub>v (r *\<^sub>R t)\) ` (ball 0 1))\ + using f_x1 by (simp add: comp_assoc) + also have \\ = Sup ((\t. \r\ *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ + using norm_scaleR f_x1 by auto + also have \\ = Sup ((\t. r *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ + using \r > 0\ by auto + also have \\ = r * Sup ((\t. \f *\<^sub>v t\) ` (ball 0 1))\ + apply (rule cSup_eq_non_empty) apply simp using s2 apply auto using s3 by auto + also have \\ = r * \f\\ + using onorm_f by auto + finally have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 r) = r * \f\\ + by blast + thus \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ using \r > 0\ by simp +qed + +text\Pointwise convergence\ +definition pointwise_convergent_to :: + \( nat \ ('a \ 'b::topological_space) ) \ ('a \ 'b) \ bool\ + (\((_)/ \pointwise\ (_))\ [60, 60] 60) where + \pointwise_convergent_to x l = (\ t::'a. (\ n. (x n) t) \ l t)\ + +lemma linear_limit_linear: + fixes f :: \_ \ ('a::real_vector \ 'b::real_normed_vector)\ + assumes \\n. linear (f n)\ and \f \pointwise\ F\ + shows \linear F\ + text\ + Explanation: If a family of linear operators converges pointwise, then the limit is also a linear + operator. +\ +proof + show "F (x + y) = F x + F y" for x y + proof- + have "\a. F a = lim (\n. f n a)" + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (full_types) limI) + moreover have "\f b c g. (lim (\n. g n + f n) = (b::'b) + c \ \ f \ c) \ \ g \ b" + by (metis (no_types) limI tendsto_add) + moreover have "\a. (\n. f n a) \ F a" + using assms(2) pointwise_convergent_to_def by force + ultimately have + lim_sum: \lim (\ n. (f n) x + (f n) y) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ + by metis + have \(f n) (x + y) = (f n) x + (f n) y\ for n + using \\ n. linear (f n)\ unfolding linear_def using Real_Vector_Spaces.linear_iff assms(1) + by auto + hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x + (f n) y)\ + by simp + hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ + using lim_sum by simp + moreover have \(\ n. (f n) (x + y)) \ F (x + y)\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast + moreover have \(\ n. (f n) x) \ F x\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast + moreover have \(\ n. (f n) y) \ F y\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast + ultimately show ?thesis + by (metis limI) + qed + show "F (r *\<^sub>R x) = r *\<^sub>R F x" for r and x + proof- + have \(f n) (r *\<^sub>R x) = r *\<^sub>R (f n) x\ for n + using \\ n. linear (f n)\ + by (simp add: Real_Vector_Spaces.linear_def real_vector.linear_scale) + hence \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ + by simp + have \convergent (\ n. (f n) x)\ + by (metis assms(2) convergentI pointwise_convergent_to_def) + moreover have \isCont (\ t::'b. r *\<^sub>R t) tt\ for tt + by (simp add: bounded_linear_scaleR_right) + ultimately have \lim (\ n. r *\<^sub>R ((f n) x)) = r *\<^sub>R lim (\ n. (f n) x)\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def + by (metis (mono_tags) isCont_tendsto_compose limI) + hence \lim (\ n. (f n) (r *\<^sub>R x)) = r *\<^sub>R lim (\ n. (f n) x)\ + using \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp + moreover have \(\ n. (f n) x) \ F x\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast + moreover have \(\ n. (f n) (r *\<^sub>R x)) \ F (r *\<^sub>R x)\ + using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast + ultimately show ?thesis + by (metis limI) + qed +qed + + +lemma non_Cauchy_unbounded: + fixes a ::\_ \ real\ + assumes \\n. a n \ 0\ and \e > 0\ + and \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ + shows \(\n. (sum a {0..n})) \ \\ + text\ + Explanation: If the sequence of partial sums of nonnegative terms is not Cauchy, then it converges + to infinite. +\ +proof- + define S::"ereal set" where \S = range (\n. sum a {0..n})\ + have \\s\S. k*e \ s\ for k::nat + proof(induction k) + case 0 + from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ + obtain m n where \m \ 0\ and \n \ 0\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast + have \n < Suc n\ + by simp + hence \{0..n} \ {Suc n..m} = {0..m}\ + using Set_Interval.ivl_disj_un(7) \n < m\ by auto + moreover have \finite {0..n}\ + by simp + moreover have \finite {Suc n..m}\ + by simp + moreover have \{0..n} \ {Suc n..m} = {}\ + by simp + ultimately have \sum a {0..n} + sum a {Suc n..m} = sum a {0..m}\ + by (metis sum.union_disjoint) + moreover have \sum a {Suc n..m} > 0\ + using \e > 0\ \sum a {Suc n..m} \ e\ by linarith + moreover have \sum a {0..n} \ 0\ + by (simp add: assms(1) sum_nonneg) + ultimately have \sum a {0..m} > 0\ + by linarith + moreover have \sum a {0..m} \ S\ + unfolding S_def by blast + ultimately have \\s\S. 0 \ s\ + using ereal_less_eq(5) by fastforce + thus ?case + by (simp add: zero_ereal_def) + next + case (Suc k) + assume \\s\S. k*e \ s\ + then obtain s where \s\S\ and \ereal (k * e) \ s\ + by blast + have \\N. s = sum a {0..N}\ + using \s\S\ unfolding S_def by blast + then obtain N where \s = sum a {0..N}\ + by blast + from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ + obtain m n where \m \ Suc N\ and \n \ Suc N\ and \m > n\ and \sum a {Suc n..m} \ e\ + by blast + have \finite {Suc N..n}\ + by simp + moreover have \finite {Suc n..m}\ + by simp + moreover have \{Suc N..n} \ {Suc n..m} = {Suc N..m}\ + using Set_Interval.ivl_disj_un + by (smt \Suc N \ n\ \n < m\ atLeastSucAtMost_greaterThanAtMost less_imp_le_nat) + moreover have \{} = {Suc N..n} \ {Suc n..m}\ + by simp + ultimately have \sum a {Suc N..m} = sum a {Suc N..n} + sum a {Suc n..m}\ + by (metis sum.union_disjoint) + moreover have \sum a {Suc N..n} \ 0\ + using \\n. a n \ 0\ by (simp add: sum_nonneg) + ultimately have \sum a {Suc N..m} \ e\ + using \e \ sum a {Suc n..m}\ by linarith + have \finite {0..N}\ + by simp + have \finite {Suc N..m}\ + by simp + moreover have \{0..N} \ {Suc N..m} = {0..m}\ + using Set_Interval.ivl_disj_un(7) \Suc N \ m\ by auto + moreover have \{0..N} \ {Suc N..m} = {}\ + by simp + ultimately have \sum a {0..N} + sum a {Suc N..m} = sum a {0..m}\ + by (metis \finite {0..N}\ sum.union_disjoint) + hence \e + k * e \ sum a {0..m}\ + using \ereal (real k * e) \ s\ \s = ereal (sum a {0..N})\ \e \ sum a {Suc N..m}\ by auto + moreover have \e + k * e = (Suc k) * e\ + by (simp add: semiring_normalization_rules(3)) + ultimately have \(Suc k) * e \ sum a {0..m}\ + by linarith + hence \ereal ((Suc k) * e) \ sum a {0..m}\ + by auto + moreover have \sum a {0..m}\S\ + unfolding S_def by blast + ultimately show ?case by blast + qed + hence \\s\S. (real n) \ s\ for n + by (meson assms(2) ereal_le_le ex_less_of_nat_mult less_le_not_le) + hence \Sup S = \\ + using Sup_le_iff Sup_subset_mono dual_order.strict_trans1 leD less_PInf_Ex_of_nat subsetI + by metis + hence Sup: \Sup ((range (\ n. (sum a {0..n})))::ereal set) = \\ using S_def + by blast + have \incseq (\n. (sum a {.. + using \\n. a n \ 0\ using Extended_Real.incseq_sumI by auto + hence \incseq (\n. (sum a {..< Suc n}))\ + by (meson incseq_Suc_iff) + hence \incseq (\n. (sum a {0..n})::ereal)\ + using incseq_ereal by (simp add: atLeast0AtMost lessThan_Suc_atMost) + hence \(\n. sum a {0..n}) \ Sup (range (\n. (sum a {0..n})::ereal))\ + using LIMSEQ_SUP by auto + thus ?thesis using Sup PInfty_neq_ereal by auto +qed + +lemma sum_Cauchy_positive: + fixes a ::\_ \ real\ + assumes \\n. a n \ 0\ and \\K. \n. (sum a {0..n}) \ K\ + shows \Cauchy (\n. sum a {0..n})\ + text\ + Explanation: If a series of nonnegative reals is bounded, then the series is + Cauchy. +\ +proof (unfold Cauchy_altdef2, rule, rule) + fix e::real + assume \e>0\ + have \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ + proof(rule classical) + assume \\(\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e)\ + hence \\M. \m. \n. m \ M \ n \ M \ m > n \ \(sum a {Suc n..m} < e)\ + by blast + hence \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ + by fastforce + hence \(\n. (sum a {0..n}) ) \ \\ + using non_Cauchy_unbounded \0 < e\ assms(1) by blast + from \\K. \n. sum a {0..n} \ K\ + obtain K where \\n. sum a {0..n} \ K\ + by blast + from \(\n. sum a {0..n}) \ \\ + have \\B. \N. \n\N. (\ n. (sum a {0..n}) ) n \ B\ + using Lim_PInfty by simp + hence \\n. (sum a {0..n}) \ K+1\ + using ereal_less_eq(3) by blast + thus ?thesis using \\n. (sum a {0..n}) \ K\ by smt + qed + have \sum a {Suc n..m} = sum a {0..m} - sum a {0..n}\ + if "m > n" for m n + apply (simp add: that atLeast0AtMost) using sum_up_index_split + by (smt less_imp_add_positive that) + hence \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ + using \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by smt + from \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ + obtain M where \\m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ + by blast + moreover have \m > n \ sum a {0..m} \ sum a {0..n}\ for m n + using \\ n. a n \ 0\ by (simp add: sum_mono2) + ultimately have \\M. \m\M. \n\M. m > n \ \sum a {0..m} - sum a {0..n}\ < e\ + by auto + hence \\M. \m\M. \n\M. m \ n \ \sum a {0..m} - sum a {0..n}\ < e\ + by (metis \0 < e\ abs_zero cancel_comm_monoid_add_class.diff_cancel diff_is_0_eq' + less_irrefl_nat linorder_neqE_nat zero_less_diff) + hence \\M. \m\M. \n\M. \sum a {0..m} - sum a {0..n}\ < e\ + by (metis abs_minus_commute nat_le_linear) + hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ + by (simp add: dist_real_def) + hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by blast + thus \\N. \n\N. dist (sum a {0..n}) (sum a {0..N}) < e\ by auto +qed + +lemma convergent_series_Cauchy: + fixes a::\nat \ real\ and \::\nat \ 'a::metric_space\ + assumes \\M. \n. sum a {0..n} \ M\ and \\n. dist (\ (Suc n)) (\ n) \ a n\ + shows \Cauchy \\ + text\ + Explanation: Let \<^term>\a\ be a real-valued sequence and let \<^term>\\\ be sequence in a metric space. + If the partial sums of \<^term>\a\ are uniformly bounded and the distance between consecutive terms of \<^term>\\\ + are bounded by the sequence \<^term>\a\, then \<^term>\\\ is Cauchy.\ +proof (unfold Cauchy_altdef2, rule, rule) + fix e::real + assume \e > 0\ + have \\k. a k \ 0\ + using \\n. dist (\ (Suc n)) (\ n) \ a n\ dual_order.trans zero_le_dist by blast + hence \Cauchy (\k. sum a {0..k})\ + using \\M. \n. sum a {0..n} \ M\ sum_Cauchy_positive by blast + hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ + unfolding Cauchy_def using \e > 0\ by blast + hence \\M. \m\M. \n\M. m > n \ dist (sum a {0..m}) (sum a {0..n}) < e\ + by blast + have \dist (sum a {0..m}) (sum a {0..n}) = sum a {Suc n..m}\ if \n for m n + proof - + have \n < Suc n\ + by simp + have \finite {0..n}\ + by simp + moreover have \finite {Suc n..m}\ + by simp + moreover have \{0..n} \ {Suc n..m} = {0..m}\ + using \n < Suc n\ \n < m\ by auto + moreover have \{0..n} \ {Suc n..m} = {}\ + by simp + ultimately have sum_plus: \(sum a {0..n}) + sum a {Suc n..m} = (sum a {0..m})\ + by (metis sum.union_disjoint) + have \dist (sum a {0..m}) (sum a {0..n}) = \(sum a {0..m}) - (sum a {0..n})\\ + using dist_real_def by blast + moreover have \(sum a {0..m}) - (sum a {0..n}) = sum a {Suc n..m}\ + using sum_plus by linarith + ultimately show ?thesis + by (simp add: \\k. 0 \ a k\ sum_nonneg) + qed + hence sum_a: \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ + by (metis \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\) + obtain M where \\m\M. \n\M. m > n \ sum a {Suc n..m} < e\ + using sum_a \e > 0\ by blast + hence \\m. \n. Suc m \ Suc M \ Suc n \ Suc M \ Suc m > Suc n \ sum a {Suc n..Suc m - 1} < e\ + by simp + hence \\m\1. \n\1. m \ Suc M \ n \ Suc M \ m > n \ sum a {n..m - 1} < e\ + by (metis Suc_le_D) + hence sum_a2: \\M. \m\M. \n\M. m > n \ sum a {n..m-1} < e\ + by (meson add_leE) + have \dist (\ (n+p+1)) (\ n) \ sum a {n..n+p}\ for p n :: nat + proof(induction p) + case 0 thus ?case by (simp add: assms(2)) + next + case (Suc p) thus ?case + by (smt Suc_eq_plus1 add_Suc_right add_less_same_cancel1 assms(2) dist_self dist_triangle2 + gr_implies_not0 sum.cl_ivl_Suc) + qed + hence \m > n \ dist (\ m) (\ n) \ sum a {n..m-1}\ for m n :: nat + by (metis Suc_eq_plus1 Suc_le_D diff_Suc_1 gr0_implies_Suc less_eq_Suc_le less_imp_Suc_add + zero_less_Suc) + hence \\M. \m\M. \n\M. m > n \ dist (\ m) (\ n) < e\ + using sum_a2 \e > 0\ by smt + thus "\N. \n\N. dist (\ n) (\ N) < e" + using \0 < e\ by fastforce +qed + +unbundle notation_blinfun_apply + +unbundle no_notation_norm + +end diff --git a/thys/Banach_Steinhaus/ROOT b/thys/Banach_Steinhaus/ROOT --- a/thys/Banach_Steinhaus/ROOT +++ b/thys/Banach_Steinhaus/ROOT @@ -1,13 +1,13 @@ -chapter AFP - -session Banach_Steinhaus (AFP) = "HOL-Analysis" + - options [timeout = 300] - sessions - "HOL-Types_To_Sets" - "HOL-ex" - theories - Banach_Steinhaus - Banach_Steinhaus_Missing - document_files - "root.tex" - "root.bib" +chapter AFP + +session Banach_Steinhaus (AFP) = "HOL-Analysis" + + options [timeout = 300] + sessions + "HOL-Types_To_Sets" + "HOL-ex" + theories + Banach_Steinhaus + Banach_Steinhaus_Missing + document_files + "root.tex" + "root.bib" diff --git a/thys/Banach_Steinhaus/document/root.bib b/thys/Banach_Steinhaus/document/root.bib --- a/thys/Banach_Steinhaus/document/root.bib +++ b/thys/Banach_Steinhaus/document/root.bib @@ -1,26 +1,26 @@ -@article{banach1927principe, - title={Sur le principe de la condensation de singularit{\'e}s}, - author={Banach, Stefan and Steinhaus, Hugo}, - journal={Fundamenta Mathematicae}, - volume={1}, - number={9}, - pages={50--61}, - year={1927} -} - -@article{sokal2011really, - title={A really simple elementary proof of the uniform boundedness theorem}, - author={Sokal, Alan D}, - journal={The American Mathematical Monthly}, - volume={118}, - number={5}, - pages={450--452}, - year={2011}, - publisher={Taylor \& Francis} -} - -@article{Weisstein_UBP, - title={Uniform Boundedness Principle}, - author={Moslehian, Mohammad Sal and Weisstein, Eric W.}, - journal={From MathWorld--A Wolfram Web Resource. http://mathworld.wolfram.com/UniformBoundednessPrinciple.html} -} +@article{banach1927principe, + title={Sur le principe de la condensation de singularit{\'e}s}, + author={Banach, Stefan and Steinhaus, Hugo}, + journal={Fundamenta Mathematicae}, + volume={1}, + number={9}, + pages={50--61}, + year={1927} +} + +@article{sokal2011really, + title={A really simple elementary proof of the uniform boundedness theorem}, + author={Sokal, Alan D}, + journal={The American Mathematical Monthly}, + volume={118}, + number={5}, + pages={450--452}, + year={2011}, + publisher={Taylor \& Francis} +} + +@article{Weisstein_UBP, + title={Uniform Boundedness Principle}, + author={Moslehian, Mohammad Sal and Weisstein, Eric W.}, + journal={From MathWorld--A Wolfram Web Resource. http://mathworld.wolfram.com/UniformBoundednessPrinciple.html} +} diff --git a/thys/Banach_Steinhaus/document/root.tex b/thys/Banach_Steinhaus/document/root.tex --- a/thys/Banach_Steinhaus/document/root.tex +++ b/thys/Banach_Steinhaus/document/root.tex @@ -1,30 +1,30 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym} -\usepackage{amsmath,amssymb} - -%this should be the last package used -\usepackage{pdfsetup} - -% urls in roman style, theory text in math-similar italics -\urlstyle{rm} -\isabellestyle{it} - -\begin{document} - -\title{Banach-Steinhaus theorem} -\author{Dominique Unruh \and Jos\'e Manuel Rodr\'iguez Caballero} -\maketitle - -\begin{abstract} -We formalize in Isabelle/HOL a result \cite{Weisstein_UBP} due to S. Banach and H. Steinhaus \cite{banach1927principe} known as Banach-Steinhaus theorem or Uniform boundedness principle: a pointwise-bounded family of continuous linear operators from a Banach space to a normed space is uniformly bounded. Our approach is an adaptation to Isabelle/HOL of a proof due to A. Sokal \cite{sokal2011really}. -\end{abstract} - -\tableofcontents - -\input{session} - - -\bibliographystyle{abbrv} -\bibliography{root} - -\end{document} +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsmath,amssymb} + +%this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Banach-Steinhaus theorem} +\author{Dominique Unruh \and Jos\'e Manuel Rodr\'iguez Caballero} +\maketitle + +\begin{abstract} +We formalize in Isabelle/HOL a result \cite{Weisstein_UBP} due to S. Banach and H. Steinhaus \cite{banach1927principe} known as Banach-Steinhaus theorem or Uniform boundedness principle: a pointwise-bounded family of continuous linear operators from a Banach space to a normed space is uniformly bounded. Our approach is an adaptation to Isabelle/HOL of a proof due to A. Sokal \cite{sokal2011really}. +\end{abstract} + +\tableofcontents + +\input{session} + + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/FOL_Harrison/document/root.tex b/thys/FOL_Harrison/document/root.tex --- a/thys/FOL_Harrison/document/root.tex +++ b/thys/FOL_Harrison/document/root.tex @@ -1,80 +1,80 @@ -\documentclass[a4paper]{article} - -\title{First-Order Logic According to Harrison} - -\author{Alexander Birch Jensen, Anders Schlichtkrull \&\\ J{\o}rgen Villadsen, DTU Compute, Denmark} - -\date{\isadate\today} - -\usepackage[left=15mm,right=15mm,top=20mm,bottom=27mm]{geometry} - -\usepackage{datetime,isabelle,isabellesym,parskip} - -\newdateformat{isadate}{\THEDAY\ \monthname[\THEMONTH] \THEYEAR} - -\usepackage{pdfsetup} - -\isabellestyle{tt} - -\urlstyle{rm} - -%\renewcommand{\isachardoublequote}{} -%\renewcommand{\isachardoublequoteopen}{} -%\renewcommand{\isachardoublequoteclose}{} - -\renewcommand{\isamarkupchapter}[1]{\clearpage\isamarkupsection{#1}\medskip} -\renewcommand{\isamarkupsection}[1]{\medskip\section*{#1}\addcontentsline{toc}{section}{#1}\medskip} -\renewcommand{\isamarkupsubsection}[1]{\medskip\subsection*{#1}\medskip} -\renewcommand{\isamarkupsubsubsection}[1]{\medskip\subsubsection*{#1}\medskip} - -\renewcommand{\isabeginpar}{\par\ifisamarkup\relax\else\bigskip\fi} -\renewcommand{\isaendpar}{\par\bigskip} - -\begin{document} - -\makeatletter -\parbox[t]{\textwidth}{\centering\Huge\bfseries\@title}\par\kern5mm -\parbox[t]{\textwidth}{\centering\Large\bfseries\@author}\par\kern3mm -\parbox[t]{\textwidth}{\centering\bfseries\@date}\par\kern8mm -\makeatother - -\begin{abstract}\normalsize\noindent -We present a certified declarative first-order prover with equality based on John Harrison's -Handbook of Practical Logic and Automated Reasoning, Cambridge University Press, 2009. ML code -reflection is used such that the entire prover can be executed within Isabelle as a very simple -interactive proof assistant. As examples we consider Pelletier's problems 1-46. -\end{abstract} - -\tableofcontents - -\isamarkupsection{Preample} - -Preliminary formalizations are described here: -\begin{trivlist} -\item -Alexander Birch Jensen: -\emph{Development and Verification of a Proof Assistant.} -Master's Thesis, Technical University of Denmark, 2016. -\url{http://findit.dtu.dk/en/catalog/2345011633} -\item -Alexander Birch Jensen, Anders Schlichtkrull \&\ J{\o}rgen Villadsen: -\emph{Verification of an LCF-Style First-Order Prover with Equality.} -Isabelle Workshop 2016. -\url{https://github.com/logic-tools/sml-handbook} -\end{trivlist} - -\clearpage - -\input{session} - -\clearpage - -\isamarkupsection{Acknowledgements} - -The SML code is based on the OCaml code accompanying John Harrison's -Handbook of Practical Logic and Automated Reasoning, Cambridge University Press, 2009 - -Thanks to Jasmin Blanchette, Asta Halkj{\ae}r From, John Bruntse Larsen, Andrei Popescu and -Tom Ridge for discussions. - -\end{document} +\documentclass[a4paper]{article} + +\title{First-Order Logic According to Harrison} + +\author{Alexander Birch Jensen, Anders Schlichtkrull \&\\ J{\o}rgen Villadsen, DTU Compute, Denmark} + +\date{\isadate\today} + +\usepackage[left=15mm,right=15mm,top=20mm,bottom=27mm]{geometry} + +\usepackage{datetime,isabelle,isabellesym,parskip} + +\newdateformat{isadate}{\THEDAY\ \monthname[\THEMONTH] \THEYEAR} + +\usepackage{pdfsetup} + +\isabellestyle{tt} + +\urlstyle{rm} + +%\renewcommand{\isachardoublequote}{} +%\renewcommand{\isachardoublequoteopen}{} +%\renewcommand{\isachardoublequoteclose}{} + +\renewcommand{\isamarkupchapter}[1]{\clearpage\isamarkupsection{#1}\medskip} +\renewcommand{\isamarkupsection}[1]{\medskip\section*{#1}\addcontentsline{toc}{section}{#1}\medskip} +\renewcommand{\isamarkupsubsection}[1]{\medskip\subsection*{#1}\medskip} +\renewcommand{\isamarkupsubsubsection}[1]{\medskip\subsubsection*{#1}\medskip} + +\renewcommand{\isabeginpar}{\par\ifisamarkup\relax\else\bigskip\fi} +\renewcommand{\isaendpar}{\par\bigskip} + +\begin{document} + +\makeatletter +\parbox[t]{\textwidth}{\centering\Huge\bfseries\@title}\par\kern5mm +\parbox[t]{\textwidth}{\centering\Large\bfseries\@author}\par\kern3mm +\parbox[t]{\textwidth}{\centering\bfseries\@date}\par\kern8mm +\makeatother + +\begin{abstract}\normalsize\noindent +We present a certified declarative first-order prover with equality based on John Harrison's +Handbook of Practical Logic and Automated Reasoning, Cambridge University Press, 2009. ML code +reflection is used such that the entire prover can be executed within Isabelle as a very simple +interactive proof assistant. As examples we consider Pelletier's problems 1-46. +\end{abstract} + +\tableofcontents + +\isamarkupsection{Preample} + +Preliminary formalizations are described here: +\begin{trivlist} +\item +Alexander Birch Jensen: +\emph{Development and Verification of a Proof Assistant.} +Master's Thesis, Technical University of Denmark, 2016. +\url{http://findit.dtu.dk/en/catalog/2345011633} +\item +Alexander Birch Jensen, Anders Schlichtkrull \&\ J{\o}rgen Villadsen: +\emph{Verification of an LCF-Style First-Order Prover with Equality.} +Isabelle Workshop 2016. +\url{https://github.com/logic-tools/sml-handbook} +\end{trivlist} + +\clearpage + +\input{session} + +\clearpage + +\isamarkupsection{Acknowledgements} + +The SML code is based on the OCaml code accompanying John Harrison's +Handbook of Practical Logic and Automated Reasoning, Cambridge University Press, 2009 + +Thanks to Jasmin Blanchette, Asta Halkj{\ae}r From, John Bruntse Larsen, Andrei Popescu and +Tom Ridge for discussions. + +\end{document} diff --git a/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy b/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy --- a/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy +++ b/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy @@ -1,903 +1,903 @@ -theory IMP2_Binary_Heap - imports IMP2.IMP2 IMP2.IMP2_Aux_Lemmas -begin -section \Introduction\ -text \In this submission imperative versions of the following array-based binary minimum heap - functions are implemented and verified: insert, get-min, delete-min, make-heap. - The latter three are then used to prove the correctness of an in-place heapsort, which sorts - an array in descending order. To do that in Isabelle/HOL, the proof framework IMP2 - \cite{IMP2-AFP} is used. Here arrays are modeled by \int \ int\ functions. The imperative - implementations are iterative versions of the partly recursive algorithms described in - \cite{MS} and \cite{CLRS}. - - This submission starts with the basic definitions and lemmas, which are needed for array-based - binary heaps. These definitions and lemmas are parameterised with an arbitrary (transitive) - comparison function (where such a function is needed), so they are not only applicable to - minimum heaps. After some more general, useful lemmas on arrays, the imperative minimum heap - functions and the heapsort are implemented and verified.\ - -section \Heap Related Definitions and Theorems\ -subsection \Array Bounds\ -text \A small helper function is used to define valid array indices. Note that the lower index - bound \l\ is arbitrary and not fixed to 0 or 1. The upper index bound \r\ is not a valid - index itself, so that the empty array can be denoted by having \l = r\.\ -abbreviation bounded :: "int \ int \ int \ bool" where - "bounded l r x \ l \ x \ x < r" - -subsection \Parent and Children\ - -subsubsection \Definitions\ -text \For the notion of an array-based binary heap, the parent and child relations on the array - indices need to be defined.\ -definition parent :: "int \ int \ int" where - "parent l c = l + (c - l - 1) div 2" - -definition l_child :: "int \ int \ int" where - "l_child l p = 2 * p - l + 1" - -definition r_child :: "int \ int \ int" where - "r_child l p = 2 * p - l + 2" - -subsubsection \Lemmas\ -lemma parent_upper_bound: "parent l c < c \ l \ c" - unfolding parent_def by auto - -lemma parent_upper_bound_alt: "l \ parent l c \ parent l c < c" - unfolding parent_def by simp - -lemma parent_lower_bound: "l \ parent l c \ l < c" - unfolding parent_def by linarith - -lemma grand_parent_upper_bound: "parent l (parent l c) < c \ l \ c" - unfolding parent_def by linarith - -corollary parent_bounds: "l < x \ x < r \ bounded l r (parent l x)" - using parent_lower_bound parent_upper_bound_alt by fastforce - - -lemma l_child_lower_bound: " p < l_child l p \ l \ p" - unfolding l_child_def by simp - -corollary l_child_lower_bound_alt: "l \ x \ x \ p \ x < l_child l p" - using l_child_lower_bound[of p l] by linarith - -lemma parent_l_child[simp]: "parent l (l_child l n) = n" - unfolding parent_def l_child_def by simp - - -lemma r_child_lower_bound: "l \ p \ p < r_child l p" - unfolding r_child_def by simp - -corollary r_child_lower_bound_alt: "l \ x \ x \ p \ x < r_child l p" - using r_child_lower_bound[of l p] by linarith - -lemma parent_r_child[simp]: "parent l (r_child l n) = n" - unfolding parent_def r_child_def by simp - - -lemma smaller_l_child: "l_child l x < r_child l x" - unfolding l_child_def r_child_def by simp - -lemma parent_two_children: - "(c = l_child l p \ c = r_child l p) \ parent l c = p" - unfolding parent_def l_child_def r_child_def by linarith - -subsection \Heap Invariants\ -subsubsection \Definitions\ -text \The following heap invariants and the following lemmas are parameterised with an arbitrary - (transitive) comparison function. For the concrete function implementations at the end of - this submission \\\ on ints is used.\ - -text \For the \make_heap\ function, which transforms an unordered array into a valid heap, - the notion of a partial heap is needed. Here the heap invariant only holds for array indices - between a certain valid array index \m\ and \r\. The standard heap invariant is then - simply the special case where \m = l\.\ -definition is_partial_heap - :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where - "is_partial_heap cmp heap l m r = (\ x. bounded m r x \ - bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x))" - -abbreviation is_heap - :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ bool" where - "is_heap cmp heap l r \ is_partial_heap cmp heap l l r" - - -text \During all of the modifying heap functions the heap invariant is temporarily violated at - a single index \i\ and it is then gradually restored by either \sift_down\ or - \sift_up\. The following definitions formalize these weakened invariants. - - The second part of the conjunction in the following definitions states, that the comparison - between the parent of \i\ and each of the children of \i\ evaluates to \True\ without - explicitly using the child relations.\ -definition is_partial_heap_except_down - :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ int \ bool" where - "is_partial_heap_except_down cmp heap l m r i = (\ x. bounded m r x \ - ((parent l x \ i \ bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ - (parent l x = i \ bounded m r (parent l (parent l x)) - \ cmp (heap (parent l (parent l x))) (heap x))))" - -abbreviation is_heap_except_down - :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where - "is_heap_except_down cmp heap l r i \ is_partial_heap_except_down cmp heap l l r i" - - -text \As mentioned the notion of a partial heap is only needed for \make_heap\, - which only uses \sift_down\ internally, so there doesn't need to be an additional - definition for the partial heap version of the \sift_up\ invariant.\ -definition is_heap_except_up - :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where - "is_heap_except_up cmp heap l r i = (\ x. bounded l r x \ - ((x \ i \ bounded l r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ - (parent l x = i \ bounded l r (parent l (parent l x)) - \ cmp (heap (parent l (parent l x))) (heap x))))" - -subsubsection \Lemmas\ - -lemma empty_partial_heap[simp]: "is_partial_heap cmp heap l r r" - unfolding is_partial_heap_def by linarith - -lemma is_partial_heap_smaller_back: - "is_partial_heap cmp heap l m r \ r' \ r \ is_partial_heap cmp heap l m r'" - unfolding is_partial_heap_def by simp - -lemma is_partial_heap_smaller_front: - "is_partial_heap cmp heap l m r \ m \ m' \ is_partial_heap cmp heap l m' r" - unfolding is_partial_heap_def by simp - -text \The second half of each array is a is a partial binary heap, since it contains only leafs, - which are all trivial binary heaps.\ -lemma snd_half_is_partial_heap: - "(l + r) div 2 \ m \ is_partial_heap cmp heap l m r" - unfolding is_partial_heap_def parent_def by linarith - -lemma modify_outside_partial_heap: - assumes - "heap = heap' on {m..The next few lemmas formalize how the heap invariant is weakened, when the heap is modified - in a certain way.\ - -text \This lemma is used by \make_heap\.\ -lemma partial_heap_added_first_el: - assumes - "l \ m" "m \ r" - "is_partial_heap cmp heap l (m + 1) r" - shows "is_partial_heap_except_down cmp heap l m r m" - unfolding is_partial_heap_except_down_def -proof - fix x - let ?p_x = "parent l x" - let ?gp_x = "parent l ?p_x" - show "bounded m r x \ - (?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ - (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" - proof - assume x_bound: "bounded m r x" - have p_x_lower: "?p_x \ m \ bounded m r ?p_x \ ?p_x \ m + 1" - by simp - hence "?p_x \ m \ bounded m r ?p_x \ x \ m + 1" - using parent_upper_bound[of l x] x_bound assms(1) by linarith - hence p_invariant: "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x))" - using assms(3) is_partial_heap_def p_x_lower x_bound by blast - - have gp_up_bound: "(?p_x = m \ ?gp_x < m)" - by (simp add: assms(1) parent_upper_bound) - show "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ - (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" - using gp_up_bound p_invariant by linarith - qed -qed - -text \This lemma is used by \del_min\.\ -lemma heap_changed_first_el: - assumes "is_heap cmp heap l r" "l \ r" - shows "is_heap_except_down cmp (heap(l := b)) l r l" -proof - - have "is_partial_heap cmp heap l (l + 1) r" - using assms(1) is_partial_heap_smaller_front by fastforce - hence "is_partial_heap cmp (heap(l := b)) l (l + 1) r" - using modify_outside_partial_heap[of heap] by simp - thus ?thesis - by (simp add: assms(2) partial_heap_added_first_el) -qed - -text \This lemma is used by \insert\.\ -lemma heap_appended_el: - assumes - "is_heap cmp heap l r" - "heap = heap' on {l..First Heap Element\ -text \The next step is to show that the first element of the heap is always the ``smallest'' - according to the given comparison function. For the proof a rule for strong induction on lower - bounded integers is needed. Its proof is based on the proof of strong induction on natural - numbers found in \cite{Str_Ind}.\ -lemma strong_int_gr_induct_helper: - assumes "k < (i::int)" "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" - shows "\j. k < j \ j < i \ P j" - using assms -proof(induction i rule: int_gr_induct) - case base - then show ?case by linarith -next - case (step i) - then show ?case - proof(cases "j = i") - case True - then show ?thesis using step.IH step.prems(1,3) by blast - next - case False - hence "j < i" using step.prems(2) by simp - then show ?thesis using step.IH step.prems(1,3) by blast - qed -qed - -theorem strong_int_gr_induct: - assumes - "k < (i::int)" - "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" - shows "P i" - using assms less_induct strong_int_gr_induct_helper by blast - -text \Now the main theorem, that the first heap element is the ``smallest'' according to the - given comparison function, can be proven.\ -theorem heap_first_el: - assumes - "is_heap cmp heap l r" - "transp cmp" - "l < x" "x < r" - shows "cmp (heap l) (heap x)" - using assms unfolding is_partial_heap_def -proof(induction x rule: strong_int_gr_induct[of l]) - case 1 - then show ?case using assms(3) by simp -next - case (2 i) - have cmp_pi_i: "cmp (heap (parent l i)) (heap i)" - using "2.hyps" "2.prems"(1,4) parent_bounds by simp - then show ?case - proof(cases) - assume "parent l i > l" - then have "cmp (heap l) (heap (parent l i))" - using "2.IH" "2.prems"(1,2,4) parent_upper_bound_alt by simp - then show ?thesis - using "2.prems"(2) cmp_pi_i transpE by metis - next - assume "\ parent l i > l" - then have "parent l i = l" - using "2.hyps" dual_order.order_iff_strict parent_lower_bound by metis - then show ?thesis - using cmp_pi_i by simp - qed -qed - - - -section \General Lemmas on Arrays\ -text \Some additional lemmas on @{const "mset_ran"}, @{const "swap"} and @{const "eq_on"} are needed - for the final proofs.\ - -subsection \Lemmas on @{const "mset_ran"}\ -abbreviation arr_mset :: "(int \ 'a) \ int \ int \ 'a multiset" where - "arr_mset arr l r \ mset_ran arr {l..# (arr_mset arr l r) \ (\i. bounded l r i \ arr i = x)" - unfolding mset_ran_def by fastforce - -lemma arr_mset_remove_last: - "l \ r \ arr_mset arr l r = arr_mset arr l (r + 1) - {#arr r#}" - by (simp add: intvs_upper_decr mset_ran_def) - -lemma arr_mset_append: - "l \ r \ arr_mset arr l (r + 1) = arr_mset arr l r + {#arr r#}" - using arr_mset_remove_last[of l r arr] by simp - -corollary arr_mset_append_alt: - "l \ r \ arr_mset (arr(r := b)) l (r + 1) = arr_mset arr l r + {#b#}" - by (simp add: arr_mset_append mset_ran_subst_outside) - -lemma arr_mset_remove_first: - "i \ r \ arr_mset arr (i - 1) r = arr_mset arr i r + {#arr (i - 1)#}" - by(induction r rule: int_ge_induct) (auto simp add: arr_mset_append) - -lemma arr_mset_split: - assumes "l \ m" "m \ r" - shows "arr_mset arr l r = arr_mset arr l m + arr_mset arr m r" - using assms -proof(induction m rule: int_ge_induct[of l]) - case (step i) - have add_last: "arr_mset arr l (i + 1) = arr_mset arr l i + {#arr i#}" - using step arr_mset_append by blast - have rem_first: "arr_mset arr (i+1) r = arr_mset arr i r - {#arr i#}" - by (metis step.prems arr_mset_remove_first add_diff_cancel_right') - show ?case - using step add_last rem_first by fastforce -qed (simp) - -text \That the first element in a heap is the ``smallest'', can now be expressed using multisets.\ -corollary heap_first_el_alt: - assumes - "transp cmp" - "is_heap cmp heap l r" - "x \# (arr_mset heap l r)" - "heap l \ x" - shows "cmp (heap l) x" - by(metis assms heap_first_el in_mset_imp_in_array le_less) - - -subsection \Lemmas on @{term "swap"} and @{term "eq_on"}\ - -lemma eq_on_subset: - "arr1 = arr2 on R \ S \ R \ arr1 = arr2 on S" - by (simp add: eq_on_def set_mp) - -lemma swap_swaps: - "arr' = swap arr x y \ arr' y = arr x \ arr' x = arr y" - unfolding swap_def by simp - -lemma swap_only_swaps: - "arr' = swap arr x y \ z \ x \ z \ y \ arr' z = arr z" - unfolding swap_def by simp - -lemma swap_commute: "swap arr x y = swap arr y x" - unfolding swap_def by fastforce - -lemma swap_eq_on: - "arr1 = arr2 on S \ x \ S \ y \ S \ arr1 = swap arr2 x y on S" - unfolding swap_def by simp - -corollary swap_parent_eq_on: - assumes - "arr1 = arr2 on - {l.. c = r_child l p" - "l \ p" "c < r" - shows "arr1 = swap arr2 p c on - {l.. c = r_child l p" - "l \ p" "c < r" - shows "arr_mset arr1 l r = arr_mset (swap arr2 p c) l r" -proof - - have child_bounded: "l < c \ c < r" - by (metis assms(2-4) parent_lower_bound parent_two_children) - have parent_bounded: "bounded l r p" - by (metis assms(2-4) dual_order.strict_trans parent_two_children parent_upper_bound_alt) - thus ?thesis - using assms(1) child_bounded mset_ran_swap[of p "{l..The following lemma shows, which propositions have to hold on the pre-swap array, so that - a comparison between two elements holds on the post-swap array. This is useful for the - proofs of the loop invariants of \sift_up\ and \sift_down\. The lemma is kept - quite general (except for the argument order) and could probably be more closely - related to the parent relation for more concise proofs.\ -lemma cmp_swapI: - fixes arr::"'a::order \ 'a::order" - assumes - "m < n \ x < y" - "m < n \ x < y \ x = m \ y = n \ P (arr n) (arr m)" - "m < n \ x < y \ x \ m \ x \ n \ y \ m \ y \ n \ P (arr m) (arr n)" - "m < n \ x < y \ x = m \ y \ m \ y \ n \ P (arr y) (arr n)" - "m < n \ x < y \ x = n \ y \ m \ y \ n \ P (arr m) (arr y)" - "m < n \ x < y \ x \ m \ x \ n \ y = n \ P (arr m) (arr x)" - "m < n \ x < y \ x \ m \ x \ n \ y = m \ P (arr x) (arr n)" - shows "P (swap arr x y m) (swap arr x y n)" - by (metis assms order.asym swap_only_swaps swap_swaps) - -section \Imperative Heap Implementation\ -text \The following imperative heap functions are based on \cite{MS} and \cite{CLRS}. All functions, - that are recursive in these books, are iterative in the following implementations. The - function definitions are done with IMP2 \cite{IMP2-AFP}. From now on the heaps only contain - ints and only use \\\ as comparison function. The auxiliary lemmas used from now on are - heavily modeled after the proof goals, that are generated by the vcg tool (also part of IMP2).\ - -subsection \Simple Functions\ -subsubsection \Parent, Children and Swap\ -text \In this section the parent and children relations are expressed as IMP2 procedures. - Additionally a simple procedure, that swaps two array elements, is defined.\ -procedure_spec prnt (l, x) returns p - assumes True - ensures "p = parent l\<^sub>0 x\<^sub>0" - defines \p = ((x - l - 1) / 2 + l)\ - by vcg (simp add: parent_def) - -procedure_spec left_child (l, x) returns lc - assumes True - ensures "lc = l_child l\<^sub>0 x\<^sub>0" - defines \lc = 2 * x - l + 1\ - by vcg (simp add: l_child_def) - -procedure_spec right_child (l, x) returns rc - assumes True - ensures "rc = r_child l\<^sub>0 x\<^sub>0" - defines \rc = 2 * x - l + 2\ - by vcg (simp add: r_child_def) - -procedure_spec swp (heap, x, y) returns heap - assumes True - ensures "heap = swap heap\<^sub>0 x\<^sub>0 y\<^sub>0 " - defines \tmp = heap[x]; heap[x] = heap[y]; heap[y] = tmp\ - by vcg (simp add: swap_def) - -subsubsection \\get_min\\ -text \In this section \get_min\ is defined, which simply returns the first element (the minimum) of - the heap. For this definition an additional theorem is proven, which enables the use of - @{const "Min_mset"} in the postcondition.\ - -theorem heap_minimum: - assumes - "l < r" - "is_heap (\) heap l r" - shows "heap l = Min_mset (arr_mset heap l r)" -proof - - have "(\x \# (arr_mset heap l r). (heap l) \ x)" - using assms(2) heap_first_el_alt transp_le by blast - thus ?thesis - by (simp add: assms(1) dual_order.antisym) -qed - -procedure_spec get_min (heap, l, r) returns min - assumes "l < r \ is_heap (\) heap l r" - ensures "min = Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)" - for heap[] l r - defines \min = heap[l]\ - by vcg (simp add: heap_minimum) - -subsection \Modifying Functions\ - -subsubsection \\sift_up\ and \insert\\ -text \The next heap function is \insert\, which internally uses \sift_up\. In the beginning of - this section \sift_up_step\ is proven, which states that each \sift_up\ loop iteration - correctly transforms the weakened heap invariant. For its proof two additional - auxiliary lemmas are used. After \sift_up_step\ \sift_up\ and then \insert\ are verified.\ - -text \\sift_up_step\ can be proven directly by the smt-solver without auxiliary lemmas, but they - were introduced to show the proof details. The analogous proofs for \sift_down\ were - just solved with smt, since the proof structure should be very similar, even though the - \sift_down\ proof goals are slightly more complex.\ -lemma sift_up_step_aux1: - fixes heap::"int \ int" - assumes - "is_heap_except_up (\) heap l r x" - "parent l x \ l" - "(heap x) \ (heap (parent l x))" - "bounded l r k" - "k \ (parent l x)" - "bounded l r (parent l k)" - shows "(swap heap (parent l x) x (parent l k)) \ (swap heap (parent l x) x k)" - apply(intro cmp_swapI[of "(parent l k)" k "(parent l x)" x "(\)" heap]) - subgoal using assms(2,6) parent_upper_bound_alt by blast - subgoal using assms(3) by blast - subgoal using assms(1,4,6) unfolding is_heap_except_up_def by blast - subgoal using assms(1,3,4,6) unfolding is_heap_except_up_def by fastforce - subgoal using assms(5) by blast - subgoal by blast - subgoal using assms(1,2,4) unfolding is_heap_except_up_def by simp - done - -lemma sift_up_step_aux2: - fixes heap::"int \ int" - assumes - "is_heap_except_up (\) heap l r x" - "parent l x \ l" - "heap x \ (heap (parent l x))" - "bounded l r k" - "parent l k = parent l x" - "bounded l r (parent l (parent l k))" - shows - "swap heap (parent l x) x (parent l (parent l k)) \ swap heap (parent l x) x k" - using assms unfolding is_heap_except_up_def -proof- - let ?gp_k = "parent l (parent l k)" - let ?gp_x = "parent l (parent l x)" - have gp_k_eq_gp_x: "swap heap (parent l x) x ?gp_k = heap ?gp_x" - by (metis assms(2,5) grand_parent_upper_bound less_irrefl swap_only_swaps) - show ?thesis - using assms unfolding is_heap_except_up_def - proof(cases) - assume k_eq_x: "k = x" - have "swap heap (parent l x) x k = heap (parent l x)" - by (metis k_eq_x swap_swaps) - then show ?thesis - using assms(1,2,4,6) unfolding is_heap_except_up_def - by (metis gp_k_eq_gp_x k_eq_x parent_bounds parent_lower_bound) - next - assume k_neq_x: "k \ x" - have "swap heap (parent l x) x k = heap k" - by (metis assms(5) gp_k_eq_gp_x k_neq_x swap_only_swaps) - then show ?thesis using assms unfolding is_heap_except_up_def - by (metis gp_k_eq_gp_x k_neq_x order_trans parent_bounds parent_lower_bound) - qed -qed - -lemma sift_up_step: - fixes heap::"int \ int" - assumes - "is_heap_except_up (\) heap l r x" - "parent l x \ l" - "(heap x) \ (heap (parent l x))" - shows "is_heap_except_up (\) (swap heap (parent l x) x) l r (parent l x)" - using assms sift_up_step_aux1 sift_up_step_aux2 - unfolding is_heap_except_up_def by blast - -text \\sift_up\ restores the heap invariant, that is only violated at the current position, by - iteratively swapping the current element with its parent until the beginning of the array is - reached or the current element is bigger than its parent.\ -procedure_spec sift_up (heap, l, r, x) returns heap - assumes "is_heap_except_up (\) heap l r x \ bounded l r x" - ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ - arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ - heap\<^sub>0 = heap on - {l\<^sub>0..0}" - for heap[] l x r - defines \ - p = prnt(l, x); - while (x > l \ heap[x] \ heap[p]) - @variant \x - l\ - @invariant \is_heap_except_up (\) heap l r x \ p = parent l x \ - bounded l r x \ arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l r \ - heap\<^sub>0 = heap on - {l.. - { - heap = swp(heap, p, x); - x = p; - p = prnt(l, x) - }\ - apply vcg_cs - apply(intro conjI) - subgoal using parent_lower_bound sift_up_step by blast - subgoal using parent_lower_bound by blast - subgoal using parent_bounds by blast - subgoal using parent_bounds by (simp add: mset_ran_swap) - subgoal using swap_parent_eq_on by blast - subgoal using parent_upper_bound by simp - subgoal unfolding is_heap_except_up_def is_partial_heap_def - by (metis le_less not_less parent_lower_bound) - done - -text \\insert\ inserts an element into a heap by appending it to the heap and restoring the heap - invariant with @{const "sift_up"}.\ -procedure_spec insert (heap, l, r, el) returns (heap, l, r) - assumes "is_heap (\) heap l r \ l \ r" - ensures "is_heap (\) heap l r \ - arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 + {#el\<^sub>0#} \ - l = l\<^sub>0 \ r = r\<^sub>0 + 1 \ heap\<^sub>0 = heap on - {l.. - heap[r] = el; - x = r; - r = r + 1; - heap = sift_up(heap, l, r, x) - \ - apply vcg_cs - subgoal by (simp add: heap_appended_el) - subgoal by (metis arr_mset_append_alt add_mset_add_single) - done - -subsubsection \\sift_down\, \del_min\ and \make_heap\\ -text \The next heap functions are \del_min\ and \make_heap\, which both use \sift_down\ to - restore/establish the heap invariant. \sift_down\ is proven first (this time without - additional auxiliary lemmas) followed by \del_min\ and \make_heap\.\ - -text \\sift_down\ restores the heap invariant, that is only violated at the current position, by - iteratively swapping the current element with its smallest child until the end of - the array is reached or the current element is smaller than its children.\ -procedure_spec sift_down(heap, l, r, x) returns heap - assumes "is_partial_heap_except_down (\) heap l x r x \ l \ x \ x \ r" - ensures "is_partial_heap (\) heap l\<^sub>0 x\<^sub>0 r\<^sub>0 \ - arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ - heap\<^sub>0 = heap on - {l\<^sub>0..0}" - defines \ - lc = left_child(l, x); - rc = right_child(l, x); - while (lc < r \ (heap[lc] < heap[x] \ (rc < r \ heap[rc] < heap[x]))) - @variant \r - x\ - @invariant \is_partial_heap_except_down (\) heap l x\<^sub>0 r x \ - x\<^sub>0 \ x \ x \ r \ lc = l_child l x \ rc = r_child l x \ - arr_mset heap\<^sub>0 l r = arr_mset heap l r \ - heap\<^sub>0 = heap on - {l.. - { - smallest = lc; - if (rc < r \ heap[rc] < heap[lc]) { - smallest = rc - }; - heap = swp(heap, x, smallest); - x = smallest; - lc = left_child(l, x); - rc = right_child(l, x) - }\ - apply vcg_cs - subgoal - apply(intro conjI) - subgoal unfolding is_partial_heap_except_down_def - by (smt parent_two_children swap_swaps swap_only_swaps - swap_commute parent_upper_bound_alt) - subgoal using r_child_lower_bound_alt by fastforce - subgoal using swap_child_mset order_trans by blast - subgoal using swap_child_eq_on by fastforce - done - subgoal - by (meson less_le_trans not_le order.asym r_child_lower_bound) - subgoal - apply(intro conjI) - subgoal unfolding is_partial_heap_except_down_def - by (smt parent_two_children swap_swaps swap_only_swaps - swap_commute parent_upper_bound_alt) - subgoal using l_child_lower_bound_alt by fastforce - subgoal using swap_child_mset order_trans by blast - subgoal using swap_child_eq_on by fastforce - done - subgoal - by (meson less_le_trans not_le order.asym l_child_lower_bound) - subgoal unfolding is_partial_heap_except_down_def is_partial_heap_def - by (metis dual_order.strict_trans not_less parent_two_children smaller_l_child) - done - - -text \\del_min\ needs an additional lemma which shows, that it actually removes (only) the minimum - from the heap.\ -lemma del_min_mset: - fixes heap::"int \ int" - assumes - "l < r" - "is_heap (\) heap l r" - "mod_heap = heap(l := heap (r - 1))" - "arr_mset mod_heap l (r - 1) = arr_mset new_heap l (r - 1)" - shows - "arr_mset new_heap l (r - 1) = arr_mset heap l r - {#Min_mset (arr_mset heap l r)#}" -proof - - let ?heap_mset = "arr_mset heap l r" - have l_is_min: "heap l = Min_mset ?heap_mset" - using assms(1,2) heap_minimum by blast - have "(arr_mset mod_heap l r) = ?heap_mset + {#heap (r-1)#} - {#heap l#}" - by (simp add: assms(1,3) mset_ran_subst_inside) - hence "(arr_mset mod_heap l (r - 1)) = ?heap_mset - {#heap l#}" - by (simp add: assms(1,3) arr_mset_remove_last) - thus ?thesis - using assms(4) l_is_min by simp -qed - -text \\del_min\ removes the minimum element from the heap by replacing the first element with the - last element, shrinking the array by one and subsequently restoring the heap invariant - with @{const "sift_down"}.\ -procedure_spec del_min (heap, l, r) returns (heap, l, r) - assumes "l < r \ is_heap (\) heap l r" - ensures "is_heap (\) heap l r \ - arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 - {#Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)#} \ - l = l\<^sub>0 \ r = r\<^sub>0 - 1 \ - heap\<^sub>0 = heap on - {l\<^sub>0..0}" - for heap l r - defines \ - r = r - 1; - heap[l] = heap[r]; - heap = sift_down(heap, l, r, l) - \ - apply vcg_cs - subgoal by (simp add: heap_changed_first_el is_partial_heap_smaller_back) - subgoal - apply(rule conjI) - subgoal using del_min_mset by blast - subgoal by (simp add: eq_on_def intvs_incdec(3) intvs_lower_incr) - done - done - - -text \\make_heap\ transforms an arbitrary array into a heap by iterating through all array - positions from the middle of the array up to the beginning of the array and calling - @{const "sift_down"} for each one.\ -procedure_spec make_heap (heap, l, r) returns heap - assumes "l \ r" - ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ - arr_mset heap l\<^sub>0 r\<^sub>0 = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ - heap\<^sub>0 = heap on - {l\<^sub>0..< r\<^sub>0}" - for heap[] l r - defines \ - y = (r + l)/2 - 1; - while (y \ l) - @variant \y - l + 1\ - @invariant \is_partial_heap (\) heap l (y + 1) r \ - arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ - l - 1 \ y \ y < r \ heap\<^sub>0 = heap on - {l.. - { - heap = sift_down(heap, l, r, y); - y = y - 1 - }\ - apply(vcg_cs) - subgoal - apply(rule conjI) - subgoal by (simp add: snd_half_is_partial_heap add.commute) - subgoal by linarith - done - subgoal using partial_heap_added_first_el le_less by blast - subgoal using eq_on_trans by blast - subgoal using dual_order.antisym by fastforce - done - -subsection \Heapsort Implementation\ - -text \The final part of this submission is the implementation of the in-place heapsort. Firstly it - builds the \\\-heap and then it iteratively removes the minimum of the heap, which is put at - the now vacant end of the shrinking heap. This is done until the heap is empty, which leaves - the array sorted in descending order.\ -subsubsection \Auxiliary Lemmas\ -text \Firstly the notion of a sorted array is needed. This is more or less the same as - @{const "ran_sorted"} generalized for arbitrary comparison functions.\ -definition array_is_sorted :: "(int \ int \ bool) \ (int \ int) \ int \ int \ bool" where - "array_is_sorted cmp a l r \ \i. \j. bounded l r i \ bounded l r j \ i < j \ cmp (a i) (a j)" - -text \This lemma states, that the heapsort doesn't change the elements contained in the array during - the loop iterations.\ -lemma heap_sort_mset_step: - fixes arr::"int \ int" - assumes - "l < m" "m \ r" - "arr_mset arr' l (m - 1) = arr_mset arr l m - {#Min_mset (arr_mset arr l m)#}" - "arr = arr' on - {l..This lemma states, that each loop iteration leaves the growing second half of the array - sorted in descending order.\ -lemma heap_sort_second_half_sorted_step: - fixes arr::"int \ int" - assumes - "l\<^sub>0 < m" "m \ r\<^sub>0" - "arr = arr' on - {l\<^sub>0..i. \j. bounded m r\<^sub>0 i \ bounded m r\<^sub>0 j \ i < j \ arr j \ arr i" - "\x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" - "bounded (m - 1) r\<^sub>0 i" - "bounded (m - 1) r\<^sub>0 j" - "i < j" - "mod_arr = (arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m)))" - shows "mod_arr j \ mod_arr i" -proof - - have second_half_eq: "mod_arr = arr on {m..< r\<^sub>0}" - using assms(3, 9) unfolding eq_on_def by simp - have j_stricter_bound: "bounded m r\<^sub>0 j" - using assms(6-8) by simp - then have el_at_j: "mod_arr j \# arr_mset arr m r\<^sub>0" - using eq_onD second_half_eq by fastforce - then show ?thesis - proof(cases) - assume "i = (m-1)" - then have "mod_arr i \# arr_mset arr l\<^sub>0 m" - by (simp add: assms(1, 9)) - then show ?thesis - using assms(5) el_at_j not_less by blast - next - assume "i \ (m-1)" - then have "bounded m r\<^sub>0 i" - using assms(6) by simp - then show ?thesis - using assms(4, 8) eq_on_def j_stricter_bound second_half_eq by force - qed -qed - - -text \The following lemma shows that all elements in the first part of the array (the binary heap) - are bigger than the elements in the second part (the sorted part) after every iteration. This - lemma and the invariant of the \heap_sort\ loop use \\ x < y\ instead of \x \ y\ since - \vcg_cs\ doesn't terminate in the latter case.\ -lemma heap_sort_fst_part_bigger_snd_part_step: - fixes arr::"int \ int" - assumes - "l\<^sub>0 < m" - "m \ r\<^sub>0" - "arr_mset arr' l\<^sub>0 (m - 1) = arr_mset arr l\<^sub>0 m - {#Min_mset (arr_mset arr l\<^sub>0 m)#}" - "arr = arr' on - {l\<^sub>0..x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" - "mod_arr = arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m))" - "x\#arr_mset mod_arr l\<^sub>0 (m - 1)" - "y\#arr_mset mod_arr (m - 1) r\<^sub>0" - shows "\ x < y" -proof - - have "{m..0} \ - {l\<^sub>0..0}" - using assms(4) eq_on_sym eq_on_subset by blast - hence arr_eq_on: "mod_arr = arr on {m..0}" - by (simp add: assms(6)) - hence same_mset: "arr_mset mod_arr m r\<^sub>0 = arr_mset arr m r\<^sub>0" - using mset_ran_cong by blast - have "x \# arr_mset arr l\<^sub>0 m" using same_mset assms - by (metis assms(3,6,7) add_mset_remove_trivial_eq lran_upd_outside(2) - mset_lran cancel_ab_semigroup_add_class.diff_right_commute - diff_single_trivial multi_self_add_other_not_self order_refl) - then have x_bigger_min: "x \ Min_mset (arr_mset arr l\<^sub>0 m)" - using Min_le by blast - have y_smaller_min: "y \ Min_mset (arr_mset arr l\<^sub>0 m)" - proof(cases "y = mod_arr (m - 1)") - case False - hence "y\#arr_mset mod_arr (m - 1) r\<^sub>0 - {#mod_arr (m - 1)#}" - by (metis assms(8) diff_single_trivial insert_DiffM insert_noteq_member) - then have "y\#arr_mset arr m r\<^sub>0" - by (simp add: assms(2) intvs_decr_l mset_ran_insert same_mset) - then show ?thesis - using assms(1) assms(5) by fastforce - qed (simp add: assms(6)) - then show ?thesis - using x_bigger_min by linarith -qed - -subsubsection \Implementation\ -text \Now finally the correctness of the \heap_sort\ is shown. As mentioned, it starts by - transforming the array into a minimum heap using @{const "make_heap"}. Then in each - iteration it removes the first element from the heap with @{const "del_min"} after its value - was retrieved with @{const "get_min"}. This value is then put at the position freed by - @{const "del_min"}.\ -program_spec heap_sort - assumes "l \ r" - ensures "array_is_sorted (\) arr l\<^sub>0 r\<^sub>0 \ - arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ - arr\<^sub>0 = arr on - {l\<^sub>0 ..0 } \ l = l\<^sub>0 \ r = r\<^sub>0" - for l r arr[] - defines \ - arr = make_heap(arr, l, r); - m = r; - while (m > l) - @variant \m - l + 1\ - @invariant \is_heap (\) arr l m \ - array_is_sorted (\) arr m r\<^sub>0 \ - (\x \# arr_mset arr l\<^sub>0 m. \y \# arr_mset arr m r\<^sub>0. \ x < y) \ - arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ - l \ m \ m \ r\<^sub>0 \ l = l\<^sub>0 \ arr\<^sub>0 = arr on - {l\<^sub>0 ..0}\ - { - min = get_min(arr, l, m); - (arr, l, m) = del_min(arr, l, m); - arr[m] = min - } - \ - apply vcg_cs - subgoal unfolding array_is_sorted_def by simp - subgoal - apply(intro conjI) - subgoal unfolding is_partial_heap_def by simp - subgoal unfolding array_is_sorted_def using heap_sort_second_half_sorted_step - by blast - subgoal using heap_sort_fst_part_bigger_snd_part_step by blast - subgoal using heap_sort_mset_step by blast - subgoal unfolding eq_on_def - by (metis ComplD ComplI atLeastLessThan_iff less_le_trans) - done - done -end +theory IMP2_Binary_Heap + imports IMP2.IMP2 IMP2.IMP2_Aux_Lemmas +begin +section \Introduction\ +text \In this submission imperative versions of the following array-based binary minimum heap + functions are implemented and verified: insert, get-min, delete-min, make-heap. + The latter three are then used to prove the correctness of an in-place heapsort, which sorts + an array in descending order. To do that in Isabelle/HOL, the proof framework IMP2 + \cite{IMP2-AFP} is used. Here arrays are modeled by \int \ int\ functions. The imperative + implementations are iterative versions of the partly recursive algorithms described in + \cite{MS} and \cite{CLRS}. + + This submission starts with the basic definitions and lemmas, which are needed for array-based + binary heaps. These definitions and lemmas are parameterised with an arbitrary (transitive) + comparison function (where such a function is needed), so they are not only applicable to + minimum heaps. After some more general, useful lemmas on arrays, the imperative minimum heap + functions and the heapsort are implemented and verified.\ + +section \Heap Related Definitions and Theorems\ +subsection \Array Bounds\ +text \A small helper function is used to define valid array indices. Note that the lower index + bound \l\ is arbitrary and not fixed to 0 or 1. The upper index bound \r\ is not a valid + index itself, so that the empty array can be denoted by having \l = r\.\ +abbreviation bounded :: "int \ int \ int \ bool" where + "bounded l r x \ l \ x \ x < r" + +subsection \Parent and Children\ + +subsubsection \Definitions\ +text \For the notion of an array-based binary heap, the parent and child relations on the array + indices need to be defined.\ +definition parent :: "int \ int \ int" where + "parent l c = l + (c - l - 1) div 2" + +definition l_child :: "int \ int \ int" where + "l_child l p = 2 * p - l + 1" + +definition r_child :: "int \ int \ int" where + "r_child l p = 2 * p - l + 2" + +subsubsection \Lemmas\ +lemma parent_upper_bound: "parent l c < c \ l \ c" + unfolding parent_def by auto + +lemma parent_upper_bound_alt: "l \ parent l c \ parent l c < c" + unfolding parent_def by simp + +lemma parent_lower_bound: "l \ parent l c \ l < c" + unfolding parent_def by linarith + +lemma grand_parent_upper_bound: "parent l (parent l c) < c \ l \ c" + unfolding parent_def by linarith + +corollary parent_bounds: "l < x \ x < r \ bounded l r (parent l x)" + using parent_lower_bound parent_upper_bound_alt by fastforce + + +lemma l_child_lower_bound: " p < l_child l p \ l \ p" + unfolding l_child_def by simp + +corollary l_child_lower_bound_alt: "l \ x \ x \ p \ x < l_child l p" + using l_child_lower_bound[of p l] by linarith + +lemma parent_l_child[simp]: "parent l (l_child l n) = n" + unfolding parent_def l_child_def by simp + + +lemma r_child_lower_bound: "l \ p \ p < r_child l p" + unfolding r_child_def by simp + +corollary r_child_lower_bound_alt: "l \ x \ x \ p \ x < r_child l p" + using r_child_lower_bound[of l p] by linarith + +lemma parent_r_child[simp]: "parent l (r_child l n) = n" + unfolding parent_def r_child_def by simp + + +lemma smaller_l_child: "l_child l x < r_child l x" + unfolding l_child_def r_child_def by simp + +lemma parent_two_children: + "(c = l_child l p \ c = r_child l p) \ parent l c = p" + unfolding parent_def l_child_def r_child_def by linarith + +subsection \Heap Invariants\ +subsubsection \Definitions\ +text \The following heap invariants and the following lemmas are parameterised with an arbitrary + (transitive) comparison function. For the concrete function implementations at the end of + this submission \\\ on ints is used.\ + +text \For the \make_heap\ function, which transforms an unordered array into a valid heap, + the notion of a partial heap is needed. Here the heap invariant only holds for array indices + between a certain valid array index \m\ and \r\. The standard heap invariant is then + simply the special case where \m = l\.\ +definition is_partial_heap + :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where + "is_partial_heap cmp heap l m r = (\ x. bounded m r x \ + bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x))" + +abbreviation is_heap + :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ bool" where + "is_heap cmp heap l r \ is_partial_heap cmp heap l l r" + + +text \During all of the modifying heap functions the heap invariant is temporarily violated at + a single index \i\ and it is then gradually restored by either \sift_down\ or + \sift_up\. The following definitions formalize these weakened invariants. + + The second part of the conjunction in the following definitions states, that the comparison + between the parent of \i\ and each of the children of \i\ evaluates to \True\ without + explicitly using the child relations.\ +definition is_partial_heap_except_down + :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ int \ bool" where + "is_partial_heap_except_down cmp heap l m r i = (\ x. bounded m r x \ + ((parent l x \ i \ bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ + (parent l x = i \ bounded m r (parent l (parent l x)) + \ cmp (heap (parent l (parent l x))) (heap x))))" + +abbreviation is_heap_except_down + :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where + "is_heap_except_down cmp heap l r i \ is_partial_heap_except_down cmp heap l l r i" + + +text \As mentioned the notion of a partial heap is only needed for \make_heap\, + which only uses \sift_down\ internally, so there doesn't need to be an additional + definition for the partial heap version of the \sift_up\ invariant.\ +definition is_heap_except_up + :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where + "is_heap_except_up cmp heap l r i = (\ x. bounded l r x \ + ((x \ i \ bounded l r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ + (parent l x = i \ bounded l r (parent l (parent l x)) + \ cmp (heap (parent l (parent l x))) (heap x))))" + +subsubsection \Lemmas\ + +lemma empty_partial_heap[simp]: "is_partial_heap cmp heap l r r" + unfolding is_partial_heap_def by linarith + +lemma is_partial_heap_smaller_back: + "is_partial_heap cmp heap l m r \ r' \ r \ is_partial_heap cmp heap l m r'" + unfolding is_partial_heap_def by simp + +lemma is_partial_heap_smaller_front: + "is_partial_heap cmp heap l m r \ m \ m' \ is_partial_heap cmp heap l m' r" + unfolding is_partial_heap_def by simp + +text \The second half of each array is a is a partial binary heap, since it contains only leafs, + which are all trivial binary heaps.\ +lemma snd_half_is_partial_heap: + "(l + r) div 2 \ m \ is_partial_heap cmp heap l m r" + unfolding is_partial_heap_def parent_def by linarith + +lemma modify_outside_partial_heap: + assumes + "heap = heap' on {m..The next few lemmas formalize how the heap invariant is weakened, when the heap is modified + in a certain way.\ + +text \This lemma is used by \make_heap\.\ +lemma partial_heap_added_first_el: + assumes + "l \ m" "m \ r" + "is_partial_heap cmp heap l (m + 1) r" + shows "is_partial_heap_except_down cmp heap l m r m" + unfolding is_partial_heap_except_down_def +proof + fix x + let ?p_x = "parent l x" + let ?gp_x = "parent l ?p_x" + show "bounded m r x \ + (?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ + (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" + proof + assume x_bound: "bounded m r x" + have p_x_lower: "?p_x \ m \ bounded m r ?p_x \ ?p_x \ m + 1" + by simp + hence "?p_x \ m \ bounded m r ?p_x \ x \ m + 1" + using parent_upper_bound[of l x] x_bound assms(1) by linarith + hence p_invariant: "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x))" + using assms(3) is_partial_heap_def p_x_lower x_bound by blast + + have gp_up_bound: "(?p_x = m \ ?gp_x < m)" + by (simp add: assms(1) parent_upper_bound) + show "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ + (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" + using gp_up_bound p_invariant by linarith + qed +qed + +text \This lemma is used by \del_min\.\ +lemma heap_changed_first_el: + assumes "is_heap cmp heap l r" "l \ r" + shows "is_heap_except_down cmp (heap(l := b)) l r l" +proof - + have "is_partial_heap cmp heap l (l + 1) r" + using assms(1) is_partial_heap_smaller_front by fastforce + hence "is_partial_heap cmp (heap(l := b)) l (l + 1) r" + using modify_outside_partial_heap[of heap] by simp + thus ?thesis + by (simp add: assms(2) partial_heap_added_first_el) +qed + +text \This lemma is used by \insert\.\ +lemma heap_appended_el: + assumes + "is_heap cmp heap l r" + "heap = heap' on {l..First Heap Element\ +text \The next step is to show that the first element of the heap is always the ``smallest'' + according to the given comparison function. For the proof a rule for strong induction on lower + bounded integers is needed. Its proof is based on the proof of strong induction on natural + numbers found in \cite{Str_Ind}.\ +lemma strong_int_gr_induct_helper: + assumes "k < (i::int)" "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" + shows "\j. k < j \ j < i \ P j" + using assms +proof(induction i rule: int_gr_induct) + case base + then show ?case by linarith +next + case (step i) + then show ?case + proof(cases "j = i") + case True + then show ?thesis using step.IH step.prems(1,3) by blast + next + case False + hence "j < i" using step.prems(2) by simp + then show ?thesis using step.IH step.prems(1,3) by blast + qed +qed + +theorem strong_int_gr_induct: + assumes + "k < (i::int)" + "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" + shows "P i" + using assms less_induct strong_int_gr_induct_helper by blast + +text \Now the main theorem, that the first heap element is the ``smallest'' according to the + given comparison function, can be proven.\ +theorem heap_first_el: + assumes + "is_heap cmp heap l r" + "transp cmp" + "l < x" "x < r" + shows "cmp (heap l) (heap x)" + using assms unfolding is_partial_heap_def +proof(induction x rule: strong_int_gr_induct[of l]) + case 1 + then show ?case using assms(3) by simp +next + case (2 i) + have cmp_pi_i: "cmp (heap (parent l i)) (heap i)" + using "2.hyps" "2.prems"(1,4) parent_bounds by simp + then show ?case + proof(cases) + assume "parent l i > l" + then have "cmp (heap l) (heap (parent l i))" + using "2.IH" "2.prems"(1,2,4) parent_upper_bound_alt by simp + then show ?thesis + using "2.prems"(2) cmp_pi_i transpE by metis + next + assume "\ parent l i > l" + then have "parent l i = l" + using "2.hyps" dual_order.order_iff_strict parent_lower_bound by metis + then show ?thesis + using cmp_pi_i by simp + qed +qed + + + +section \General Lemmas on Arrays\ +text \Some additional lemmas on @{const "mset_ran"}, @{const "swap"} and @{const "eq_on"} are needed + for the final proofs.\ + +subsection \Lemmas on @{const "mset_ran"}\ +abbreviation arr_mset :: "(int \ 'a) \ int \ int \ 'a multiset" where + "arr_mset arr l r \ mset_ran arr {l..# (arr_mset arr l r) \ (\i. bounded l r i \ arr i = x)" + unfolding mset_ran_def by fastforce + +lemma arr_mset_remove_last: + "l \ r \ arr_mset arr l r = arr_mset arr l (r + 1) - {#arr r#}" + by (simp add: intvs_upper_decr mset_ran_def) + +lemma arr_mset_append: + "l \ r \ arr_mset arr l (r + 1) = arr_mset arr l r + {#arr r#}" + using arr_mset_remove_last[of l r arr] by simp + +corollary arr_mset_append_alt: + "l \ r \ arr_mset (arr(r := b)) l (r + 1) = arr_mset arr l r + {#b#}" + by (simp add: arr_mset_append mset_ran_subst_outside) + +lemma arr_mset_remove_first: + "i \ r \ arr_mset arr (i - 1) r = arr_mset arr i r + {#arr (i - 1)#}" + by(induction r rule: int_ge_induct) (auto simp add: arr_mset_append) + +lemma arr_mset_split: + assumes "l \ m" "m \ r" + shows "arr_mset arr l r = arr_mset arr l m + arr_mset arr m r" + using assms +proof(induction m rule: int_ge_induct[of l]) + case (step i) + have add_last: "arr_mset arr l (i + 1) = arr_mset arr l i + {#arr i#}" + using step arr_mset_append by blast + have rem_first: "arr_mset arr (i+1) r = arr_mset arr i r - {#arr i#}" + by (metis step.prems arr_mset_remove_first add_diff_cancel_right') + show ?case + using step add_last rem_first by fastforce +qed (simp) + +text \That the first element in a heap is the ``smallest'', can now be expressed using multisets.\ +corollary heap_first_el_alt: + assumes + "transp cmp" + "is_heap cmp heap l r" + "x \# (arr_mset heap l r)" + "heap l \ x" + shows "cmp (heap l) x" + by(metis assms heap_first_el in_mset_imp_in_array le_less) + + +subsection \Lemmas on @{term "swap"} and @{term "eq_on"}\ + +lemma eq_on_subset: + "arr1 = arr2 on R \ S \ R \ arr1 = arr2 on S" + by (simp add: eq_on_def set_mp) + +lemma swap_swaps: + "arr' = swap arr x y \ arr' y = arr x \ arr' x = arr y" + unfolding swap_def by simp + +lemma swap_only_swaps: + "arr' = swap arr x y \ z \ x \ z \ y \ arr' z = arr z" + unfolding swap_def by simp + +lemma swap_commute: "swap arr x y = swap arr y x" + unfolding swap_def by fastforce + +lemma swap_eq_on: + "arr1 = arr2 on S \ x \ S \ y \ S \ arr1 = swap arr2 x y on S" + unfolding swap_def by simp + +corollary swap_parent_eq_on: + assumes + "arr1 = arr2 on - {l.. c = r_child l p" + "l \ p" "c < r" + shows "arr1 = swap arr2 p c on - {l.. c = r_child l p" + "l \ p" "c < r" + shows "arr_mset arr1 l r = arr_mset (swap arr2 p c) l r" +proof - + have child_bounded: "l < c \ c < r" + by (metis assms(2-4) parent_lower_bound parent_two_children) + have parent_bounded: "bounded l r p" + by (metis assms(2-4) dual_order.strict_trans parent_two_children parent_upper_bound_alt) + thus ?thesis + using assms(1) child_bounded mset_ran_swap[of p "{l..The following lemma shows, which propositions have to hold on the pre-swap array, so that + a comparison between two elements holds on the post-swap array. This is useful for the + proofs of the loop invariants of \sift_up\ and \sift_down\. The lemma is kept + quite general (except for the argument order) and could probably be more closely + related to the parent relation for more concise proofs.\ +lemma cmp_swapI: + fixes arr::"'a::order \ 'a::order" + assumes + "m < n \ x < y" + "m < n \ x < y \ x = m \ y = n \ P (arr n) (arr m)" + "m < n \ x < y \ x \ m \ x \ n \ y \ m \ y \ n \ P (arr m) (arr n)" + "m < n \ x < y \ x = m \ y \ m \ y \ n \ P (arr y) (arr n)" + "m < n \ x < y \ x = n \ y \ m \ y \ n \ P (arr m) (arr y)" + "m < n \ x < y \ x \ m \ x \ n \ y = n \ P (arr m) (arr x)" + "m < n \ x < y \ x \ m \ x \ n \ y = m \ P (arr x) (arr n)" + shows "P (swap arr x y m) (swap arr x y n)" + by (metis assms order.asym swap_only_swaps swap_swaps) + +section \Imperative Heap Implementation\ +text \The following imperative heap functions are based on \cite{MS} and \cite{CLRS}. All functions, + that are recursive in these books, are iterative in the following implementations. The + function definitions are done with IMP2 \cite{IMP2-AFP}. From now on the heaps only contain + ints and only use \\\ as comparison function. The auxiliary lemmas used from now on are + heavily modeled after the proof goals, that are generated by the vcg tool (also part of IMP2).\ + +subsection \Simple Functions\ +subsubsection \Parent, Children and Swap\ +text \In this section the parent and children relations are expressed as IMP2 procedures. + Additionally a simple procedure, that swaps two array elements, is defined.\ +procedure_spec prnt (l, x) returns p + assumes True + ensures "p = parent l\<^sub>0 x\<^sub>0" + defines \p = ((x - l - 1) / 2 + l)\ + by vcg (simp add: parent_def) + +procedure_spec left_child (l, x) returns lc + assumes True + ensures "lc = l_child l\<^sub>0 x\<^sub>0" + defines \lc = 2 * x - l + 1\ + by vcg (simp add: l_child_def) + +procedure_spec right_child (l, x) returns rc + assumes True + ensures "rc = r_child l\<^sub>0 x\<^sub>0" + defines \rc = 2 * x - l + 2\ + by vcg (simp add: r_child_def) + +procedure_spec swp (heap, x, y) returns heap + assumes True + ensures "heap = swap heap\<^sub>0 x\<^sub>0 y\<^sub>0 " + defines \tmp = heap[x]; heap[x] = heap[y]; heap[y] = tmp\ + by vcg (simp add: swap_def) + +subsubsection \\get_min\\ +text \In this section \get_min\ is defined, which simply returns the first element (the minimum) of + the heap. For this definition an additional theorem is proven, which enables the use of + @{const "Min_mset"} in the postcondition.\ + +theorem heap_minimum: + assumes + "l < r" + "is_heap (\) heap l r" + shows "heap l = Min_mset (arr_mset heap l r)" +proof - + have "(\x \# (arr_mset heap l r). (heap l) \ x)" + using assms(2) heap_first_el_alt transp_le by blast + thus ?thesis + by (simp add: assms(1) dual_order.antisym) +qed + +procedure_spec get_min (heap, l, r) returns min + assumes "l < r \ is_heap (\) heap l r" + ensures "min = Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)" + for heap[] l r + defines \min = heap[l]\ + by vcg (simp add: heap_minimum) + +subsection \Modifying Functions\ + +subsubsection \\sift_up\ and \insert\\ +text \The next heap function is \insert\, which internally uses \sift_up\. In the beginning of + this section \sift_up_step\ is proven, which states that each \sift_up\ loop iteration + correctly transforms the weakened heap invariant. For its proof two additional + auxiliary lemmas are used. After \sift_up_step\ \sift_up\ and then \insert\ are verified.\ + +text \\sift_up_step\ can be proven directly by the smt-solver without auxiliary lemmas, but they + were introduced to show the proof details. The analogous proofs for \sift_down\ were + just solved with smt, since the proof structure should be very similar, even though the + \sift_down\ proof goals are slightly more complex.\ +lemma sift_up_step_aux1: + fixes heap::"int \ int" + assumes + "is_heap_except_up (\) heap l r x" + "parent l x \ l" + "(heap x) \ (heap (parent l x))" + "bounded l r k" + "k \ (parent l x)" + "bounded l r (parent l k)" + shows "(swap heap (parent l x) x (parent l k)) \ (swap heap (parent l x) x k)" + apply(intro cmp_swapI[of "(parent l k)" k "(parent l x)" x "(\)" heap]) + subgoal using assms(2,6) parent_upper_bound_alt by blast + subgoal using assms(3) by blast + subgoal using assms(1,4,6) unfolding is_heap_except_up_def by blast + subgoal using assms(1,3,4,6) unfolding is_heap_except_up_def by fastforce + subgoal using assms(5) by blast + subgoal by blast + subgoal using assms(1,2,4) unfolding is_heap_except_up_def by simp + done + +lemma sift_up_step_aux2: + fixes heap::"int \ int" + assumes + "is_heap_except_up (\) heap l r x" + "parent l x \ l" + "heap x \ (heap (parent l x))" + "bounded l r k" + "parent l k = parent l x" + "bounded l r (parent l (parent l k))" + shows + "swap heap (parent l x) x (parent l (parent l k)) \ swap heap (parent l x) x k" + using assms unfolding is_heap_except_up_def +proof- + let ?gp_k = "parent l (parent l k)" + let ?gp_x = "parent l (parent l x)" + have gp_k_eq_gp_x: "swap heap (parent l x) x ?gp_k = heap ?gp_x" + by (metis assms(2,5) grand_parent_upper_bound less_irrefl swap_only_swaps) + show ?thesis + using assms unfolding is_heap_except_up_def + proof(cases) + assume k_eq_x: "k = x" + have "swap heap (parent l x) x k = heap (parent l x)" + by (metis k_eq_x swap_swaps) + then show ?thesis + using assms(1,2,4,6) unfolding is_heap_except_up_def + by (metis gp_k_eq_gp_x k_eq_x parent_bounds parent_lower_bound) + next + assume k_neq_x: "k \ x" + have "swap heap (parent l x) x k = heap k" + by (metis assms(5) gp_k_eq_gp_x k_neq_x swap_only_swaps) + then show ?thesis using assms unfolding is_heap_except_up_def + by (metis gp_k_eq_gp_x k_neq_x order_trans parent_bounds parent_lower_bound) + qed +qed + +lemma sift_up_step: + fixes heap::"int \ int" + assumes + "is_heap_except_up (\) heap l r x" + "parent l x \ l" + "(heap x) \ (heap (parent l x))" + shows "is_heap_except_up (\) (swap heap (parent l x) x) l r (parent l x)" + using assms sift_up_step_aux1 sift_up_step_aux2 + unfolding is_heap_except_up_def by blast + +text \\sift_up\ restores the heap invariant, that is only violated at the current position, by + iteratively swapping the current element with its parent until the beginning of the array is + reached or the current element is bigger than its parent.\ +procedure_spec sift_up (heap, l, r, x) returns heap + assumes "is_heap_except_up (\) heap l r x \ bounded l r x" + ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ + arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ + heap\<^sub>0 = heap on - {l\<^sub>0..0}" + for heap[] l x r + defines \ + p = prnt(l, x); + while (x > l \ heap[x] \ heap[p]) + @variant \x - l\ + @invariant \is_heap_except_up (\) heap l r x \ p = parent l x \ + bounded l r x \ arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l r \ + heap\<^sub>0 = heap on - {l.. + { + heap = swp(heap, p, x); + x = p; + p = prnt(l, x) + }\ + apply vcg_cs + apply(intro conjI) + subgoal using parent_lower_bound sift_up_step by blast + subgoal using parent_lower_bound by blast + subgoal using parent_bounds by blast + subgoal using parent_bounds by (simp add: mset_ran_swap) + subgoal using swap_parent_eq_on by blast + subgoal using parent_upper_bound by simp + subgoal unfolding is_heap_except_up_def is_partial_heap_def + by (metis le_less not_less parent_lower_bound) + done + +text \\insert\ inserts an element into a heap by appending it to the heap and restoring the heap + invariant with @{const "sift_up"}.\ +procedure_spec insert (heap, l, r, el) returns (heap, l, r) + assumes "is_heap (\) heap l r \ l \ r" + ensures "is_heap (\) heap l r \ + arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 + {#el\<^sub>0#} \ + l = l\<^sub>0 \ r = r\<^sub>0 + 1 \ heap\<^sub>0 = heap on - {l.. + heap[r] = el; + x = r; + r = r + 1; + heap = sift_up(heap, l, r, x) + \ + apply vcg_cs + subgoal by (simp add: heap_appended_el) + subgoal by (metis arr_mset_append_alt add_mset_add_single) + done + +subsubsection \\sift_down\, \del_min\ and \make_heap\\ +text \The next heap functions are \del_min\ and \make_heap\, which both use \sift_down\ to + restore/establish the heap invariant. \sift_down\ is proven first (this time without + additional auxiliary lemmas) followed by \del_min\ and \make_heap\.\ + +text \\sift_down\ restores the heap invariant, that is only violated at the current position, by + iteratively swapping the current element with its smallest child until the end of + the array is reached or the current element is smaller than its children.\ +procedure_spec sift_down(heap, l, r, x) returns heap + assumes "is_partial_heap_except_down (\) heap l x r x \ l \ x \ x \ r" + ensures "is_partial_heap (\) heap l\<^sub>0 x\<^sub>0 r\<^sub>0 \ + arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ + heap\<^sub>0 = heap on - {l\<^sub>0..0}" + defines \ + lc = left_child(l, x); + rc = right_child(l, x); + while (lc < r \ (heap[lc] < heap[x] \ (rc < r \ heap[rc] < heap[x]))) + @variant \r - x\ + @invariant \is_partial_heap_except_down (\) heap l x\<^sub>0 r x \ + x\<^sub>0 \ x \ x \ r \ lc = l_child l x \ rc = r_child l x \ + arr_mset heap\<^sub>0 l r = arr_mset heap l r \ + heap\<^sub>0 = heap on - {l.. + { + smallest = lc; + if (rc < r \ heap[rc] < heap[lc]) { + smallest = rc + }; + heap = swp(heap, x, smallest); + x = smallest; + lc = left_child(l, x); + rc = right_child(l, x) + }\ + apply vcg_cs + subgoal + apply(intro conjI) + subgoal unfolding is_partial_heap_except_down_def + by (smt parent_two_children swap_swaps swap_only_swaps + swap_commute parent_upper_bound_alt) + subgoal using r_child_lower_bound_alt by fastforce + subgoal using swap_child_mset order_trans by blast + subgoal using swap_child_eq_on by fastforce + done + subgoal + by (meson less_le_trans not_le order.asym r_child_lower_bound) + subgoal + apply(intro conjI) + subgoal unfolding is_partial_heap_except_down_def + by (smt parent_two_children swap_swaps swap_only_swaps + swap_commute parent_upper_bound_alt) + subgoal using l_child_lower_bound_alt by fastforce + subgoal using swap_child_mset order_trans by blast + subgoal using swap_child_eq_on by fastforce + done + subgoal + by (meson less_le_trans not_le order.asym l_child_lower_bound) + subgoal unfolding is_partial_heap_except_down_def is_partial_heap_def + by (metis dual_order.strict_trans not_less parent_two_children smaller_l_child) + done + + +text \\del_min\ needs an additional lemma which shows, that it actually removes (only) the minimum + from the heap.\ +lemma del_min_mset: + fixes heap::"int \ int" + assumes + "l < r" + "is_heap (\) heap l r" + "mod_heap = heap(l := heap (r - 1))" + "arr_mset mod_heap l (r - 1) = arr_mset new_heap l (r - 1)" + shows + "arr_mset new_heap l (r - 1) = arr_mset heap l r - {#Min_mset (arr_mset heap l r)#}" +proof - + let ?heap_mset = "arr_mset heap l r" + have l_is_min: "heap l = Min_mset ?heap_mset" + using assms(1,2) heap_minimum by blast + have "(arr_mset mod_heap l r) = ?heap_mset + {#heap (r-1)#} - {#heap l#}" + by (simp add: assms(1,3) mset_ran_subst_inside) + hence "(arr_mset mod_heap l (r - 1)) = ?heap_mset - {#heap l#}" + by (simp add: assms(1,3) arr_mset_remove_last) + thus ?thesis + using assms(4) l_is_min by simp +qed + +text \\del_min\ removes the minimum element from the heap by replacing the first element with the + last element, shrinking the array by one and subsequently restoring the heap invariant + with @{const "sift_down"}.\ +procedure_spec del_min (heap, l, r) returns (heap, l, r) + assumes "l < r \ is_heap (\) heap l r" + ensures "is_heap (\) heap l r \ + arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 - {#Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)#} \ + l = l\<^sub>0 \ r = r\<^sub>0 - 1 \ + heap\<^sub>0 = heap on - {l\<^sub>0..0}" + for heap l r + defines \ + r = r - 1; + heap[l] = heap[r]; + heap = sift_down(heap, l, r, l) + \ + apply vcg_cs + subgoal by (simp add: heap_changed_first_el is_partial_heap_smaller_back) + subgoal + apply(rule conjI) + subgoal using del_min_mset by blast + subgoal by (simp add: eq_on_def intvs_incdec(3) intvs_lower_incr) + done + done + + +text \\make_heap\ transforms an arbitrary array into a heap by iterating through all array + positions from the middle of the array up to the beginning of the array and calling + @{const "sift_down"} for each one.\ +procedure_spec make_heap (heap, l, r) returns heap + assumes "l \ r" + ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ + arr_mset heap l\<^sub>0 r\<^sub>0 = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ + heap\<^sub>0 = heap on - {l\<^sub>0..< r\<^sub>0}" + for heap[] l r + defines \ + y = (r + l)/2 - 1; + while (y \ l) + @variant \y - l + 1\ + @invariant \is_partial_heap (\) heap l (y + 1) r \ + arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ + l - 1 \ y \ y < r \ heap\<^sub>0 = heap on - {l.. + { + heap = sift_down(heap, l, r, y); + y = y - 1 + }\ + apply(vcg_cs) + subgoal + apply(rule conjI) + subgoal by (simp add: snd_half_is_partial_heap add.commute) + subgoal by linarith + done + subgoal using partial_heap_added_first_el le_less by blast + subgoal using eq_on_trans by blast + subgoal using dual_order.antisym by fastforce + done + +subsection \Heapsort Implementation\ + +text \The final part of this submission is the implementation of the in-place heapsort. Firstly it + builds the \\\-heap and then it iteratively removes the minimum of the heap, which is put at + the now vacant end of the shrinking heap. This is done until the heap is empty, which leaves + the array sorted in descending order.\ +subsubsection \Auxiliary Lemmas\ +text \Firstly the notion of a sorted array is needed. This is more or less the same as + @{const "ran_sorted"} generalized for arbitrary comparison functions.\ +definition array_is_sorted :: "(int \ int \ bool) \ (int \ int) \ int \ int \ bool" where + "array_is_sorted cmp a l r \ \i. \j. bounded l r i \ bounded l r j \ i < j \ cmp (a i) (a j)" + +text \This lemma states, that the heapsort doesn't change the elements contained in the array during + the loop iterations.\ +lemma heap_sort_mset_step: + fixes arr::"int \ int" + assumes + "l < m" "m \ r" + "arr_mset arr' l (m - 1) = arr_mset arr l m - {#Min_mset (arr_mset arr l m)#}" + "arr = arr' on - {l..This lemma states, that each loop iteration leaves the growing second half of the array + sorted in descending order.\ +lemma heap_sort_second_half_sorted_step: + fixes arr::"int \ int" + assumes + "l\<^sub>0 < m" "m \ r\<^sub>0" + "arr = arr' on - {l\<^sub>0..i. \j. bounded m r\<^sub>0 i \ bounded m r\<^sub>0 j \ i < j \ arr j \ arr i" + "\x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" + "bounded (m - 1) r\<^sub>0 i" + "bounded (m - 1) r\<^sub>0 j" + "i < j" + "mod_arr = (arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m)))" + shows "mod_arr j \ mod_arr i" +proof - + have second_half_eq: "mod_arr = arr on {m..< r\<^sub>0}" + using assms(3, 9) unfolding eq_on_def by simp + have j_stricter_bound: "bounded m r\<^sub>0 j" + using assms(6-8) by simp + then have el_at_j: "mod_arr j \# arr_mset arr m r\<^sub>0" + using eq_onD second_half_eq by fastforce + then show ?thesis + proof(cases) + assume "i = (m-1)" + then have "mod_arr i \# arr_mset arr l\<^sub>0 m" + by (simp add: assms(1, 9)) + then show ?thesis + using assms(5) el_at_j not_less by blast + next + assume "i \ (m-1)" + then have "bounded m r\<^sub>0 i" + using assms(6) by simp + then show ?thesis + using assms(4, 8) eq_on_def j_stricter_bound second_half_eq by force + qed +qed + + +text \The following lemma shows that all elements in the first part of the array (the binary heap) + are bigger than the elements in the second part (the sorted part) after every iteration. This + lemma and the invariant of the \heap_sort\ loop use \\ x < y\ instead of \x \ y\ since + \vcg_cs\ doesn't terminate in the latter case.\ +lemma heap_sort_fst_part_bigger_snd_part_step: + fixes arr::"int \ int" + assumes + "l\<^sub>0 < m" + "m \ r\<^sub>0" + "arr_mset arr' l\<^sub>0 (m - 1) = arr_mset arr l\<^sub>0 m - {#Min_mset (arr_mset arr l\<^sub>0 m)#}" + "arr = arr' on - {l\<^sub>0..x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" + "mod_arr = arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m))" + "x\#arr_mset mod_arr l\<^sub>0 (m - 1)" + "y\#arr_mset mod_arr (m - 1) r\<^sub>0" + shows "\ x < y" +proof - + have "{m..0} \ - {l\<^sub>0..0}" + using assms(4) eq_on_sym eq_on_subset by blast + hence arr_eq_on: "mod_arr = arr on {m..0}" + by (simp add: assms(6)) + hence same_mset: "arr_mset mod_arr m r\<^sub>0 = arr_mset arr m r\<^sub>0" + using mset_ran_cong by blast + have "x \# arr_mset arr l\<^sub>0 m" using same_mset assms + by (metis assms(3,6,7) add_mset_remove_trivial_eq lran_upd_outside(2) + mset_lran cancel_ab_semigroup_add_class.diff_right_commute + diff_single_trivial multi_self_add_other_not_self order_refl) + then have x_bigger_min: "x \ Min_mset (arr_mset arr l\<^sub>0 m)" + using Min_le by blast + have y_smaller_min: "y \ Min_mset (arr_mset arr l\<^sub>0 m)" + proof(cases "y = mod_arr (m - 1)") + case False + hence "y\#arr_mset mod_arr (m - 1) r\<^sub>0 - {#mod_arr (m - 1)#}" + by (metis assms(8) diff_single_trivial insert_DiffM insert_noteq_member) + then have "y\#arr_mset arr m r\<^sub>0" + by (simp add: assms(2) intvs_decr_l mset_ran_insert same_mset) + then show ?thesis + using assms(1) assms(5) by fastforce + qed (simp add: assms(6)) + then show ?thesis + using x_bigger_min by linarith +qed + +subsubsection \Implementation\ +text \Now finally the correctness of the \heap_sort\ is shown. As mentioned, it starts by + transforming the array into a minimum heap using @{const "make_heap"}. Then in each + iteration it removes the first element from the heap with @{const "del_min"} after its value + was retrieved with @{const "get_min"}. This value is then put at the position freed by + @{const "del_min"}.\ +program_spec heap_sort + assumes "l \ r" + ensures "array_is_sorted (\) arr l\<^sub>0 r\<^sub>0 \ + arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ + arr\<^sub>0 = arr on - {l\<^sub>0 ..0 } \ l = l\<^sub>0 \ r = r\<^sub>0" + for l r arr[] + defines \ + arr = make_heap(arr, l, r); + m = r; + while (m > l) + @variant \m - l + 1\ + @invariant \is_heap (\) arr l m \ + array_is_sorted (\) arr m r\<^sub>0 \ + (\x \# arr_mset arr l\<^sub>0 m. \y \# arr_mset arr m r\<^sub>0. \ x < y) \ + arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ + l \ m \ m \ r\<^sub>0 \ l = l\<^sub>0 \ arr\<^sub>0 = arr on - {l\<^sub>0 ..0}\ + { + min = get_min(arr, l, m); + (arr, l, m) = del_min(arr, l, m); + arr[m] = min + } + \ + apply vcg_cs + subgoal unfolding array_is_sorted_def by simp + subgoal + apply(intro conjI) + subgoal unfolding is_partial_heap_def by simp + subgoal unfolding array_is_sorted_def using heap_sort_second_half_sorted_step + by blast + subgoal using heap_sort_fst_part_bigger_snd_part_step by blast + subgoal using heap_sort_mset_step by blast + subgoal unfolding eq_on_def + by (metis ComplD ComplI atLeastLessThan_iff less_le_trans) + done + done +end diff --git a/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy b/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy --- a/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy +++ b/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy @@ -1,2017 +1,2017 @@ -(* Title: Irrational_Series_Erdos_Straus.thy - Author: Angeliki Koutsoukou-Argyraki and Wenda Li, University of Cambridge, UK. - -We formalise certain irrationality criteria for infinite series by P. Erdos and E.G. Straus. -In particular, we formalise Theorem 2.1, Corollary 2.10 and Theorem 3.1 in [1]. The latter is an -application of Theorem 2.1 involving the prime numbers. - -References: -[1] P. Erdos and E.G. Straus, On the irrationality of certain series, Pacific Journal of -Mathematics, Vol. 55, No 1, 1974 https://projecteuclid.org/euclid.pjm/1102911140 -*) - -theory "Irrational_Series_Erdos_Straus" imports - Prime_Number_Theorem.Prime_Number_Theorem - Prime_Distribution_Elementary.PNT_Consequences -begin - -section \Miscellaneous\ - -lemma suminf_comparison: - assumes "summable f" and gf: "\n. norm (g n) \ f n" - shows "suminf g \ suminf f" -proof (rule suminf_le) - show "g n \ f n" for n - using gf[of n] by auto - show "summable g" - using assms summable_comparison_test' by blast - show "summable f" using assms(1) . -qed - -lemma tendsto_of_int_diff_0: - assumes "(\n. f n - of_int(g n)) \ (0::real)" "\\<^sub>F n in sequentially. f n > 0" - shows "\\<^sub>F n in sequentially. 0 \ g n" -proof - - have "\\<^sub>F n in sequentially. \f n - of_int(g n)\ < 1 / 2" - using assms(1)[unfolded tendsto_iff,rule_format,of "1/2"] by auto - then show ?thesis using assms(2) - apply eventually_elim - by linarith -qed - -lemma eventually_mono_sequentially: - assumes "eventually P sequentially" - assumes "\x. P (x+k) \ Q (x+k)" - shows "eventually Q sequentially" - using sequentially_offset[OF assms(1),of k] - apply (subst eventually_sequentially_seg[symmetric,of _ k]) - apply (elim eventually_mono) - by fact - -lemma frequently_eventually_at_top: - fixes P Q::"'a::linorder \ bool" - assumes "frequently P at_top" "eventually Q at_top" - shows "frequently (\x. P x \ (\y\x. Q y) ) at_top" - using assms - unfolding frequently_def eventually_at_top_linorder - by (metis (mono_tags, hide_lams) le_cases order_trans) - -lemma eventually_at_top_mono: - fixes P Q::"'a::linorder \ bool" - assumes event_P:"eventually P at_top" - assumes PQ_imp:"\x. x\z \ \y\x. P y \ Q x" - shows "eventually Q at_top" -proof - - obtain N where N_P:"\n\N. P n" - using event_P[unfolded eventually_at_top_linorder] by auto - define N' where "N' = max N z" - have "Q x" when "x\N'" for x - apply (rule PQ_imp) - using N_P that unfolding N'_def by auto - then show ?thesis unfolding eventually_at_top_linorder - by auto -qed - -lemma frequently_at_top_elim: - fixes P Q::"'a::linorder \ bool" - assumes "\\<^sub>Fx in at_top. P x" - assumes "\i. P i \ \j>i. Q j" - shows "\\<^sub>Fx in at_top. Q x" - using assms unfolding frequently_def eventually_at_top_linorder - by (meson leD le_cases less_le_trans) - -lemma less_Liminf_iff: - fixes X :: "_ \ _ :: complete_linorder" - shows "Liminf F X < C \ (\yx. y \ X x) F)" - apply (subst Not_eq_iff[symmetric]) - apply (simp add:not_less not_frequently not_le le_Liminf_iff) - by force - -lemma sequentially_even_odd_imp: - assumes "\\<^sub>F N in sequentially. P (2*N)" "\\<^sub>F N in sequentially. P (2*N+1)" - shows "\\<^sub>F n in sequentially. P n" -proof - - obtain N where N_P:"\x\N. P (2 * x) \ P (2 * x + 1)" - using eventually_conj[OF assms] - unfolding eventually_at_top_linorder by auto - define N' where "N'=2*N " - have "P n" when "n\2*N" for n - proof - - define n' where "n'= n div 2" - then have "n'\N" using that by auto - then have "P (2 * n') \ P (2 * n' + 1)" - using N_P by auto - then show ?thesis unfolding n'_def - apply (cases "even n") - by auto - qed - then show ?thesis unfolding eventually_at_top_linorder by auto -qed - - -section \Theorem 2.1 and Corollary 2.10\ - -context - fixes a b ::"nat\int " - assumes a_pos:"\ n. a n >0 " and a_large:"\\<^sub>F n in sequentially. a n > 1" - and ab_tendsto: "(\n. \b n\ / (a (n-1)*a n)) \ 0" -begin - -private lemma aux_series_summable: "summable (\n. b n / (\k\n. a k))" -proof - - have "\e>0. \\<^sub>F x in sequentially. \b x\ / (a (x-1) * a x) < e" - using ab_tendsto[unfolded tendsto_iff] - apply (simp add:of_int_abs[symmetric] abs_mult del:of_int_abs) - by (subst (asm) (2) abs_of_pos,use \\ n. a n > 0\ in auto)+ - from this[rule_format,of 1] - have "\\<^sub>F x in sequentially. \real_of_int(b x)\ < (a (x-1) * a x)" - using \\ n. a n >0\ by auto - moreover have "\n. (\k\n. real_of_int (a k)) > 0" - using a_pos by (auto intro!:linordered_semidom_class.prod_pos) - ultimately have "\\<^sub>F n in sequentially. \b n\ / (\k\n. a k) - < (a (n-1) * a n) / (\k\n. a k)" - apply (elim eventually_mono) - by (auto simp add:field_simps) - moreover have "\b n\ / (\k\n. a k) = norm (b n / (\k\n. a k))" for n - using \\n. (\k\n. real_of_int (a k)) > 0\[rule_format,of n] by auto - ultimately have "\\<^sub>F n in sequentially. norm (b n / (\k\n. a k)) - < (a (n-1) * a n) / (\k\n. a k)" - by algebra - moreover have "summable (\n. (a (n-1) * a n) / (\k\n. a k))" - proof - - obtain s where a_gt_1:"\ n\s. a n >1" - using a_large[unfolded eventually_at_top_linorder] by auto - define cc where "cc= (\k0" - unfolding cc_def - apply (rule linordered_semidom_class.prod_pos) - using a_pos by auto - have "(\k\n+s. a k) \ cc * 2^n" for n - proof - - have "prod a {s.. 2^n" - proof (induct n) - case 0 - then show ?case using a_gt_1 by auto - next - case (Suc n) - moreover have "a (s + Suc n) \ 2" - using a_gt_1 by (smt le_add1) - ultimately show ?case - apply (subst prod.atLeastLessThan_Suc,simp) - using mult_mono'[of 2 "a (Suc (s + n))" " 2 ^ n" "prod a {s..cc>0\ unfolding cc_def by (simp add: atLeast0AtMost) - qed - then have "1/(\k\n+s. a k) \ 1/(cc * 2^n)" for n - proof - - assume asm:"\n. cc * 2 ^ n \ prod a {..n + s}" - then have "real_of_int (cc * 2 ^ n) \ prod a {..n + s}" using of_int_le_iff by blast - moreover have "prod a {..n + s} >0" using \cc>0\ by (simp add: a_pos prod_pos) - ultimately show ?thesis using \cc>0\ - by (auto simp:field_simps simp del:of_int_prod) - qed - moreover have "summable (\n. 1/(cc * 2^n))" - proof - - have "summable (\n. 1/(2::int)^n)" - using summable_geometric[of "1/(2::int)"] by (simp add:power_one_over) - from summable_mult[OF this,of "1/cc"] show ?thesis by auto - qed - ultimately have "summable (\n. 1 / (\k\n+s. a k))" - apply (elim summable_comparison_test'[where N=0]) - apply (unfold real_norm_def, subst abs_of_pos) - by (auto simp add: \\n. 0 < (\k\n. real_of_int (a k))\) - then have "summable (\n. 1 / (\k\n. a k))" - apply (subst summable_iff_shift[where k=s,symmetric]) - by simp - then have "summable (\n. (a (n+1) * a (n+2)) / (\k\n+2. a k))" - proof - - assume asm:"summable (\n. 1 / real_of_int (prod a {..n}))" - have "1 / real_of_int (prod a {..n}) = (a (n+1) * a (n+2)) / (\k\n+2. a k)" for n - proof - - have "a (Suc (Suc n)) \ 0" "a (Suc n) \0" - using a_pos by (metis less_irrefl)+ - then show ?thesis - by (simp add: atLeast0_atMost_Suc atMost_atLeast0) - qed - then show ?thesis using asm by auto - qed - then show "summable (\n. (a (n-1) * a n) / (\k\n. a k))" - apply (subst summable_iff_shift[symmetric,of _ 2]) - by auto - qed - ultimately show ?thesis - apply (elim summable_comparison_test_ev[rotated]) - by (simp add: eventually_mono) -qed - -private fun get_c::"(nat \ int) \ (nat \ int) \ int \ nat \ (nat \ int)" where - "get_c a' b' B N 0 = round (B * b' N / a' N)"| - "get_c a' b' B N (Suc n) = get_c a' b' B N n * a' (n+N) - B * b' (n+N)" - -lemma ab_rationality_imp: - assumes ab_rational:"(\n. (b n / (\i \ n. a i))) \ \" - shows "\ (B::int)>0. \ c::nat\ int. - (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\ (\n. c (Suc n) / a n) \ 0" -proof - - have [simp]:"a n \ 0" for n using a_pos by (metis less_numeral_extra(3)) - obtain A::int and B::int where - AB_eq:"(\n. real_of_int (b n) / real_of_int (prod a {..n})) = A / B" and "B>0" - proof - - obtain q::rat where "(\n. real_of_int (b n) / real_of_int (prod a {..n})) = real_of_rat q" - using ab_rational by (rule Rats_cases) simp - moreover obtain A::int and B::int where "q = Rat.Fract A B" "B > 0" "coprime A B" - by (rule Rat_cases) auto - ultimately show ?thesis by (auto intro!: that[of A B] simp:of_rat_rat) - qed - define f where "f = (\n. b n / real_of_int (prod a {..n}))" - define R where "R = (\N. (\n. B*b (n+N+1) / prod a {N..n+N+1}))" - have all_e_ubound:"\e>0. \\<^sub>F M in sequentially. \n. \B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" - proof safe - fix e::real assume "e>0" - obtain N where N_a2: "\n \ N. a n \ 2" - and N_ba: "\n \ N. \b n\ / (a (n-1) * a n) < e/(4*B)" - proof - - have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B)" - using order_topology_class.order_tendstoD[OF ab_tendsto,of "e/(4*B)"] \B>0\ \e>0\ - by auto - moreover have "\\<^sub>F n in sequentially. a n \ 2" - using a_large by (auto elim: eventually_mono) - ultimately have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B) \ a n \ 2" - by eventually_elim auto - then show ?thesis unfolding eventually_at_top_linorder using that - by auto - qed - have geq_N_bound:"\B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" when "M\N" for n M - proof - - define D where "D = B*b (n+M+1)/ (a (n+M) * a (n+M+1))" - have "\B*b (n+M+1) / prod a {M..n+M+1}\ = \D / prod a {M.." - proof - - have "{M..n+M+1} = {M.. {n+M,n+M+1}" by auto - then have "prod a {M..n+M+1} = a (n+M) * a (n+M+1)* prod a {M..e/4 * (1/prod a {M.." - proof - - have "\D\ < e/ 4" - unfolding D_def using N_ba[rule_format, of "n+M+1"] \B>0\ \M \ N\ \e>0\ a_pos - by (auto simp:field_simps abs_mult abs_of_pos) - from mult_strict_right_mono[OF this,of "1/prod a {M..e>0\ - show ?thesis - apply (auto simp:abs_prod abs_mult prod_pos) - by (subst (2) abs_of_pos,auto)+ - qed - also have "... \ e/4 * 1/2^n" - proof - - have "prod a {M.. 2^n" - proof (induct n) - case 0 - then show ?case by simp - next - case (Suc n) - then show ?case - using \M\N\ by (simp add: N_a2 mult.commute mult_mono' prod.atLeastLessThan_Suc) - qed - then have "real_of_int (prod a {M.. 2^n" - using numeral_power_le_of_int_cancel_iff by blast - then show ?thesis using \e>0\ by (auto simp add:divide_simps) - qed - finally show ?thesis . - qed - show "\\<^sub>F M in sequentially. \n. \real_of_int (B * b (n + M + 1)) - / real_of_int (prod a {M..n + M + 1})\ < e / 4 * 1 / 2 ^ n" - apply (rule eventually_sequentiallyI[of N]) - using geq_N_bound by blast - qed - have R_tendsto_0:"R \ 0" - proof (rule tendstoI) - fix e::real assume "e>0" - show "\\<^sub>F x in sequentially. dist (R x) 0 < e" using all_e_ubound[rule_format,OF \e>0\] - proof eventually_elim - case (elim M) - define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" - have g_lt:"\g n\ < e/4 * 1/2^n" for n - using elim unfolding g_def by auto - have g_abs_summable:"summable (\n. \g n\)" - proof - - have "summable (\n. e/4 * 1/2^n)" - using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] - by (auto simp add:algebra_simps power_divide) - then show ?thesis - apply (elim summable_comparison_test') - using g_lt less_eq_real_def by auto - qed - have "\\n. g n\ \ (\n. \g n\)" by (rule summable_rabs[OF g_abs_summable]) - also have "... \(\n. e/4 * 1/2^n)" - proof (rule suminf_comparison) - show "summable (\n. e/4 * 1/2^n)" - using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] - by (auto simp add:algebra_simps power_divide) - show "\n. norm \g n\ \ e / 4 * 1 / 2 ^ n" using g_lt less_eq_real_def by auto - qed - also have "... = (e/4) * (\n. (1/2)^n)" - apply (subst suminf_mult[symmetric]) - subgoal - apply (rule complete_algebra_summable_geometric) - by simp - subgoal by (auto simp:algebra_simps power_divide) - done - also have "... = e/2" by (simp add:suminf_geometric[of "1/2"]) - finally have "\\n. g n\ \ e / 2" . - then show "dist (R M) 0 < e" unfolding R_def g_def using \e>0\ by auto - qed - qed - - obtain N where R_N_bound:"\M \ N. \R M\ \ 1 / 4" - and N_geometric:"\M\N. \n. \real_of_int (B * b (n + M + 1)) / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" - proof - - obtain N1 where N1:"\M \ N1. \R M\ \ 1 / 4" - using metric_LIMSEQ_D[OF R_tendsto_0,of "1/4"] all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] - by (auto simp:less_eq_real_def) - obtain N2 where N2:"\M\N2. \n. \real_of_int (B * b (n + M + 1)) - / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" - using all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] - by (auto simp:less_eq_real_def) - define N where "N=max N1 N2" - show ?thesis using that[of N] N1 N2 unfolding N_def by simp - qed - - define C where "C = B * prod a {..nn. f n)" - unfolding AB_eq f_def using \B>0\ by auto - also have "... = B * prod a {..nn. f (n+N+1)))" - using suminf_split_initial_segment[OF \summable f\, of "N+1"] by auto - also have "... = B * prod a {..nn. f (n+N+1)))" - using sum.atLeast0_lessThan_Suc by simp - also have "... = C + B * b N / a N + (\n. B*b (n+N+1) / prod a {N..n+N+1})" - proof - - have "B * prod a {.. {N}" using ivl_disj_un_singleton(2) by blast - then show ?thesis unfolding f_def by auto - qed - moreover have "B * prod a {..n. f (n+N+1)) = (\n. B*b (n+N+1) / prod a {N..n+N+1})" - proof - - have "summable (\n. f (n + N + 1))" - using \summable f\ summable_iff_shift[of f "N+1"] by auto - moreover have "prod a {.. {N..n + N + 1}" by auto - then show ?thesis - unfolding f_def - apply simp - apply (subst prod.union_disjoint) - by auto - qed - ultimately show ?thesis - apply (subst suminf_mult[symmetric]) - by (auto simp add: mult.commute mult.left_commute) - qed - ultimately show ?thesis unfolding C_def by (auto simp:algebra_simps) - qed - also have "... = C +B * b N / a N + R N" - unfolding R_def by simp - finally show ?thesis . - qed - have R_bound:"\R M\ \ 1 / 4" and R_Suc:"R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" - when "M \ N" for M - proof - - define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" - have g_abs_summable:"summable (\n. \g n\)" - proof - - have "summable (\n.(1::real)/2^n)" - using summable_geometric[of "(1::real)/2",simplified] - by (auto elim!: back_subst[of "summable"] simp:field_simps) - moreover have "\g n\ < 1/2^n" for n - using N_geometric[rule_format,OF that] unfolding g_def by simp - ultimately show ?thesis - apply (elim summable_comparison_test') - using less_eq_real_def by auto - qed - show "\R M\ \ 1 / 4" using R_N_bound[rule_format,OF that] . - have "R M = (\n. g n)" unfolding R_def g_def by simp - also have "... = g 0 + (\n. g (Suc n))" - apply (subst suminf_split_head) - using summable_rabs_cancel[OF g_abs_summable] by auto - also have "... = g 0 + 1/a M * (\n. a M * g (Suc n))" - apply (subst suminf_mult) - by (auto simp add: g_abs_summable summable_Suc_iff summable_rabs_cancel) - also have "... = g 0 + 1/a M * R (Suc M)" - proof - - have "a M * g (Suc n) = B * b (n + M + 2) / prod a {Suc M..n + M + 2}" for n - proof - - have "{M..Suc (Suc (M + n))} = {M} \ {Suc M..Suc (Suc (M + n))}" by auto - then show ?thesis - unfolding g_def using \B>0\ by (auto simp add:algebra_simps) - qed - then have "(\n. a M * g (Suc n)) = R (Suc M)" - unfolding R_def by auto - then show ?thesis by auto - qed - finally have "R M = g 0 + 1 / a M * R (Suc M)" . - then have "R (Suc M) = a M * R M - g 0 * a M" - by (auto simp add:algebra_simps) - moreover have "{M..Suc M} = {M,Suc M}" by auto - ultimately show "R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" - unfolding g_def by auto - qed - - define c where "c = (\n. if n\N then get_c a b B N (n-N) else undefined)" - have c_rec:"c (n+1) = c n * a n - B * b n" when "n \ N" for n - unfolding c_def using that by (auto simp:Suc_diff_le) - have c_R:"c (Suc n) / a n = R n" when "n \ N" for n - using that - proof (induct rule:nat_induct_at_least) - case base - have "\ c (N+1) / a N \ \ 1/2" - proof - - have "c N = round (B * b N / a N)" unfolding c_def by simp - moreover have "c (N+1) / a N = c N - B * b N / a N" - using a_pos[rule_format,of N] - by (auto simp add:c_rec[of N,simplified] divide_simps) - ultimately show ?thesis using of_int_round_abs_le by auto - qed - moreover have "\R N\ \ 1 / 4" using R_bound[of N] by simp - ultimately have "\c (N+1) / a N - R N \ < 1" by linarith - moreover have "c (N+1) / a N - R N \ \" - proof - - have "c (N+1) / a N = c N - B * b N / a N" - using a_pos[rule_format,of N] - by (auto simp add:c_rec[of N,simplified] divide_simps) - moreover have " B * b N / a N + R N \ \" - proof - - have "C = B * (\nn {..n}" if "nB>0\ - apply simp - apply (subst prod.union_disjoint) - by auto - qed - finally have "C = real_of_int (B * (\n \" using Ints_of_int by blast - moreover note \A * prod a {.. - ultimately show ?thesis - by (metis Ints_diff Ints_of_int add.assoc add_diff_cancel_left') - qed - ultimately show ?thesis by (simp add: diff_diff_add) - qed - ultimately have "c (N+1) / a N - R N = 0" - by (metis Ints_cases less_irrefl of_int_0 of_int_lessD) - then show ?case by simp - next - case (Suc n) - have "c (Suc (Suc n)) / a (Suc n) = c (Suc n) - B * b (Suc n) / a (Suc n)" - apply (subst c_rec[of "Suc n",simplified]) - using \N \ n\ by (auto simp add: divide_simps) - also have "... = a n * R n - B * b (Suc n) / a (Suc n)" - using Suc by (auto simp: divide_simps) - also have "... = R (Suc n)" - using R_Suc[OF \N \ n\] by simp - finally show ?case . - qed - have ca_tendsto_zero:"(\n. c (Suc n) / a n) \ 0" - using R_tendsto_0 - apply (elim filterlim_mono_eventually) - using c_R by (auto intro!:eventually_sequentiallyI[of N]) - have ca_bound:"\c (n + 1)\ < a n / 2" when "n \ N" for n - proof - - have "\c (Suc n)\ / a n = \c (Suc n) / a n\" using a_pos[rule_format,of n] by auto - also have "... = \R n\" using c_R[OF that] by auto - also have "... < 1/2" using R_bound[OF that] by auto - finally have "\c (Suc n)\ / a n < 1/2" . - then show ?thesis using a_pos[rule_format,of n] by auto - qed - -(* (* the following part corresponds to (2.7) (2.8) in the original paper, but turns out to be - not necessary. *) - have c_round:"c n = round (B * b n / a n)" when "n \ N" for n - proof (cases "n=N") - case True - then show ?thesis unfolding c_def by simp - next - case False - with \n\N\ obtain n' where n_Suc:"n=Suc n'" and "n' \ N" - by (metis le_eq_less_or_eq lessE less_imp_le_nat) - have "B * b n / a n = c n - R n" - proof - - have "R n = c n - B * b n / a n" - using c_R[OF \n'\N\,symmetric,folded n_Suc] R_Suc[OF \n'\N\,folded n_Suc] - by (auto simp:field_simps) - then show ?thesis by (auto simp:field_simps) - qed - then have "\B * b n / a n - c n\ = \R n\" by auto - then have "\B * b n / a n - c n\ < 1/2" using R_bound[OF \n \ N\] by auto - from round_unique'[OF this] show ?thesis by auto - qed - *) - show "\B>0. \c. (\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) - \ real_of_int \c (n + 1)\ < a n / 2) \ (\n. c (Suc n) / a n) \ 0" - unfolding eventually_at_top_linorder - apply (rule exI[of _ B],use \B>0\ in simp) - apply (intro exI[of _c] exI[of _ N]) - using c_rec ca_bound ca_tendsto_zero - by fastforce -qed - -private lemma imp_ab_rational: - assumes "\ (B::int)>0. \ c::nat\ int. - (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\n. (b n / (\i \ n. a i))) \ \" -proof - - obtain B::int and c::"nat\int" and N::nat where "B>0" and - large_n:"\n\N. B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < a n / 2 - \ a n\2" - proof - - obtain B c where "B>0" and event1:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) - \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2" - using assms by auto - from eventually_conj[OF event1 a_large,unfolded eventually_at_top_linorder] - obtain N where "\n\N. (B * b n = c n * a n - c (n + 1) - \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ 2 \ a n" - by fastforce - then show ?thesis using that[of B N c] \B>0\ by auto - qed - define f where "f=(\n. real_of_int (b n) / real_of_int (prod a {..n}))" - define S where "S = (\n. f n)" - have "summable f" - unfolding f_def by (rule aux_series_summable) - define C where "C=B*prod a {..nn. (c (n+N) * a (n+N)) / prod a {N..n+N})" - define h2 where "h2 = (\n. c (n+N+1) / prod a {N..n+N})" - have f_h12:"B*prod a {..n. B * b (n+N))" - define g2 where "g2 = (\n. prod a {.. {N..n + N}) = prod a {.. {N..n + N}) = prod a {..n+N}" - by (simp add: ivl_disj_un_one(4)) - ultimately show ?thesis - unfolding g2_def - apply simp - using a_pos by (metis less_irrefl) - qed - ultimately have "B*prod a {..nn. f (n+N)))" - using suminf_split_initial_segment[OF \summable f\,of N] - unfolding S_def by (auto simp:algebra_simps) - also have "... = C + B*prod a {..n. f (n+N))" - unfolding C_def by (auto simp:algebra_simps) - also have "... = C + (\n. h1 n - h2 n)" - apply (subst suminf_mult[symmetric]) - subgoal using \summable f\ by (simp add: summable_iff_shift) - subgoal using f_h12 by auto - done - also have "... = C + h1 0" - proof - - have "(\n. \i\n. h1 i - h2 i) \ (\i. h1 i - h2 i)" - proof (rule summable_LIMSEQ') - have "(\i. h1 i - h2 i) = (\i. real_of_int (B * prod a {..i. h1 i - h2 i)" - using \summable f\ by (simp add: summable_iff_shift summable_mult) - qed - moreover have "(\i\n. h1 i - h2 i) = h1 0 - h2 n" for n - proof (induct n) - case 0 - then show ?case by simp - next - case (Suc n) - have "(\i\Suc n. h1 i - h2 i) = (\i\n. h1 i - h2 i) + h1 (n+1) - h2 (n+1)" - by auto - also have "... = h1 0 - h2 n + h1 (n+1) - h2 (n+1)" using Suc by auto - also have "... = h1 0 - h2 (n+1)" - proof - - have "h2 n = h1 (n+1)" - unfolding h2_def h1_def - apply (auto simp:prod.nat_ivl_Suc') - using a_pos by (metis less_numeral_extra(3)) - then show ?thesis by auto - qed - finally show ?case by simp - qed - ultimately have "(\n. h1 0 - h2 n) \ (\i. h1 i - h2 i)" by simp - then have "h2 \ (h1 0 - (\i. h1 i - h2 i))" - apply (elim metric_tendsto_imp_tendsto) - by (auto intro!:eventuallyI simp add:dist_real_def) - moreover have "h2 \ 0" - proof - - have h2_n:"\h2 n\ < (1 / 2)^(n+1)" for n - proof - - have "\h2 n\ = \c (n + N + 1)\ / prod a {N..n + N}" - unfolding h2_def abs_divide - using a_pos by (simp add: abs_of_pos prod_pos) - also have "... < (a (N+n) / 2) / prod a {N..n + N}" - unfolding h2_def - apply (rule divide_strict_right_mono) - subgoal using large_n[rule_format,of "N+n"] by (auto simp add:algebra_simps) - subgoal using a_pos by (simp add: prod_pos) - done - also have "... = 1 / (2*prod a {N.. (1/2)^(n+1)" - proof (induct n) - case 0 - then show ?case by auto - next - case (Suc n) - define P where "P=1 / real_of_int (2 * prod a {N.. ( (1 / 2) ^ (n + 1) ) / a (n+N) " - apply (rule divide_right_mono) - subgoal unfolding P_def using Suc by auto - subgoal by (simp add: a_pos less_imp_le) - done - also have "... \ ( (1 / 2) ^ (n + 1) ) / 2 " - apply (rule divide_left_mono) - using large_n[rule_format,of "n+N",simplified] by auto - also have "... = (1 / 2) ^ (n + 2)" by auto - finally show ?case by simp - qed - finally show ?thesis . - qed - have "(\n. (1 / 2)^(n+1)) \ (0::real)" - using tendsto_mult_right_zero[OF LIMSEQ_abs_realpow_zero2[of "1/2",simplified],of "1/2"] - by auto - then show ?thesis - apply (elim Lim_null_comparison[rotated]) - using h2_n less_eq_real_def by (auto intro!:eventuallyI) - qed - ultimately have "(\i. h1 i - h2 i) = h1 0" - using LIMSEQ_unique by fastforce - then show ?thesis by simp - qed - also have "... = C + c N" - unfolding h1_def using a_pos - by auto (metis less_irrefl) - finally show ?thesis . - qed - then have "S = (C + real_of_int (c N)) / (B*prod a {..0 < B\ a_pos less_irrefl mult.commute mult_pos_pos - nonzero_mult_div_cancel_right of_int_eq_0_iff prod_pos) - moreover have "... \ \" - unfolding C_def f_def by (intro Rats_divide Rats_add Rats_mult Rats_of_int Rats_sum) - ultimately show "S \ \" by auto -qed - -theorem theorem_2_1_Erdos_Straus : - "(\n. (b n / (\i \ n. a i))) \ \ \ (\ (B::int)>0. \ c::nat\ int. - (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\The following is a Corollary to Theorem 2.1. \ - -corollary corollary_2_10_Erdos_Straus: - assumes ab_event:"\\<^sub>F n in sequentially. b n > 0 \ a (n+1) \ a n" - and ba_lim_leq:"lim (\n. (b(n+1) - b n )/a n) \ 0" - and ba_lim_exist:"convergent (\n. (b(n+1) - b n )/a n)" - and "liminf (\n. a n / b n) = 0 " - shows "(\n. (b n / (\i \ n. a i))) \ \" -proof - assume "(\n. (b n / (\i \ n. a i))) \ \" - then obtain B c where "B>0" and abc_event:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) - \ \c (n + 1)\ < a n / 2" and ca_vanish: "(\n. c (Suc n) / a n) \ 0" - using ab_rationality_imp by auto - - have bac_close:"(\n. B * b n / a n - c n) \ 0" - proof - - have "\\<^sub>F n in sequentially. B * b n - c n * a n + c (n + 1) = 0" - using abc_event by (auto elim!:eventually_mono) - then have "\\<^sub>F n in sequentially. (B * b n - c n * a n + c (n+1)) / a n = 0 " - apply eventually_elim - by auto - then have "\\<^sub>F n in sequentially. B * b n / a n - c n + c (n + 1) / a n = 0" - apply eventually_elim - using a_pos by (auto simp:divide_simps) (metis less_irrefl) - then have "(\n. B * b n / a n - c n + c (n + 1) / a n) \ 0" - by (simp add: eventually_mono tendsto_iff) - from tendsto_diff[OF this ca_vanish] - show ?thesis by auto - qed - - have c_pos:"\\<^sub>F n in sequentially. c n > 0" - proof - - from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" - apply (elim tendsto_of_int_diff_0) - using ab_event a_large apply (eventually_elim) - using \B>0\ by auto - show ?thesis - proof (rule ccontr) - assume "\ (\\<^sub>F n in sequentially. c n > 0)" - moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" - using * eventually_sequentially_Suc[of "\n. c n\0"] - by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) - ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" - using eventually_elim2 frequently_def by fastforce - moreover have "\\<^sub>F n in sequentially. b n > 0 \ B * b n = c n * a n - c (n + 1)" - using ab_event abc_event by eventually_elim auto - ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 \ b n > 0 - \ B * b n = c n * a n - c (n + 1)" - using frequently_eventually_frequently by fastforce - from frequently_ex[OF this] - obtain n where "c n = 0" "c (Suc n) \ 0" "b n > 0" - "B * b n = c n * a n - c (n + 1)" - by auto - then have "B * b n \ 0" by auto - then show False using \b n>0\ \B > 0\ using mult_pos_pos not_le by blast - qed - qed - - have bc_epsilon:"\\<^sub>F n in sequentially. b (n+1) / b n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real - proof - - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" - using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto - moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" - apply (subst (asm) eventually_sequentially_Suc[symmetric]) - by simp - moreover have "\\<^sub>F n in sequentially. B * b (n+1) = c (n+1) * a (n+1) - c (n + 2)" - using abc_event - apply (subst (asm) eventually_sequentially_Suc[symmetric]) - by (auto elim:eventually_mono) - moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" - proof - - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" - using c_pos by (subst eventually_sequentially_Suc) simp - moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" - using c_pos by (subst eventually_sequentially_Suc) simp - ultimately show ?thesis using c_pos by eventually_elim auto - qed - ultimately show ?thesis using ab_event abc_event - proof eventually_elim - case (elim n) - define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" - have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" using a_pos elim by (auto simp add: \\<^sub>0_def \\<^sub>1_def) - have "(\ - \\<^sub>1) * c n > 0" - apply (rule mult_pos_pos) - using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto - moreover have "\\<^sub>0 * (c (n+1) - \) > 0" - apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) - using elim(4) that(2) by linarith - ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto - moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith - moreover have "c n > 0" by (simp add: elim(4)) - ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" - by (auto simp add: field_simps) - also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" - proof - - have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" - by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ - divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) - moreover have "(a (n+1) / a n) \ 1" - using a_pos elim(5) by auto - ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) - qed - also have "... = (B * b (n+1)) / (B * b n)" - proof - - have "B * b n = c n * a n - c (n + 1)" - using elim by auto - also have "... = a n * (c n - \\<^sub>0)" - using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) - finally have "B * b n = a n * (c n - \\<^sub>0)" . - moreover have "B * b (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" - unfolding \\<^sub>1_def - using a_pos[rule_format,of "n+1"] - apply (subst \B * b (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) - by (auto simp:field_simps) - ultimately show ?thesis by (simp add: mult.commute) - qed - also have "... = b (n+1) / b n" - using \B>0\ by auto - finally show ?case . - qed - qed - - have eq_2_11:"\\<^sub>F n in sequentially. b (n+1) > b n + (1 - \)^2 * a n / B" - when "\>0" "\<1" "\ (\\<^sub>F n in sequentially. c (n+1) \ c n)" for \::real - proof - - have "\\<^sub>F x in sequentially. c x < c (Suc x) " using that(3) - by (simp add:not_eventually frequently_elim1) - moreover have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" - using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto - moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" - proof - - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" - using c_pos by (subst eventually_sequentially_Suc) simp - then show ?thesis using c_pos by eventually_elim auto - qed - ultimately show ?thesis using ab_event abc_event bc_epsilon[OF \\>0\ \\<1\] - proof (elim frequently_rev_mp,eventually_elim) - case (elim n) - then have "c (n+1) / a n < \" - using a_pos[rule_format,of n] by auto - also have "... \ \ * c n" using elim(7) that(1) by auto - finally have "c (n+1) / a n < \ * c n" . - then have "c (n+1) / c n < \ * a n" - using a_pos[rule_format,of n] elim by (auto simp:field_simps) - then have "(1 - \) * a n < a n - c (n+1) / c n" - by (auto simp:algebra_simps) - then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" - apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) - using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) - then have "b n + (1 - \)^2 * a n / B < b n + (1 - \) * (a n - c (n+1) / c n) / B" - using \B>0\ by auto - also have "... = b n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" - using elim by (auto simp:field_simps) - also have "... = b n + (1 - \) * (b n / c n)" - proof - - have "B * b n = c n * a n - c (n + 1)" using elim by auto - from this[symmetric] show ?thesis - using \B>0\ by simp - qed - also have "... = (1+(1-\)/c n) * b n" - by (auto simp:algebra_simps) - also have "... = ((c n+1-\)/c n) * b n" - using elim by (auto simp:divide_simps) - also have "... \ ((c (n+1) -\)/c n) * b n" - proof - - define cp where "cp = c n+1" - have "c (n+1) \ cp" unfolding cp_def using \c n < c (Suc n)\ by auto - moreover have "c n>0" "b n>0" using elim by auto - ultimately show ?thesis - apply (fold cp_def) - by (auto simp:divide_simps) - qed - also have "... < b (n+1)" - using elim by (auto simp:divide_simps) - finally show ?case . - qed - qed - - have "\\<^sub>F n in sequentially. c (n+1) \ c n" - proof (rule ccontr) - assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" - from eq_2_11[OF _ _ this,of "1/2"] - have "\\<^sub>F n in sequentially. b (n+1) > b n + 1/4 * a n / B" - by (auto simp:algebra_simps power2_eq_square) - then have *:"\\<^sub>F n in sequentially. (b (n+1) - b n) / a n > 1 / (B * 4)" - apply (elim frequently_elim1) - subgoal for n - using a_pos[rule_format,of n] by (auto simp:field_simps) - done - define f where "f = (\n. (b (n+1) - b n) / a n)" - have "f \ lim f" - using convergent_LIMSEQ_iff ba_lim_exist unfolding f_def by auto - from this[unfolded tendsto_iff,rule_format, of "1 / (B*4)"] - have "\\<^sub>F x in sequentially. \f x - lim f\ < 1 / (B * 4)" - using \B>0\ by (auto simp:dist_real_def) - moreover have "\\<^sub>F n in sequentially. f n > 1 / (B * 4)" - using * unfolding f_def by auto - ultimately have "\\<^sub>F n in sequentially. f n > 1 / (B * 4) \ \f n - lim f\ < 1 / (B * 4)" - by (auto elim:frequently_eventually_frequently[rotated]) - from frequently_ex[OF this] - obtain n where "f n > 1 / (B * 4)" "\f n - lim f\ < 1 / (B * 4)" - by auto - moreover have "lim f \ 0" using ba_lim_leq unfolding f_def by auto - ultimately show False by linarith - qed - then obtain N where N_dec:"\n\N. c (n+1) \ c n" by (meson eventually_at_top_linorder) - define max_c where "max_c = (MAX n \ {..N}. c n)" - have max_c:"c n \ max_c" for n - proof (cases "n\N") - case True - then show ?thesis unfolding max_c_def by simp - next - case False - then have "n\N" by auto - then have "c n\c N" - proof (induct rule:nat_induct_at_least) - case base - then show ?case by simp - next - case (Suc n) - then have "c (n+1) \ c n" using N_dec by auto - then show ?case using \c n \ c N\ by auto - qed - moreover have "c N \ max_c" unfolding max_c_def by auto - ultimately show ?thesis by auto - qed - have "max_c > 0 " - proof - - obtain N where "\n\N. 0 < c n" - using c_pos[unfolded eventually_at_top_linorder] by auto - then have "c N > 0" by auto - then show ?thesis using max_c[of N] by simp - qed - have ba_limsup_bound:"1/(B*(B+1)) \ limsup (\n. b n/a n)" - "limsup (\n. b n/a n) \ max_c / B + 1 / (B+1)" - proof - - define f where "f = (\n. b n/a n)" - from tendsto_mult_right_zero[OF bac_close,of "1/B"] - have "(\n. f n - c n / B) \ 0" - unfolding f_def using \B>0\ by (auto simp:algebra_simps) - from this[unfolded tendsto_iff,rule_format,of "1/(B+1)"] - have "\\<^sub>F x in sequentially. \f x - c x / B\ < 1 / (B+1)" - using \B>0\ by auto - then have *:"\\<^sub>F n in sequentially. 1/(B*(B+1)) \ ereal (f n) \ ereal (f n) \ max_c / B + 1 / (B+1)" - using c_pos - proof eventually_elim - case (elim n) - then have "f n - c n / B < 1 / (B+1)" by auto - then have "f n < c n / B + 1 / (B+1)" by simp - also have "... \ max_c / B + 1 / (B+1)" - using max_c[of n] using \B>0\ by (auto simp:divide_simps) - finally have *:"f n < max_c / B + 1 / (B+1)" . - - have "1/(B*(B+1)) = 1/B - 1 / (B+1)" - using \B>0\ by (auto simp:divide_simps) - also have "... \ c n/B - 1 / (B+1)" - using \0 < c n\ \B>0\ by (auto,auto simp:divide_simps) - also have "... < f n" using elim by auto - finally have "1/(B*(B+1)) < f n" . - with * show ?case by simp - qed - show "limsup f \ max_c / B + 1 / (B+1)" - apply (rule Limsup_bounded) - using * by (auto elim:eventually_mono) - have "1/(B*(B+1)) \ liminf f" - apply (rule Liminf_bounded) - using * by (auto elim:eventually_mono) - also have "liminf f \ limsup f" by (simp add: Liminf_le_Limsup) - finally show "1/(B*(B+1)) \ limsup f" . - qed - - have "0 < inverse (ereal (max_c / B + 1 / (B+1)))" - using \max_c > 0\ \B>0\ - by (simp add: pos_add_strict) - also have "... \ inverse (limsup (\n. b n/a n))" - proof (rule ereal_inverse_antimono[OF _ ba_limsup_bound(2)]) - have "0<1/(B*(B+1))" using \B>0\ by auto - also have "... \ limsup (\n. b n/a n)" using ba_limsup_bound(1) . - finally show "0\limsup (\n. b n/a n)" using zero_ereal_def by auto - qed - also have "... = liminf (\n. inverse (ereal ( b n/a n)))" - apply (subst Liminf_inverse_ereal[symmetric]) - using a_pos ab_event by (auto elim!:eventually_mono simp:divide_simps) - also have "... = liminf (\n. ( a n/b n))" - apply (rule Liminf_eq) - using a_pos ab_event - apply (auto elim!:eventually_mono) - by (metis less_int_code(1)) - finally have "liminf (\n. ( a n/b n)) > 0" . - then show False using \liminf (\n. a n / b n) = 0\ by simp -qed - -end - -section\Some auxiliary results on the prime numbers. \ - -lemma nth_prime_nonzero[simp]:"nth_prime n \ 0" - by (simp add: prime_gt_0_nat prime_nth_prime) - -lemma nth_prime_gt_zero[simp]:"nth_prime n >0" - by (simp add: prime_gt_0_nat prime_nth_prime) - -lemma ratio_of_consecutive_primes: - "(\n. nth_prime (n+1)/nth_prime n) \1" -proof - - define f where "f=(\x. real (nth_prime (Suc x)) /real (nth_prime x))" - define g where "g=(\x. (real x * ln (real x)) - / (real (Suc x) * ln (real (Suc x))))" - have p_n:"(\x. real (nth_prime x) / (real x * ln (real x))) \ 1" - using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified] . - moreover have p_sn:"(\n. real (nth_prime (Suc n)) - / (real (Suc n) * ln (real (Suc n)))) \ 1" - using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified - ,THEN LIMSEQ_Suc] . - ultimately have "(\x. f x * g x) \ 1" - using tendsto_divide[OF p_sn p_n] - unfolding f_def g_def by (auto simp:algebra_simps) - moreover have "g \ 1" unfolding g_def - by real_asymp - ultimately have "(\x. if g x = 0 then 0 else f x) \ 1" - apply (drule_tac tendsto_divide[OF _ \g \ 1\]) - by auto - then have "f \ 1" - proof (elim filterlim_mono_eventually) - have "\\<^sub>F x in sequentially. (if g (x+3) = 0 then 0 - else f (x+3)) = f (x+3)" - unfolding g_def by auto - then show "\\<^sub>F x in sequentially. (if g x = 0 then 0 else f x) = f x" - apply (subst (asm) eventually_sequentially_seg) - by simp - qed auto - then show ?thesis unfolding f_def by auto -qed - -lemma nth_prime_double_sqrt_less: - assumes "\ > 0" - shows "\\<^sub>F n in sequentially. (nth_prime (2*n) - nth_prime n) - / sqrt (nth_prime n) < n powr (1/2+\)" -proof - - define pp ll where - "pp=(\n. (nth_prime (2*n) - nth_prime n) / sqrt (nth_prime n))" and - "ll=(\x::nat. x * ln x)" - have pp_pos:"pp (n+1) > 0" for n - unfolding pp_def by simp - - have "(\x. nth_prime (2 * x)) \[sequentially] (\x. (2 * x) * ln (2 * x))" - using nth_prime_asymptotics[THEN asymp_equiv_compose - ,of "(*) 2" sequentially,unfolded comp_def] - using mult_nat_left_at_top pos2 by blast - also have "... \[sequentially] (\x. 2 *x * ln x)" - by real_asymp - finally have "(\x. nth_prime (2 * x)) \[sequentially] (\x. 2 *x * ln x)" . - from this[unfolded asymp_equiv_def, THEN tendsto_mult_left,of 2] - have "(\x. nth_prime (2 * x) / (x * ln x)) \ 2" - unfolding asymp_equiv_def by auto - moreover have *:"(\x. nth_prime x / (x * ln x)) \ 1" - using nth_prime_asymptotics unfolding asymp_equiv_def by auto - ultimately - have "(\x. (nth_prime (2 * x) - nth_prime x) / ll x) \ 1" - unfolding ll_def - apply - - apply (drule (1) tendsto_diff) - apply (subst of_nat_diff,simp) - by (subst diff_divide_distrib,simp) - moreover have "(\x. sqrt (nth_prime x) / sqrt (ll x)) \ 1" - unfolding ll_def - using tendsto_real_sqrt[OF *] - by (auto simp: real_sqrt_divide) - ultimately have "(\x. pp x * (sqrt (ll x) / (ll x))) \ 1" - apply - - apply (drule (1) tendsto_divide,simp) - by (auto simp:field_simps of_nat_diff pp_def) - moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / ll x = 1/sqrt (ll x)" - apply (subst eventually_sequentially_Suc[symmetric]) - by (auto intro!:eventuallyI simp:ll_def divide_simps) - ultimately have "(\x. pp x / sqrt (ll x)) \ 1" - apply (elim filterlim_mono_eventually) - by (auto elim!:eventually_mono) (metis mult.right_neutral times_divide_eq_right) - moreover have "(\x. sqrt (ll x) / x powr (1/2+\)) \ 0" - unfolding ll_def using \\>0\ by real_asymp - ultimately have "(\x. pp x / x powr (1/2+\) * - (sqrt (ll x) / sqrt (ll x))) \ 0" - apply - - apply (drule (1) tendsto_mult) - by (auto elim:filterlim_mono_eventually) - moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / sqrt (ll x) = 1" - apply (subst eventually_sequentially_Suc[symmetric]) - by (auto intro!:eventuallyI simp:ll_def ) - ultimately have "(\x. pp x / x powr (1/2+\)) \ 0" - apply (elim filterlim_mono_eventually) - by (auto elim:eventually_mono) - from tendstoD[OF this, of 1,simplified] - show "\\<^sub>F x in sequentially. pp x < x powr (1 / 2 + \)" - apply (elim eventually_mono_sequentially[of _ 1]) - using pp_pos by auto -qed - - -section \Theorem 3.1\ - -text\Theorem 3.1 is an application of Theorem 2.1 with the sequences considered involving -the prime numbers.\ - -theorem theorem_3_10_Erdos_Straus: - fixes a::"nat \ int" - assumes a_pos:"\ n. a n >0" and "mono a" - and nth_1:"(\n. nth_prime n / (a n)^2) \ 0" - and nth_2:"liminf (\n. a n / nth_prime n) = 0" - shows "(\n. (nth_prime n / (\i \ n. a i))) \ \" -proof - assume asm:"(\n. (nth_prime n / (\i \ n. a i))) \ \" - - have a2_omega:"(\n. (a n)^2) \ \(\x. x * ln x)" - proof - - have "(\n. real (nth_prime n)) \ o(\n. real_of_int ((a n)\<^sup>2))" - apply (rule smalloI_tendsto[OF nth_1]) - using a_pos by (metis (mono_tags, lifting) less_int_code(1) - not_eventuallyD of_int_0_eq_iff zero_eq_power2) - moreover have "(\x. real (nth_prime x)) \ \(\x. real x * ln (real x))" - using nth_prime_bigtheta - by blast - ultimately show ?thesis - using landau_omega.small_big_trans smallo_imp_smallomega by blast - qed - - have a_gt_1:"\\<^sub>F n in sequentially. 1 < a n" - proof - - have "\\<^sub>F x in sequentially. \x * ln x\ \ (a x)\<^sup>2" - using a2_omega[unfolded smallomega_def,simplified,rule_format,of 1] - by auto - then have "\\<^sub>F x in sequentially. \(x+3) * ln (x+3)\ \ (a (x+3))\<^sup>2" - apply (subst (asm) eventually_sequentially_seg[symmetric, of _ 3]) - by simp - then have "\\<^sub>F n in sequentially. 1 < a ( n+3)" - proof (elim eventually_mono) - fix x - assume "\real (x + 3) * ln (real (x + 3))\ \ real_of_int ((a (x + 3))\<^sup>2)" - moreover have "\real (x + 3) * ln (real (x + 3))\ > 3" - proof - - have "ln (real (x + 3)) > 1" - apply simp using ln3_gt_1 ln_gt_1 by force - moreover have "real(x+3) \ 3" by simp - ultimately have "(x+3)*ln (real (x + 3)) > 3*1 " - apply (rule_tac mult_le_less_imp_less) - by auto - then show ?thesis by auto - qed - ultimately have "real_of_int ((a (x + 3))\<^sup>2) > 3" - by auto - then show "1 < a (x + 3)" - by (smt Suc3_eq_add_3 a_pos add.commute of_int_1 one_power2) - qed - then show ?thesis - apply (subst eventually_sequentially_seg[symmetric, of _ 3]) - by auto - qed - - obtain B::int and c where - "B>0" and Bc_large:"\\<^sub>F n in sequentially. B * nth_prime n - = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" - and ca_vanish: "(\n. c (Suc n) / real_of_int (a n)) \ 0" - proof - - note a_gt_1 - moreover have "(\n. real_of_int \int (nth_prime n)\ - / real_of_int (a (n - 1) * a n)) \ 0" - proof - - define f where "f=(\n. nth_prime (n+1) / (a n * a (n+1)))" - define g where "g=(\n. 2*nth_prime n / (a n)^2)" - have "\\<^sub>F x in sequentially. norm (f x) \ g x" - proof - - have "\\<^sub>F n in sequentially. nth_prime (n+1) < 2*nth_prime n" - using ratio_of_consecutive_primes[unfolded tendsto_iff - ,rule_format,of 1,simplified] - apply (elim eventually_mono) - by (auto simp :divide_simps dist_norm) - moreover have "\\<^sub>F n in sequentially. real_of_int (a n * a (n+1)) - \ (a n)^2" - apply (rule eventuallyI) - using \mono a\ by (auto simp:power2_eq_square a_pos incseq_SucD) - ultimately show ?thesis unfolding f_def g_def - apply eventually_elim - apply (subst norm_divide) - apply (rule_tac linordered_field_class.frac_le) - using a_pos[rule_format, THEN order.strict_implies_not_eq ] - by auto - qed - moreover have "g \ 0 " - using nth_1[THEN tendsto_mult_right_zero,of 2] unfolding g_def - by auto - ultimately have "f \ 0" - using Lim_null_comparison[of f g sequentially] - by auto - then show ?thesis - unfolding f_def - by (rule_tac LIMSEQ_imp_Suc) auto - qed - moreover have "(\n. real_of_int (int (nth_prime n)) - / real_of_int (prod a {..n})) \ \" - using asm by simp - ultimately have "\B>0. \c. (\\<^sub>F n in sequentially. - B * int (nth_prime n) = c n * a n - c (n + 1) \ - real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ - (\n. real_of_int (c (Suc n)) / real_of_int (a n)) \ 0" - using ab_rationality_imp[OF a_pos,of nth_prime] by fast - then show thesis - apply clarify - apply (rule_tac c=c and B=B in that) - by auto - qed - - have bac_close:"(\n. B * nth_prime n / a n - c n) \ 0" - proof - - have "\\<^sub>F n in sequentially. B * nth_prime n - c n * a n + c (n + 1) = 0" - using Bc_large by (auto elim!:eventually_mono) - then have "\\<^sub>F n in sequentially. (B * nth_prime n - c n * a n + c (n+1)) / a n = 0 " - apply eventually_elim - by auto - then have "\\<^sub>F n in sequentially. B * nth_prime n / a n - c n + c (n + 1) / a n = 0" - apply eventually_elim - using a_pos by (auto simp:divide_simps) (metis less_irrefl) - then have "(\n. B * nth_prime n / a n - c n + c (n + 1) / a n) \ 0" - by (simp add: eventually_mono tendsto_iff) - from tendsto_diff[OF this ca_vanish] - show ?thesis by auto - qed - - have c_pos:"\\<^sub>F n in sequentially. c n > 0" - proof - - from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" - apply (elim tendsto_of_int_diff_0) - using a_gt_1 apply (eventually_elim) - using \B>0\ by auto - show ?thesis - proof (rule ccontr) - assume "\ (\\<^sub>F n in sequentially. c n > 0)" - moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" - using * eventually_sequentially_Suc[of "\n. c n\0"] - by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) - ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" - using eventually_elim2 frequently_def by fastforce - moreover have "\\<^sub>F n in sequentially. nth_prime n > 0 - \ B * nth_prime n = c n * a n - c (n + 1)" - using Bc_large by eventually_elim auto - ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 - \ B * nth_prime n = c n * a n - c (n + 1)" - using frequently_eventually_frequently by fastforce - from frequently_ex[OF this] - obtain n where "c n = 0" "c (Suc n) \ 0" - "B * nth_prime n = c n * a n - c (n + 1)" - by auto - then have "B * nth_prime n \ 0" by auto - then show False using \B > 0\ - by (simp add: mult_le_0_iff) - qed - qed - - have B_nth_prime:"\\<^sub>F n in sequentially. nth_prime n > B" - proof - - have "\\<^sub>F x in sequentially. B+1 \ nth_prime x" - using nth_prime_at_top[unfolded filterlim_at_top_ge[where c="nat B+1"] - ,rule_format,of "nat B + 1",simplified] - - apply (elim eventually_mono) - using \B>0\ by auto - then show ?thesis - apply (elim eventually_mono) - by auto - qed - - have bc_epsilon:"\\<^sub>F n in sequentially. nth_prime (n+1) - / nth_prime n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real - proof - - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" - using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto - moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" - apply (subst (asm) eventually_sequentially_Suc[symmetric]) - by simp - moreover have "\\<^sub>F n in sequentially. B * nth_prime (n+1) = c (n+1) * a (n+1) - c (n + 2)" - using Bc_large - apply (subst (asm) eventually_sequentially_Suc[symmetric]) - by (auto elim:eventually_mono) - moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" - proof - - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" - using c_pos by (subst eventually_sequentially_Suc) simp - moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" - using c_pos by (subst eventually_sequentially_Suc) simp - ultimately show ?thesis using c_pos by eventually_elim auto - qed - ultimately show ?thesis using Bc_large - proof eventually_elim - case (elim n) - define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" - have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" - using a_pos elim \mono a\ - by (auto simp add: \\<^sub>0_def \\<^sub>1_def abs_of_pos) - have "(\ - \\<^sub>1) * c n > 0" - apply (rule mult_pos_pos) - using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto - moreover have "\\<^sub>0 * (c (n+1) - \) > 0" - apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) - using elim(4) that(2) by linarith - ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto - moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith - moreover have "c n > 0" by (simp add: elim(4)) - ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" - by (auto simp add:field_simps) - also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" - proof - - have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" - by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ - divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) - moreover have "(a (n+1) / a n) \ 1" - using a_pos \mono a\ by (simp add: mono_def) - ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) - qed - also have "... = (B * nth_prime (n+1)) / (B * nth_prime n)" - proof - - have "B * nth_prime n = c n * a n - c (n + 1)" - using elim by auto - also have "... = a n * (c n - \\<^sub>0)" - using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) - finally have "B * nth_prime n = a n * (c n - \\<^sub>0)" . - moreover have "B * nth_prime (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" - unfolding \\<^sub>1_def - using a_pos[rule_format,of "n+1"] - apply (subst \B * nth_prime (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) - by (auto simp:field_simps) - ultimately show ?thesis by (simp add: mult.commute) - qed - also have "... = nth_prime (n+1) / nth_prime n" - using \B>0\ by auto - finally show ?case . - qed - qed - - - have c_ubound:"\x. \n. c n > x" - proof (rule ccontr) - assume " \ (\x. \n. x < c n)" - then obtain ub where "\n. c n \ ub" "ub > 0" - by (meson dual_order.trans int_one_le_iff_zero_less le_cases not_le) - define pa where "pa = (\n. nth_prime n / a n)" - have pa_pos:"\n. pa n > 0" unfolding pa_def by (simp add: a_pos) - have "liminf (\n. 1 / pa n) = 0" - using nth_2 unfolding pa_def by auto - then have "(\y\<^sub>F x in sequentially. ereal (1 / pa x) \ y)" - apply (subst less_Liminf_iff[symmetric]) - using \0 < B\ \0 < ub\ by auto - then have "\\<^sub>F x in sequentially. 1 / pa x < B/(ub+1)" - by (meson frequently_mono le_less_trans less_ereal.simps(1)) - then have "\\<^sub>F x in sequentially. B*pa x > (ub+1)" - apply (elim frequently_elim1) - by (metis \0 < ub\ mult.left_neutral of_int_0_less_iff pa_pos pos_divide_less_eq - pos_less_divide_eq times_divide_eq_left zless_add1_eq) - moreover have "\\<^sub>F x in sequentially. c x \ ub" - using \\n. c n \ ub\ by simp - ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1" - apply (elim frequently_rev_mp eventually_mono) - by linarith - moreover have "(\n. B * pa n - c n) \0" - unfolding pa_def using bac_close by auto - from tendstoD[OF this,of 1] - have "\\<^sub>F n in sequentially. \B * pa n - c n\ < 1" - by auto - ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1 \ \B * pa x - c x\ < 1" - using frequently_eventually_frequently by blast - then show False - by (simp add: frequently_def) - qed - - have eq_2_11:"\\<^sub>F n in sequentially. c (n+1)>c n \ - nth_prime (n+1) > nth_prime n + (1 - \)^2 * a n / B" - when "\>0" "\<1" for \::real - proof - - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" - using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto - moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" - proof - - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" - using c_pos by (subst eventually_sequentially_Suc) simp - then show ?thesis using c_pos by eventually_elim auto - qed - ultimately show ?thesis using Bc_large bc_epsilon[OF \\>0\ \\<1\] - proof (eventually_elim, rule_tac impI) - case (elim n) - assume "c n < c (n + 1)" - have "c (n+1) / a n < \" - using a_pos[rule_format,of n] using elim(1,2) by auto - also have "... \ \ * c n" using elim(2) that(1) by auto - finally have "c (n+1) / a n < \ * c n" . - then have "c (n+1) / c n < \ * a n" - using a_pos[rule_format,of n] elim by (auto simp:field_simps) - then have "(1 - \) * a n < a n - c (n+1) / c n" - by (auto simp:algebra_simps) - then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" - apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) - using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) - then have "nth_prime n + (1 - \)^2 * a n / B < nth_prime n + (1 - \) * (a n - c (n+1) / c n) / B" - using \B>0\ by auto - also have "... = nth_prime n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" - using elim by (auto simp:field_simps) - also have "... = nth_prime n + (1 - \) * (nth_prime n / c n)" - proof - - have "B * nth_prime n = c n * a n - c (n + 1)" using elim by auto - from this[symmetric] show ?thesis - using \B>0\ by simp - qed - also have "... = (1+(1-\)/c n) * nth_prime n" - by (auto simp:algebra_simps) - also have "... = ((c n+1-\)/c n) * nth_prime n" - using elim by (auto simp:divide_simps) - also have "... \ ((c (n+1) -\)/c n) * nth_prime n" - proof - - define cp where "cp = c n+1" - have "c (n+1) \ cp" unfolding cp_def using \c n < c (n + 1)\ by auto - moreover have "c n>0" "nth_prime n>0" using elim by auto - ultimately show ?thesis - apply (fold cp_def) - by (auto simp:divide_simps) - qed - also have "... < nth_prime (n+1)" - using elim by (auto simp:divide_simps) - finally show "real (nth_prime n) + (1 - \)\<^sup>2 * real_of_int (a n) - / real_of_int B < real (nth_prime (n + 1))" . - qed - qed - - have c_neq_large:"\\<^sub>F n in sequentially. c (n+1) \ c n" - proof (rule ccontr) - assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" - then have that:"\\<^sub>F n in sequentially. c (n + 1) = c n" - unfolding frequently_def . - have "\\<^sub>F x in sequentially. (B * int (nth_prime x) = c x * a x - c (x + 1) - \ \real_of_int (c (x + 1))\ < real_of_int (a x) / 2) \ 0 < c x \ B < int (nth_prime x) - \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" - using Bc_large c_pos B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] - by eventually_elim (auto simp:divide_simps) - then have "\\<^sub>F m in sequentially. nth_prime (m+1) > (1+1/(2*B))*nth_prime m" - proof (elim frequently_eventually_at_top[OF that, THEN frequently_at_top_elim]) - fix n - assume "c (n + 1) = c n \ - (\y\n. (B * int (nth_prime y) = c y * a y - c (y + 1) \ - \real_of_int (c (y + 1))\ < real_of_int (a y) / 2) \ - 0 < c y \ B < int (nth_prime y) \ (c y < c (y + 1) \ - real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) - < real (nth_prime (y + 1))))" - then have "c (n + 1) = c n" - and Bc_eq:"\y\n. B * int (nth_prime y) = c y * a y - c (y + 1) \ 0 < c y - \ \real_of_int (c (y + 1))\ < real_of_int (a y) / 2 - \ B < int (nth_prime y) - \ (c y < c (y + 1) \ - real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) - < real (nth_prime (y + 1)))" - by auto - obtain m where "n c n" "c nN. N > n \ c N > c n" - using c_ubound[rule_format, of "MAX x\{..n}. c x"] - by (metis Max_ge atMost_iff dual_order.trans finite_atMost finite_imageI image_eqI - linorder_not_le order_refl) - then obtain N where "N>n" "c N>c n" by auto - define A m where "A={m. n (m+1)\N \ c (m+1) > c n}" and "m = Min A" - have "finite A" unfolding A_def - by (metis (no_types, lifting) A_def add_leE finite_nat_set_iff_bounded_le mem_Collect_eq) - moreover have "N-1\A" unfolding A_def - using \c n < c N\ \n < N\ \c (n + 1) = c n\ - by (smt Suc_diff_Suc Suc_eq_plus1 Suc_leI Suc_pred add.commute - add_diff_inverse_nat add_leD1 diff_is_0_eq' mem_Collect_eq nat_add_left_cancel_less - zero_less_one) - ultimately have "m\A" - using Min_in unfolding m_def by auto - then have "n0" - unfolding m_def A_def by auto - moreover have "c m \ c n" - proof (rule ccontr) - assume " \ c m \ c n" - then have "m-1\A" using \m\A\ \c (n + 1) = c n\ - unfolding A_def - by auto (smt One_nat_def Suc_eq_plus1 Suc_lessI less_diff_conv) - from Min_le[OF \finite A\ this,folded m_def] \m>0\ show False by auto - qed - ultimately show ?thesis using that[of m] by auto - qed - have "(1 + 1 / (2 * B)) * nth_prime m < nth_prime m + a m / (2*B)" - proof - - have "nth_prime m < a m" - proof - - have "B * int (nth_prime m) < c m * (a m - 1)" - using Bc_eq[rule_format,of m] \c m \ c n\ \c n < c (m + 1)\ \n < m\ - by (auto simp:algebra_simps) - also have "... \ c n * (a m - 1)" - by (simp add: \c m \ c n\ a_pos mult_right_mono) - finally have "B * int (nth_prime m) < c n * (a m - 1)" . - moreover have "c n\B" - proof - - have " B * int (nth_prime n) = c n * (a n - 1)" "B < int (nth_prime n)" - and c_a:"\real_of_int (c (n + 1))\ < real_of_int (a n) / 2" - using Bc_eq[rule_format,of n] \c (n + 1) = c n\ by (auto simp:algebra_simps) - from this(1) have " c n dvd (B * int (nth_prime n))" - by simp - moreover have "coprime (c n) (int (nth_prime n))" - proof - - have "c n < int (nth_prime n)" - proof (rule ccontr) - assume "\ c n < int (nth_prime n)" - then have asm:"c n \ int (nth_prime n)" by auto - then have "a n > 2 * nth_prime n" - using c_a \c (n + 1) = c n\ by auto - then have "a n -1 \ 2 * nth_prime n" - by simp - then have "a n - 1 > 2 * B" - using \B < int (nth_prime n)\ by auto - from mult_le_less_imp_less[OF asm this] \B>0\ - have "int (nth_prime n) * (2 * B) < c n * (a n - 1)" - by auto - then show False using \B * int (nth_prime n) = c n * (a n - 1)\ - by (smt \0 < B\ \B < int (nth_prime n)\ combine_common_factor - mult.commute mult_pos_pos) - qed - then have "\ nth_prime n dvd c n" - by (simp add: Bc_eq zdvd_not_zless) - then have "coprime (int (nth_prime n)) (c n)" - by (auto intro!:prime_imp_coprime_int) - then show ?thesis using coprime_commute by blast - qed - ultimately have "c n dvd B" - using coprime_dvd_mult_left_iff by auto - then show ?thesis using \0 < B\ zdvd_imp_le by blast - qed - moreover have "c n > 0 " using Bc_eq by blast - ultimately show ?thesis - using \B>0\ by (smt a_pos mult_mono) - qed - then show ?thesis using \B>0\ by (auto simp:field_simps) - qed - also have "... < nth_prime (m+1)" - using Bc_eq[rule_format, of m] \n \c m \ c n\ \c n < c (m+1)\ - by linarith - finally show "\j>n. (1 + 1 / real_of_int (2 * B)) * real (nth_prime j) - < real (nth_prime (j + 1))" using \m>n\ by auto - qed - then have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m > (1+1/(2*B))" - by (auto elim:frequently_elim1 simp:field_simps) - moreover have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m < (1+1/(2*B))" - using ratio_of_consecutive_primes[unfolded tendsto_iff,rule_format,of "1/(2*B)"] - \B>0\ - unfolding dist_real_def - by (auto elim!:eventually_mono simp:algebra_simps) - ultimately show False by (simp add: eventually_mono frequently_def) - qed - - have c_gt_half:"\\<^sub>F N in sequentially. card {n\{N..<2*N}. c n > c (n+1)} > N / 2" - proof - - define h where "h=(\n. (nth_prime (2*n) - nth_prime n) - / sqrt (nth_prime n))" - have "\\<^sub>F n in sequentially. h n < n / 2" - proof - - have "\\<^sub>F n in sequentially. h n < n powr (5/6)" - using nth_prime_double_sqrt_less[of "1/3"] - unfolding h_def by auto - moreover have "\\<^sub>F n in sequentially. n powr (5/6) < (n /2)" - by real_asymp - ultimately show ?thesis - by eventually_elim auto - qed - moreover have "\\<^sub>F n in sequentially. sqrt (nth_prime n) / a n < 1 / (2*B)" - using nth_1[THEN tendsto_real_sqrt,unfolded tendsto_iff - ,rule_format,of "1/(2*B)"] \B>0\ a_pos - by (auto simp:real_sqrt_divide abs_of_pos) - ultimately have "\\<^sub>F x in sequentially. c (x+1) \ c x - \ sqrt (nth_prime x) / a x < 1 / (2*B) - \ h x < x / 2 - \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" - using c_neq_large B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] - by eventually_elim (auto simp:divide_simps) - then show ?thesis - proof (elim eventually_at_top_mono) - fix N assume "N\1" and N_asm:"\y\N. c (y + 1) \ c y \ - sqrt (real (nth_prime y)) / real_of_int (a y) - < 1 / real_of_int (2 * B) \ h y < y / 2 \ - (c y < c (y + 1) \ - real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) - < real (nth_prime (y + 1)))" - - define S where "S={n \ {N..<2 * N}. c n < c (n + 1)}" - define g where "g=(\n. (nth_prime (n+1) - nth_prime n) - / sqrt (nth_prime n))" - define f where "f=(\n. nth_prime (n+1) - nth_prime n)" - have g_gt_1:"g n>1" when "n\N" "c n < c (n + 1)" for n - proof - - have "nth_prime n + sqrt (nth_prime n) < nth_prime (n+1)" - proof - - have "nth_prime n + sqrt (nth_prime n) < nth_prime n + a n / (2*B)" - using N_asm[rule_format,OF \n\N\] a_pos - by (auto simp:field_simps) - also have "... < nth_prime (n+1)" - using N_asm[rule_format,OF \n\N\] \c n < c (n + 1)\ by auto - finally show ?thesis . - qed - then show ?thesis unfolding g_def - using \c n < c (n + 1)\ by auto - qed - have g_geq_0:"g n \ 0" for n - unfolding g_def by auto - - have "finite S" "\x\S. x\N \ c x sum g S" - proof (induct S) - case empty - then show ?case by auto - next - case (insert x F) - moreover have "g x>1" - proof - - have "c x < c (x+1)" "x\N" using insert(4) by auto - then show ?thesis using g_gt_1 by auto - qed - ultimately show ?case by simp - qed - also have "... \ sum g {N..<2*N}" - apply (rule sum_mono2) - unfolding S_def using g_geq_0 by auto - also have "... \ sum (\n. f n/sqrt (nth_prime N)) {N..<2*N}" - apply (rule sum_mono) - unfolding f_def g_def by (auto intro!:divide_left_mono) - also have "... = sum f {N..<2*N} / sqrt (nth_prime N)" - unfolding sum_divide_distrib[symmetric] by auto - also have "... = (nth_prime (2*N) - nth_prime N) / sqrt (nth_prime N)" - proof - - have "sum f {N..<2 * N} = nth_prime (2 * N) - nth_prime N" - proof (induct N) - case 0 - then show ?case by simp - next - case (Suc N) - have ?case if "N=0" - proof - - have "sum f {Suc N..<2 * Suc N} = sum f {1}" - using that by (simp add: numeral_2_eq_2) - also have "... = nth_prime 2 - nth_prime 1" - unfolding f_def by (simp add:numeral_2_eq_2) - also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" - using that by auto - finally show ?thesis . - qed - moreover have ?case if "N\0" - proof - - have "sum f {Suc N..<2 * Suc N} = sum f {N..<2 * Suc N} - f N" - apply (subst (2) sum.atLeast_Suc_lessThan) - using that by auto - also have "... = sum f {N..<2 * N}+ f (2*N) + f(2*N+1) - f N" - by auto - also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" - using Suc unfolding f_def by auto - finally show ?thesis . - qed - ultimately show ?case by blast - qed - then show ?thesis by auto - qed - also have "... = h N" - unfolding h_def by auto - also have "... < N/2" - using N_asm by auto - finally have "card S < N/2" . - - define T where "T={n \ {N..<2 * N}. c n > c (n + 1)}" - have "T \ S = {N..<2 * N}" "T \ S = {}" "finite T" - unfolding T_def S_def using N_asm by fastforce+ - - then have "card T + card S = card {N..<2 * N}" - using card_Un_disjoint \finite S\ by metis - also have "... = N" - by simp - finally have "card T + card S = N" . - with \card S < N/2\ - show "card T > N/2" by linarith - qed - qed - - text\Inequality (3.5) in the original paper required a slight modification: \ - - have a_gt_plus:"\\<^sub>F n in sequentially. c n > c (n+1) \a (n+1) > a n + (a n - c(n+1) - 1) / c (n+1)" - proof - - note a_gt_1[THEN eventually_all_ge_at_top] c_pos[THEN eventually_all_ge_at_top] - moreover have "\\<^sub>F n in sequentially. - B * int (nth_prime (n+1)) = c (n+1) * a (n+1) - c (n + 2)" - using Bc_large - apply (subst (asm) eventually_sequentially_Suc[symmetric]) - by (auto elim:eventually_mono) - moreover have "\\<^sub>F n in sequentially. - B * int (nth_prime n) = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" - using Bc_large by (auto elim:eventually_mono) - ultimately show ?thesis - apply (eventually_elim) - proof (rule impI) - fix n - assume "\y\n. 1 < a y" "\y\n. 0 < c y" - and - Suc_n_eq:"B * int (nth_prime (n + 1)) = c (n + 1) * a (n + 1) - c (n + 2)" and - "B * int (nth_prime n) = c n * a n - c (n + 1) \ - real_of_int \c (n + 1)\ < real_of_int (a n) / 2" - and "c (n + 1) < c n" - then have n_eq:"B * int (nth_prime n) = c n * a n - c (n + 1)" and - c_less_a: "real_of_int \c (n + 1)\ < real_of_int (a n) / 2" - by auto - from \\y\n. 1 < a y\ \\y\n. 0 < c y\ - have *:"a n>1" "a (n+1) > 1" "c n > 0" - "c (n+1) > 0" "c (n+2) > 0" - by auto - then have "(1+1/c (n+1))* (a n - 1)/a (n+1) = (c (n+1)+1) * ((a n - 1) / (c (n+1) * a (n+1)))" - by (auto simp:field_simps) - also have "... \ c n * ((a n - 1) / (c (n+1) * a (n+1)))" - apply (rule mult_right_mono) - subgoal using \c (n + 1) < c n\ by auto - subgoal by (smt \0 < c (n + 1)\ a_pos divide_nonneg_pos mult_pos_pos of_int_0_le_iff - of_int_0_less_iff) - done - also have "... = (c n * (a n - 1)) / (c (n+1) * a (n+1))" by auto - also have "... < (c n * (a n - 1)) / (c (n+1) * a (n+1) - c (n+2))" - apply (rule divide_strict_left_mono) - subgoal using \c (n+2) > 0\ by auto - unfolding Suc_n_eq[symmetric] using * \B>0\ by auto - also have "... < (c n * a n - c (n+1)) / (c (n+1) * a (n+1) - c (n+2))" - apply (rule frac_less) - unfolding Suc_n_eq[symmetric] using * \B>0\ \c (n + 1) < c n\ - by (auto simp:algebra_simps) - also have "... = nth_prime n / nth_prime (n+1)" - unfolding Suc_n_eq[symmetric] n_eq[symmetric] using \B>0\ by auto - also have "... < 1" by auto - finally have "(1 + 1 / real_of_int (c (n + 1))) * real_of_int (a n - 1) - / real_of_int (a (n + 1)) < 1 " . - then show "a n + (a n - c (n + 1) - 1) / (c (n + 1)) < (a (n + 1))" - using * by (auto simp:field_simps) - qed - qed - have a_gt_1:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + 1" - using Bc_large a_gt_plus c_pos[THEN eventually_all_ge_at_top] - apply eventually_elim - proof (rule impI) - fix n assume - "c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" - "c (n + 1) < c n" and B_eq:"B * int (nth_prime n) = c n * a n - c (n + 1) \ - \real_of_int (c (n + 1))\ < real_of_int (a n) / 2" and c_pos:"\y\n. 0 < c y" - from this(1,2) - have "a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" by auto - moreover have "a n - 2 * c (n+1) > 0" - using B_eq c_pos[rule_format,of "n+1"] by auto - then have "a n - 2 * c (n+1) \ 1" by simp - then have "(a n - c (n + 1) - 1) / c (n + 1) \ 1" - using c_pos[rule_format,of "n+1"] by (auto simp:field_simps) - ultimately show "a n + 1 < a (n + 1)" by auto - qed - - text\The following corresponds to inequality (3.6) in the paper, which had to be - slightly corrected: \ - - have a_gt_sqrt:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + (sqrt n - 2)" - proof - - have a_2N:"\\<^sub>F N in sequentially. a (2*N) \ N /2 +1" - using c_gt_half a_gt_1[THEN eventually_all_ge_at_top] - proof eventually_elim - case (elim N) - define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" - define f where "f = (\n. a (Suc n) - a n)" - - have f_1:"\x\S. f x\1" and f_0:"\x. f x\0" - subgoal using elim unfolding S_def f_def by auto - subgoal using \mono a\[THEN incseq_SucD] unfolding f_def by auto - done - have "N / 2 < card S" - using elim unfolding S_def by auto - also have "... \ sum f S" - unfolding of_int_sum - apply (rule sum_bounded_below[of _ 1,simplified]) - using f_1 by auto - also have "... \ sum f {N..<2 * N}" - unfolding of_int_sum - apply (rule sum_mono2) - unfolding S_def using f_0 by auto - also have "... = a (2*N) - a N" - unfolding of_int_sum f_def of_int_diff - apply (rule sum_Suc_diff') - by auto - finally have "N / 2 < a (2*N) - a N" . - then show ?case using a_pos[rule_format,of N] by linarith - qed - - have a_n4:"\\<^sub>F n in sequentially. a n > n/4" - proof - - obtain N where a_N:"\n\N. a (2*n) \ n /2+1" - using a_2N unfolding eventually_at_top_linorder by auto - have "a n>n/4" when "n\2*N" for n - proof - - define n' where "n'=n div 2" - have "n'\N" unfolding n'_def using that by auto - have "n/4 < n' /2+1" - unfolding n'_def by auto - also have "... \ a (2*n')" - using a_N \n'\N\ by auto - also have "... \a n" unfolding n'_def - apply (cases "even n") - subgoal by simp - subgoal by (simp add: assms(2) incseqD) - done - finally show ?thesis . - qed - then show ?thesis - unfolding eventually_at_top_linorder by auto - qed - - have c_sqrt:"\\<^sub>F n in sequentially. c n < sqrt n / 4" - proof - - have "\\<^sub>F x in sequentially. x>1" by simp - moreover have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" - using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] - by simp - ultimately have "\\<^sub>F n in sequentially. c n < B*8 *ln n + 1" using a_n4 Bc_large - proof eventually_elim - case (elim n) - from this(4) have "c n=(B*nth_prime n+c (n+1))/a n" - using a_pos[rule_format,of n] - by (auto simp:divide_simps) - also have "... = (B*nth_prime n)/a n+c (n+1)/a n" - by (auto simp:divide_simps) - also have "... < (B*nth_prime n)/a n + 1" - proof - - have "c (n+1)/a n < 1" using elim(4) by auto - then show ?thesis by auto - qed - also have "... < B*8 * ln n + 1" - proof - - have "B*nth_prime n < 2*B*n*ln n" - using \real (nth_prime n) / (real n * ln (real n)) < 2\ \B>0\ \ 1 < n\ - by (auto simp:divide_simps) - moreover have "real n / 4 < real_of_int (a n)" by fact - ultimately have "(B*nth_prime n) / a n < (2*B*n*ln n) / (n/4)" - apply (rule_tac frac_less) - using \B>0\ \ 1 < n\ by auto - also have "... = B*8 * ln n" - using \ 1 < n\ by auto - finally show ?thesis by auto - qed - finally show ?case . - qed - moreover have "\\<^sub>F n in sequentially. B*8 *ln n + 1 < sqrt n / 4" - by real_asymp - ultimately show ?thesis - by eventually_elim auto - qed - - have - "\\<^sub>F n in sequentially. 0 < c (n+1)" - "\\<^sub>F n in sequentially. c (n+1) < sqrt (n+1) / 4" - "\\<^sub>F n in sequentially. n > 4" - "\\<^sub>F n in sequentially. (n - 4) / sqrt (n + 1) + 1 > sqrt n" - subgoal using c_pos[THEN eventually_all_ge_at_top] - by eventually_elim auto - subgoal using c_sqrt[THEN eventually_all_ge_at_top] - by eventually_elim (use le_add1 in blast) - subgoal by simp - subgoal - by real_asymp - done - then show ?thesis using a_gt_plus a_n4 - apply eventually_elim - proof (rule impI) - fix n assume asm:"0 < c (n + 1)" "c (n + 1) < sqrt (real (n + 1)) / 4" and - a_ineq:"c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" - "c (n + 1) < c n" and "n / 4 < a n" "n > 4" - and n_neq:" sqrt (real n) < real (n - 4) / sqrt (real (n + 1)) + 1" - - have "(n-4) / sqrt(n+1) = (n/4 - 1)/ (sqrt (real (n + 1)) / 4)" - using \n>4\ by (auto simp:divide_simps) - also have "... < (a n - 1) / c (n + 1)" - apply (rule frac_less) - using \n > 4\ \n / 4 < a n\ \0 < c (n + 1)\ \c (n + 1) < sqrt (real (n + 1)) / 4\ - by auto - also have "... - 1 = (a n - c (n + 1) - 1) / c (n + 1)" - using \0 < c (n + 1)\ by (auto simp:field_simps) - also have "a n + ... < a (n+1)" - using a_ineq by auto - finally have "a n + ((n - 4) / sqrt (n + 1) - 1) < a (n + 1)" by simp - moreover have "(n - 4) / sqrt (n + 1) - 1 > sqrt n - 2" - using n_neq[THEN diff_strict_right_mono,of 2] \n>4\ - by (auto simp:algebra_simps of_nat_diff) - ultimately show "real_of_int (a n) + (sqrt (real n) - 2) < real_of_int (a (n + 1))" - by argo - qed - qed - - text\The following corresponds to inequality $ a_{2N} > N^{3/2}/2$ in the paper, - which had to be slightly corrected: \ - - have a_2N_sqrt:"\\<^sub>F N in sequentially. a (2*N) > real N * (sqrt (real N)/2 - 1)" - using c_gt_half a_gt_sqrt[THEN eventually_all_ge_at_top] eventually_gt_at_top[of 4] - proof eventually_elim - case (elim N) - define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" - define f where "f = (\n. a (Suc n) - a n)" - - have f_N:"\x\S. f x\sqrt N - 2" - proof - fix x assume "x\S" - then have "sqrt (real x) - 2 < f x" "x\N" - using elim unfolding S_def f_def by auto - moreover have "sqrt x - 2 \ sqrt N - 2" - using \x\N\ by simp - ultimately show "sqrt (real N) - 2 \ real_of_int (f x)" by argo - qed - have f_0:"\x. f x\0" - using \mono a\[THEN incseq_SucD] unfolding f_def by auto - - have "(N / 2) * (sqrt N - 2) < card S * (sqrt N - 2)" - apply (rule mult_strict_right_mono) - subgoal using elim unfolding S_def by auto - subgoal using \N>4\ - by (metis diff_gt_0_iff_gt numeral_less_real_of_nat_iff real_sqrt_four real_sqrt_less_iff) - done - also have "... \ sum f S" - unfolding of_int_sum - apply (rule sum_bounded_below) - using f_N by auto - also have "... \ sum f {N..<2 * N}" - unfolding of_int_sum - apply (rule sum_mono2) - unfolding S_def using f_0 by auto - also have "... = a (2*N) - a N" - unfolding of_int_sum f_def of_int_diff - apply (rule sum_Suc_diff') - by auto - finally have "real N / 2 * (sqrt (real N) - 2) < real_of_int (a (2 * N) - a N)" - . - then have "real N / 2 * (sqrt (real N) - 2) < a (2 * N)" - using a_pos[rule_format,of N] by linarith - then show ?case by (auto simp:field_simps) - qed - - text\The following part is required to derive the final contradiction of the proof.\ - - have a_n_sqrt:"\\<^sub>F n in sequentially. a n > (((n-1)/2) powr (3/2) - (n-1)) /2" - proof (rule sequentially_even_odd_imp) - define f where "f=(\N. ((real (2 * N - 1) / 2) powr (3 / 2) - real (2 * N - 1)) / 2)" - define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" - have "\\<^sub>F N in sequentially. g N > f N" - unfolding f_def g_def - by real_asymp - moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" - unfolding g_def using a_2N_sqrt . - ultimately show "\\<^sub>F N in sequentially. f N < a (2 * N)" - by eventually_elim auto - next - define f where "f=(\N. ((real (2 * N + 1 - 1) / 2) powr (3 / 2) - - real (2 * N + 1 - 1)) / 2)" - define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" - have "\\<^sub>F N in sequentially. g N = f N" - using eventually_gt_at_top[of 0] - apply eventually_elim - unfolding f_def g_def - by (auto simp:algebra_simps powr_half_sqrt[symmetric] powr_mult_base) - moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" - unfolding g_def using a_2N_sqrt . - moreover have "\\<^sub>F N in sequentially. a (2 * N + 1) \ a (2*N)" - apply (rule eventuallyI) - using \mono a\ by (simp add: incseqD) - ultimately show "\\<^sub>F N in sequentially. f N < (a (2 * N + 1))" - apply eventually_elim - by auto - qed - - have a_nth_prime_gt:"\\<^sub>F n in sequentially. a n / nth_prime n > 1" - proof - - define f where "f=(\n::nat. (((n-1)/2) powr (3/2) - (n-1)) /2)" - have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" - using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] - by simp - from this[] eventually_gt_at_top[of 1] - have "\\<^sub>F n in sequentially. real (nth_prime n) < 2*(real n * ln n)" - apply eventually_elim - by (auto simp:field_simps) - moreover have *:"\\<^sub>F N in sequentially. f N >0 " - unfolding f_def - by real_asymp - moreover have " \\<^sub>F n in sequentially. f n < a n" - using a_n_sqrt unfolding f_def . - ultimately have "\\<^sub>F n in sequentially. a n / nth_prime n - > f n / (2*(real n * ln n))" - apply eventually_elim - apply (rule frac_less2) - by auto - moreover have "\\<^sub>F n in sequentially. - (f n)/ (2*(real n * ln n)) > 1" - unfolding f_def - by real_asymp - ultimately show ?thesis - by eventually_elim argo - qed - - have a_nth_prime_lt:"\\<^sub>F n in sequentially. a n / nth_prime n < 1" - proof - - have "liminf (\x. a x / nth_prime x) < 1" - using nth_2 by auto - from this[unfolded less_Liminf_iff] - show ?thesis - apply (auto elim!:frequently_elim1) - by (meson divide_less_eq_1 ereal_less_eq(7) leD leI - nth_prime_nonzero of_nat_eq_0_iff of_nat_less_0_iff order.trans) - qed - - from a_nth_prime_gt a_nth_prime_lt show False - by (simp add: eventually_mono frequently_def) -qed - -section\Acknowledgements\ - -text\A.K.-A. and W.L. were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) - funded by the European Research Council and led by Professor Lawrence Paulson - at the University of Cambridge, UK.\ - +(* Title: Irrational_Series_Erdos_Straus.thy + Author: Angeliki Koutsoukou-Argyraki and Wenda Li, University of Cambridge, UK. + +We formalise certain irrationality criteria for infinite series by P. Erdos and E.G. Straus. +In particular, we formalise Theorem 2.1, Corollary 2.10 and Theorem 3.1 in [1]. The latter is an +application of Theorem 2.1 involving the prime numbers. + +References: +[1] P. Erdos and E.G. Straus, On the irrationality of certain series, Pacific Journal of +Mathematics, Vol. 55, No 1, 1974 https://projecteuclid.org/euclid.pjm/1102911140 +*) + +theory "Irrational_Series_Erdos_Straus" imports + Prime_Number_Theorem.Prime_Number_Theorem + Prime_Distribution_Elementary.PNT_Consequences +begin + +section \Miscellaneous\ + +lemma suminf_comparison: + assumes "summable f" and gf: "\n. norm (g n) \ f n" + shows "suminf g \ suminf f" +proof (rule suminf_le) + show "g n \ f n" for n + using gf[of n] by auto + show "summable g" + using assms summable_comparison_test' by blast + show "summable f" using assms(1) . +qed + +lemma tendsto_of_int_diff_0: + assumes "(\n. f n - of_int(g n)) \ (0::real)" "\\<^sub>F n in sequentially. f n > 0" + shows "\\<^sub>F n in sequentially. 0 \ g n" +proof - + have "\\<^sub>F n in sequentially. \f n - of_int(g n)\ < 1 / 2" + using assms(1)[unfolded tendsto_iff,rule_format,of "1/2"] by auto + then show ?thesis using assms(2) + apply eventually_elim + by linarith +qed + +lemma eventually_mono_sequentially: + assumes "eventually P sequentially" + assumes "\x. P (x+k) \ Q (x+k)" + shows "eventually Q sequentially" + using sequentially_offset[OF assms(1),of k] + apply (subst eventually_sequentially_seg[symmetric,of _ k]) + apply (elim eventually_mono) + by fact + +lemma frequently_eventually_at_top: + fixes P Q::"'a::linorder \ bool" + assumes "frequently P at_top" "eventually Q at_top" + shows "frequently (\x. P x \ (\y\x. Q y) ) at_top" + using assms + unfolding frequently_def eventually_at_top_linorder + by (metis (mono_tags, hide_lams) le_cases order_trans) + +lemma eventually_at_top_mono: + fixes P Q::"'a::linorder \ bool" + assumes event_P:"eventually P at_top" + assumes PQ_imp:"\x. x\z \ \y\x. P y \ Q x" + shows "eventually Q at_top" +proof - + obtain N where N_P:"\n\N. P n" + using event_P[unfolded eventually_at_top_linorder] by auto + define N' where "N' = max N z" + have "Q x" when "x\N'" for x + apply (rule PQ_imp) + using N_P that unfolding N'_def by auto + then show ?thesis unfolding eventually_at_top_linorder + by auto +qed + +lemma frequently_at_top_elim: + fixes P Q::"'a::linorder \ bool" + assumes "\\<^sub>Fx in at_top. P x" + assumes "\i. P i \ \j>i. Q j" + shows "\\<^sub>Fx in at_top. Q x" + using assms unfolding frequently_def eventually_at_top_linorder + by (meson leD le_cases less_le_trans) + +lemma less_Liminf_iff: + fixes X :: "_ \ _ :: complete_linorder" + shows "Liminf F X < C \ (\yx. y \ X x) F)" + apply (subst Not_eq_iff[symmetric]) + apply (simp add:not_less not_frequently not_le le_Liminf_iff) + by force + +lemma sequentially_even_odd_imp: + assumes "\\<^sub>F N in sequentially. P (2*N)" "\\<^sub>F N in sequentially. P (2*N+1)" + shows "\\<^sub>F n in sequentially. P n" +proof - + obtain N where N_P:"\x\N. P (2 * x) \ P (2 * x + 1)" + using eventually_conj[OF assms] + unfolding eventually_at_top_linorder by auto + define N' where "N'=2*N " + have "P n" when "n\2*N" for n + proof - + define n' where "n'= n div 2" + then have "n'\N" using that by auto + then have "P (2 * n') \ P (2 * n' + 1)" + using N_P by auto + then show ?thesis unfolding n'_def + apply (cases "even n") + by auto + qed + then show ?thesis unfolding eventually_at_top_linorder by auto +qed + + +section \Theorem 2.1 and Corollary 2.10\ + +context + fixes a b ::"nat\int " + assumes a_pos:"\ n. a n >0 " and a_large:"\\<^sub>F n in sequentially. a n > 1" + and ab_tendsto: "(\n. \b n\ / (a (n-1)*a n)) \ 0" +begin + +private lemma aux_series_summable: "summable (\n. b n / (\k\n. a k))" +proof - + have "\e>0. \\<^sub>F x in sequentially. \b x\ / (a (x-1) * a x) < e" + using ab_tendsto[unfolded tendsto_iff] + apply (simp add:of_int_abs[symmetric] abs_mult del:of_int_abs) + by (subst (asm) (2) abs_of_pos,use \\ n. a n > 0\ in auto)+ + from this[rule_format,of 1] + have "\\<^sub>F x in sequentially. \real_of_int(b x)\ < (a (x-1) * a x)" + using \\ n. a n >0\ by auto + moreover have "\n. (\k\n. real_of_int (a k)) > 0" + using a_pos by (auto intro!:linordered_semidom_class.prod_pos) + ultimately have "\\<^sub>F n in sequentially. \b n\ / (\k\n. a k) + < (a (n-1) * a n) / (\k\n. a k)" + apply (elim eventually_mono) + by (auto simp add:field_simps) + moreover have "\b n\ / (\k\n. a k) = norm (b n / (\k\n. a k))" for n + using \\n. (\k\n. real_of_int (a k)) > 0\[rule_format,of n] by auto + ultimately have "\\<^sub>F n in sequentially. norm (b n / (\k\n. a k)) + < (a (n-1) * a n) / (\k\n. a k)" + by algebra + moreover have "summable (\n. (a (n-1) * a n) / (\k\n. a k))" + proof - + obtain s where a_gt_1:"\ n\s. a n >1" + using a_large[unfolded eventually_at_top_linorder] by auto + define cc where "cc= (\k0" + unfolding cc_def + apply (rule linordered_semidom_class.prod_pos) + using a_pos by auto + have "(\k\n+s. a k) \ cc * 2^n" for n + proof - + have "prod a {s.. 2^n" + proof (induct n) + case 0 + then show ?case using a_gt_1 by auto + next + case (Suc n) + moreover have "a (s + Suc n) \ 2" + using a_gt_1 by (smt le_add1) + ultimately show ?case + apply (subst prod.atLeastLessThan_Suc,simp) + using mult_mono'[of 2 "a (Suc (s + n))" " 2 ^ n" "prod a {s..cc>0\ unfolding cc_def by (simp add: atLeast0AtMost) + qed + then have "1/(\k\n+s. a k) \ 1/(cc * 2^n)" for n + proof - + assume asm:"\n. cc * 2 ^ n \ prod a {..n + s}" + then have "real_of_int (cc * 2 ^ n) \ prod a {..n + s}" using of_int_le_iff by blast + moreover have "prod a {..n + s} >0" using \cc>0\ by (simp add: a_pos prod_pos) + ultimately show ?thesis using \cc>0\ + by (auto simp:field_simps simp del:of_int_prod) + qed + moreover have "summable (\n. 1/(cc * 2^n))" + proof - + have "summable (\n. 1/(2::int)^n)" + using summable_geometric[of "1/(2::int)"] by (simp add:power_one_over) + from summable_mult[OF this,of "1/cc"] show ?thesis by auto + qed + ultimately have "summable (\n. 1 / (\k\n+s. a k))" + apply (elim summable_comparison_test'[where N=0]) + apply (unfold real_norm_def, subst abs_of_pos) + by (auto simp add: \\n. 0 < (\k\n. real_of_int (a k))\) + then have "summable (\n. 1 / (\k\n. a k))" + apply (subst summable_iff_shift[where k=s,symmetric]) + by simp + then have "summable (\n. (a (n+1) * a (n+2)) / (\k\n+2. a k))" + proof - + assume asm:"summable (\n. 1 / real_of_int (prod a {..n}))" + have "1 / real_of_int (prod a {..n}) = (a (n+1) * a (n+2)) / (\k\n+2. a k)" for n + proof - + have "a (Suc (Suc n)) \ 0" "a (Suc n) \0" + using a_pos by (metis less_irrefl)+ + then show ?thesis + by (simp add: atLeast0_atMost_Suc atMost_atLeast0) + qed + then show ?thesis using asm by auto + qed + then show "summable (\n. (a (n-1) * a n) / (\k\n. a k))" + apply (subst summable_iff_shift[symmetric,of _ 2]) + by auto + qed + ultimately show ?thesis + apply (elim summable_comparison_test_ev[rotated]) + by (simp add: eventually_mono) +qed + +private fun get_c::"(nat \ int) \ (nat \ int) \ int \ nat \ (nat \ int)" where + "get_c a' b' B N 0 = round (B * b' N / a' N)"| + "get_c a' b' B N (Suc n) = get_c a' b' B N n * a' (n+N) - B * b' (n+N)" + +lemma ab_rationality_imp: + assumes ab_rational:"(\n. (b n / (\i \ n. a i))) \ \" + shows "\ (B::int)>0. \ c::nat\ int. + (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\ (\n. c (Suc n) / a n) \ 0" +proof - + have [simp]:"a n \ 0" for n using a_pos by (metis less_numeral_extra(3)) + obtain A::int and B::int where + AB_eq:"(\n. real_of_int (b n) / real_of_int (prod a {..n})) = A / B" and "B>0" + proof - + obtain q::rat where "(\n. real_of_int (b n) / real_of_int (prod a {..n})) = real_of_rat q" + using ab_rational by (rule Rats_cases) simp + moreover obtain A::int and B::int where "q = Rat.Fract A B" "B > 0" "coprime A B" + by (rule Rat_cases) auto + ultimately show ?thesis by (auto intro!: that[of A B] simp:of_rat_rat) + qed + define f where "f = (\n. b n / real_of_int (prod a {..n}))" + define R where "R = (\N. (\n. B*b (n+N+1) / prod a {N..n+N+1}))" + have all_e_ubound:"\e>0. \\<^sub>F M in sequentially. \n. \B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" + proof safe + fix e::real assume "e>0" + obtain N where N_a2: "\n \ N. a n \ 2" + and N_ba: "\n \ N. \b n\ / (a (n-1) * a n) < e/(4*B)" + proof - + have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B)" + using order_topology_class.order_tendstoD[OF ab_tendsto,of "e/(4*B)"] \B>0\ \e>0\ + by auto + moreover have "\\<^sub>F n in sequentially. a n \ 2" + using a_large by (auto elim: eventually_mono) + ultimately have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B) \ a n \ 2" + by eventually_elim auto + then show ?thesis unfolding eventually_at_top_linorder using that + by auto + qed + have geq_N_bound:"\B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" when "M\N" for n M + proof - + define D where "D = B*b (n+M+1)/ (a (n+M) * a (n+M+1))" + have "\B*b (n+M+1) / prod a {M..n+M+1}\ = \D / prod a {M.." + proof - + have "{M..n+M+1} = {M.. {n+M,n+M+1}" by auto + then have "prod a {M..n+M+1} = a (n+M) * a (n+M+1)* prod a {M..e/4 * (1/prod a {M.." + proof - + have "\D\ < e/ 4" + unfolding D_def using N_ba[rule_format, of "n+M+1"] \B>0\ \M \ N\ \e>0\ a_pos + by (auto simp:field_simps abs_mult abs_of_pos) + from mult_strict_right_mono[OF this,of "1/prod a {M..e>0\ + show ?thesis + apply (auto simp:abs_prod abs_mult prod_pos) + by (subst (2) abs_of_pos,auto)+ + qed + also have "... \ e/4 * 1/2^n" + proof - + have "prod a {M.. 2^n" + proof (induct n) + case 0 + then show ?case by simp + next + case (Suc n) + then show ?case + using \M\N\ by (simp add: N_a2 mult.commute mult_mono' prod.atLeastLessThan_Suc) + qed + then have "real_of_int (prod a {M.. 2^n" + using numeral_power_le_of_int_cancel_iff by blast + then show ?thesis using \e>0\ by (auto simp add:divide_simps) + qed + finally show ?thesis . + qed + show "\\<^sub>F M in sequentially. \n. \real_of_int (B * b (n + M + 1)) + / real_of_int (prod a {M..n + M + 1})\ < e / 4 * 1 / 2 ^ n" + apply (rule eventually_sequentiallyI[of N]) + using geq_N_bound by blast + qed + have R_tendsto_0:"R \ 0" + proof (rule tendstoI) + fix e::real assume "e>0" + show "\\<^sub>F x in sequentially. dist (R x) 0 < e" using all_e_ubound[rule_format,OF \e>0\] + proof eventually_elim + case (elim M) + define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" + have g_lt:"\g n\ < e/4 * 1/2^n" for n + using elim unfolding g_def by auto + have g_abs_summable:"summable (\n. \g n\)" + proof - + have "summable (\n. e/4 * 1/2^n)" + using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] + by (auto simp add:algebra_simps power_divide) + then show ?thesis + apply (elim summable_comparison_test') + using g_lt less_eq_real_def by auto + qed + have "\\n. g n\ \ (\n. \g n\)" by (rule summable_rabs[OF g_abs_summable]) + also have "... \(\n. e/4 * 1/2^n)" + proof (rule suminf_comparison) + show "summable (\n. e/4 * 1/2^n)" + using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] + by (auto simp add:algebra_simps power_divide) + show "\n. norm \g n\ \ e / 4 * 1 / 2 ^ n" using g_lt less_eq_real_def by auto + qed + also have "... = (e/4) * (\n. (1/2)^n)" + apply (subst suminf_mult[symmetric]) + subgoal + apply (rule complete_algebra_summable_geometric) + by simp + subgoal by (auto simp:algebra_simps power_divide) + done + also have "... = e/2" by (simp add:suminf_geometric[of "1/2"]) + finally have "\\n. g n\ \ e / 2" . + then show "dist (R M) 0 < e" unfolding R_def g_def using \e>0\ by auto + qed + qed + + obtain N where R_N_bound:"\M \ N. \R M\ \ 1 / 4" + and N_geometric:"\M\N. \n. \real_of_int (B * b (n + M + 1)) / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" + proof - + obtain N1 where N1:"\M \ N1. \R M\ \ 1 / 4" + using metric_LIMSEQ_D[OF R_tendsto_0,of "1/4"] all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] + by (auto simp:less_eq_real_def) + obtain N2 where N2:"\M\N2. \n. \real_of_int (B * b (n + M + 1)) + / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" + using all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] + by (auto simp:less_eq_real_def) + define N where "N=max N1 N2" + show ?thesis using that[of N] N1 N2 unfolding N_def by simp + qed + + define C where "C = B * prod a {..nn. f n)" + unfolding AB_eq f_def using \B>0\ by auto + also have "... = B * prod a {..nn. f (n+N+1)))" + using suminf_split_initial_segment[OF \summable f\, of "N+1"] by auto + also have "... = B * prod a {..nn. f (n+N+1)))" + using sum.atLeast0_lessThan_Suc by simp + also have "... = C + B * b N / a N + (\n. B*b (n+N+1) / prod a {N..n+N+1})" + proof - + have "B * prod a {.. {N}" using ivl_disj_un_singleton(2) by blast + then show ?thesis unfolding f_def by auto + qed + moreover have "B * prod a {..n. f (n+N+1)) = (\n. B*b (n+N+1) / prod a {N..n+N+1})" + proof - + have "summable (\n. f (n + N + 1))" + using \summable f\ summable_iff_shift[of f "N+1"] by auto + moreover have "prod a {.. {N..n + N + 1}" by auto + then show ?thesis + unfolding f_def + apply simp + apply (subst prod.union_disjoint) + by auto + qed + ultimately show ?thesis + apply (subst suminf_mult[symmetric]) + by (auto simp add: mult.commute mult.left_commute) + qed + ultimately show ?thesis unfolding C_def by (auto simp:algebra_simps) + qed + also have "... = C +B * b N / a N + R N" + unfolding R_def by simp + finally show ?thesis . + qed + have R_bound:"\R M\ \ 1 / 4" and R_Suc:"R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" + when "M \ N" for M + proof - + define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" + have g_abs_summable:"summable (\n. \g n\)" + proof - + have "summable (\n.(1::real)/2^n)" + using summable_geometric[of "(1::real)/2",simplified] + by (auto elim!: back_subst[of "summable"] simp:field_simps) + moreover have "\g n\ < 1/2^n" for n + using N_geometric[rule_format,OF that] unfolding g_def by simp + ultimately show ?thesis + apply (elim summable_comparison_test') + using less_eq_real_def by auto + qed + show "\R M\ \ 1 / 4" using R_N_bound[rule_format,OF that] . + have "R M = (\n. g n)" unfolding R_def g_def by simp + also have "... = g 0 + (\n. g (Suc n))" + apply (subst suminf_split_head) + using summable_rabs_cancel[OF g_abs_summable] by auto + also have "... = g 0 + 1/a M * (\n. a M * g (Suc n))" + apply (subst suminf_mult) + by (auto simp add: g_abs_summable summable_Suc_iff summable_rabs_cancel) + also have "... = g 0 + 1/a M * R (Suc M)" + proof - + have "a M * g (Suc n) = B * b (n + M + 2) / prod a {Suc M..n + M + 2}" for n + proof - + have "{M..Suc (Suc (M + n))} = {M} \ {Suc M..Suc (Suc (M + n))}" by auto + then show ?thesis + unfolding g_def using \B>0\ by (auto simp add:algebra_simps) + qed + then have "(\n. a M * g (Suc n)) = R (Suc M)" + unfolding R_def by auto + then show ?thesis by auto + qed + finally have "R M = g 0 + 1 / a M * R (Suc M)" . + then have "R (Suc M) = a M * R M - g 0 * a M" + by (auto simp add:algebra_simps) + moreover have "{M..Suc M} = {M,Suc M}" by auto + ultimately show "R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" + unfolding g_def by auto + qed + + define c where "c = (\n. if n\N then get_c a b B N (n-N) else undefined)" + have c_rec:"c (n+1) = c n * a n - B * b n" when "n \ N" for n + unfolding c_def using that by (auto simp:Suc_diff_le) + have c_R:"c (Suc n) / a n = R n" when "n \ N" for n + using that + proof (induct rule:nat_induct_at_least) + case base + have "\ c (N+1) / a N \ \ 1/2" + proof - + have "c N = round (B * b N / a N)" unfolding c_def by simp + moreover have "c (N+1) / a N = c N - B * b N / a N" + using a_pos[rule_format,of N] + by (auto simp add:c_rec[of N,simplified] divide_simps) + ultimately show ?thesis using of_int_round_abs_le by auto + qed + moreover have "\R N\ \ 1 / 4" using R_bound[of N] by simp + ultimately have "\c (N+1) / a N - R N \ < 1" by linarith + moreover have "c (N+1) / a N - R N \ \" + proof - + have "c (N+1) / a N = c N - B * b N / a N" + using a_pos[rule_format,of N] + by (auto simp add:c_rec[of N,simplified] divide_simps) + moreover have " B * b N / a N + R N \ \" + proof - + have "C = B * (\nn {..n}" if "nB>0\ + apply simp + apply (subst prod.union_disjoint) + by auto + qed + finally have "C = real_of_int (B * (\n \" using Ints_of_int by blast + moreover note \A * prod a {.. + ultimately show ?thesis + by (metis Ints_diff Ints_of_int add.assoc add_diff_cancel_left') + qed + ultimately show ?thesis by (simp add: diff_diff_add) + qed + ultimately have "c (N+1) / a N - R N = 0" + by (metis Ints_cases less_irrefl of_int_0 of_int_lessD) + then show ?case by simp + next + case (Suc n) + have "c (Suc (Suc n)) / a (Suc n) = c (Suc n) - B * b (Suc n) / a (Suc n)" + apply (subst c_rec[of "Suc n",simplified]) + using \N \ n\ by (auto simp add: divide_simps) + also have "... = a n * R n - B * b (Suc n) / a (Suc n)" + using Suc by (auto simp: divide_simps) + also have "... = R (Suc n)" + using R_Suc[OF \N \ n\] by simp + finally show ?case . + qed + have ca_tendsto_zero:"(\n. c (Suc n) / a n) \ 0" + using R_tendsto_0 + apply (elim filterlim_mono_eventually) + using c_R by (auto intro!:eventually_sequentiallyI[of N]) + have ca_bound:"\c (n + 1)\ < a n / 2" when "n \ N" for n + proof - + have "\c (Suc n)\ / a n = \c (Suc n) / a n\" using a_pos[rule_format,of n] by auto + also have "... = \R n\" using c_R[OF that] by auto + also have "... < 1/2" using R_bound[OF that] by auto + finally have "\c (Suc n)\ / a n < 1/2" . + then show ?thesis using a_pos[rule_format,of n] by auto + qed + +(* (* the following part corresponds to (2.7) (2.8) in the original paper, but turns out to be + not necessary. *) + have c_round:"c n = round (B * b n / a n)" when "n \ N" for n + proof (cases "n=N") + case True + then show ?thesis unfolding c_def by simp + next + case False + with \n\N\ obtain n' where n_Suc:"n=Suc n'" and "n' \ N" + by (metis le_eq_less_or_eq lessE less_imp_le_nat) + have "B * b n / a n = c n - R n" + proof - + have "R n = c n - B * b n / a n" + using c_R[OF \n'\N\,symmetric,folded n_Suc] R_Suc[OF \n'\N\,folded n_Suc] + by (auto simp:field_simps) + then show ?thesis by (auto simp:field_simps) + qed + then have "\B * b n / a n - c n\ = \R n\" by auto + then have "\B * b n / a n - c n\ < 1/2" using R_bound[OF \n \ N\] by auto + from round_unique'[OF this] show ?thesis by auto + qed + *) + show "\B>0. \c. (\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) + \ real_of_int \c (n + 1)\ < a n / 2) \ (\n. c (Suc n) / a n) \ 0" + unfolding eventually_at_top_linorder + apply (rule exI[of _ B],use \B>0\ in simp) + apply (intro exI[of _c] exI[of _ N]) + using c_rec ca_bound ca_tendsto_zero + by fastforce +qed + +private lemma imp_ab_rational: + assumes "\ (B::int)>0. \ c::nat\ int. + (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\n. (b n / (\i \ n. a i))) \ \" +proof - + obtain B::int and c::"nat\int" and N::nat where "B>0" and + large_n:"\n\N. B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < a n / 2 + \ a n\2" + proof - + obtain B c where "B>0" and event1:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) + \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2" + using assms by auto + from eventually_conj[OF event1 a_large,unfolded eventually_at_top_linorder] + obtain N where "\n\N. (B * b n = c n * a n - c (n + 1) + \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ 2 \ a n" + by fastforce + then show ?thesis using that[of B N c] \B>0\ by auto + qed + define f where "f=(\n. real_of_int (b n) / real_of_int (prod a {..n}))" + define S where "S = (\n. f n)" + have "summable f" + unfolding f_def by (rule aux_series_summable) + define C where "C=B*prod a {..nn. (c (n+N) * a (n+N)) / prod a {N..n+N})" + define h2 where "h2 = (\n. c (n+N+1) / prod a {N..n+N})" + have f_h12:"B*prod a {..n. B * b (n+N))" + define g2 where "g2 = (\n. prod a {.. {N..n + N}) = prod a {.. {N..n + N}) = prod a {..n+N}" + by (simp add: ivl_disj_un_one(4)) + ultimately show ?thesis + unfolding g2_def + apply simp + using a_pos by (metis less_irrefl) + qed + ultimately have "B*prod a {..nn. f (n+N)))" + using suminf_split_initial_segment[OF \summable f\,of N] + unfolding S_def by (auto simp:algebra_simps) + also have "... = C + B*prod a {..n. f (n+N))" + unfolding C_def by (auto simp:algebra_simps) + also have "... = C + (\n. h1 n - h2 n)" + apply (subst suminf_mult[symmetric]) + subgoal using \summable f\ by (simp add: summable_iff_shift) + subgoal using f_h12 by auto + done + also have "... = C + h1 0" + proof - + have "(\n. \i\n. h1 i - h2 i) \ (\i. h1 i - h2 i)" + proof (rule summable_LIMSEQ') + have "(\i. h1 i - h2 i) = (\i. real_of_int (B * prod a {..i. h1 i - h2 i)" + using \summable f\ by (simp add: summable_iff_shift summable_mult) + qed + moreover have "(\i\n. h1 i - h2 i) = h1 0 - h2 n" for n + proof (induct n) + case 0 + then show ?case by simp + next + case (Suc n) + have "(\i\Suc n. h1 i - h2 i) = (\i\n. h1 i - h2 i) + h1 (n+1) - h2 (n+1)" + by auto + also have "... = h1 0 - h2 n + h1 (n+1) - h2 (n+1)" using Suc by auto + also have "... = h1 0 - h2 (n+1)" + proof - + have "h2 n = h1 (n+1)" + unfolding h2_def h1_def + apply (auto simp:prod.nat_ivl_Suc') + using a_pos by (metis less_numeral_extra(3)) + then show ?thesis by auto + qed + finally show ?case by simp + qed + ultimately have "(\n. h1 0 - h2 n) \ (\i. h1 i - h2 i)" by simp + then have "h2 \ (h1 0 - (\i. h1 i - h2 i))" + apply (elim metric_tendsto_imp_tendsto) + by (auto intro!:eventuallyI simp add:dist_real_def) + moreover have "h2 \ 0" + proof - + have h2_n:"\h2 n\ < (1 / 2)^(n+1)" for n + proof - + have "\h2 n\ = \c (n + N + 1)\ / prod a {N..n + N}" + unfolding h2_def abs_divide + using a_pos by (simp add: abs_of_pos prod_pos) + also have "... < (a (N+n) / 2) / prod a {N..n + N}" + unfolding h2_def + apply (rule divide_strict_right_mono) + subgoal using large_n[rule_format,of "N+n"] by (auto simp add:algebra_simps) + subgoal using a_pos by (simp add: prod_pos) + done + also have "... = 1 / (2*prod a {N.. (1/2)^(n+1)" + proof (induct n) + case 0 + then show ?case by auto + next + case (Suc n) + define P where "P=1 / real_of_int (2 * prod a {N.. ( (1 / 2) ^ (n + 1) ) / a (n+N) " + apply (rule divide_right_mono) + subgoal unfolding P_def using Suc by auto + subgoal by (simp add: a_pos less_imp_le) + done + also have "... \ ( (1 / 2) ^ (n + 1) ) / 2 " + apply (rule divide_left_mono) + using large_n[rule_format,of "n+N",simplified] by auto + also have "... = (1 / 2) ^ (n + 2)" by auto + finally show ?case by simp + qed + finally show ?thesis . + qed + have "(\n. (1 / 2)^(n+1)) \ (0::real)" + using tendsto_mult_right_zero[OF LIMSEQ_abs_realpow_zero2[of "1/2",simplified],of "1/2"] + by auto + then show ?thesis + apply (elim Lim_null_comparison[rotated]) + using h2_n less_eq_real_def by (auto intro!:eventuallyI) + qed + ultimately have "(\i. h1 i - h2 i) = h1 0" + using LIMSEQ_unique by fastforce + then show ?thesis by simp + qed + also have "... = C + c N" + unfolding h1_def using a_pos + by auto (metis less_irrefl) + finally show ?thesis . + qed + then have "S = (C + real_of_int (c N)) / (B*prod a {..0 < B\ a_pos less_irrefl mult.commute mult_pos_pos + nonzero_mult_div_cancel_right of_int_eq_0_iff prod_pos) + moreover have "... \ \" + unfolding C_def f_def by (intro Rats_divide Rats_add Rats_mult Rats_of_int Rats_sum) + ultimately show "S \ \" by auto +qed + +theorem theorem_2_1_Erdos_Straus : + "(\n. (b n / (\i \ n. a i))) \ \ \ (\ (B::int)>0. \ c::nat\ int. + (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\The following is a Corollary to Theorem 2.1. \ + +corollary corollary_2_10_Erdos_Straus: + assumes ab_event:"\\<^sub>F n in sequentially. b n > 0 \ a (n+1) \ a n" + and ba_lim_leq:"lim (\n. (b(n+1) - b n )/a n) \ 0" + and ba_lim_exist:"convergent (\n. (b(n+1) - b n )/a n)" + and "liminf (\n. a n / b n) = 0 " + shows "(\n. (b n / (\i \ n. a i))) \ \" +proof + assume "(\n. (b n / (\i \ n. a i))) \ \" + then obtain B c where "B>0" and abc_event:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) + \ \c (n + 1)\ < a n / 2" and ca_vanish: "(\n. c (Suc n) / a n) \ 0" + using ab_rationality_imp by auto + + have bac_close:"(\n. B * b n / a n - c n) \ 0" + proof - + have "\\<^sub>F n in sequentially. B * b n - c n * a n + c (n + 1) = 0" + using abc_event by (auto elim!:eventually_mono) + then have "\\<^sub>F n in sequentially. (B * b n - c n * a n + c (n+1)) / a n = 0 " + apply eventually_elim + by auto + then have "\\<^sub>F n in sequentially. B * b n / a n - c n + c (n + 1) / a n = 0" + apply eventually_elim + using a_pos by (auto simp:divide_simps) (metis less_irrefl) + then have "(\n. B * b n / a n - c n + c (n + 1) / a n) \ 0" + by (simp add: eventually_mono tendsto_iff) + from tendsto_diff[OF this ca_vanish] + show ?thesis by auto + qed + + have c_pos:"\\<^sub>F n in sequentially. c n > 0" + proof - + from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" + apply (elim tendsto_of_int_diff_0) + using ab_event a_large apply (eventually_elim) + using \B>0\ by auto + show ?thesis + proof (rule ccontr) + assume "\ (\\<^sub>F n in sequentially. c n > 0)" + moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" + using * eventually_sequentially_Suc[of "\n. c n\0"] + by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) + ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" + using eventually_elim2 frequently_def by fastforce + moreover have "\\<^sub>F n in sequentially. b n > 0 \ B * b n = c n * a n - c (n + 1)" + using ab_event abc_event by eventually_elim auto + ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 \ b n > 0 + \ B * b n = c n * a n - c (n + 1)" + using frequently_eventually_frequently by fastforce + from frequently_ex[OF this] + obtain n where "c n = 0" "c (Suc n) \ 0" "b n > 0" + "B * b n = c n * a n - c (n + 1)" + by auto + then have "B * b n \ 0" by auto + then show False using \b n>0\ \B > 0\ using mult_pos_pos not_le by blast + qed + qed + + have bc_epsilon:"\\<^sub>F n in sequentially. b (n+1) / b n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real + proof - + have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" + using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto + moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" + apply (subst (asm) eventually_sequentially_Suc[symmetric]) + by simp + moreover have "\\<^sub>F n in sequentially. B * b (n+1) = c (n+1) * a (n+1) - c (n + 2)" + using abc_event + apply (subst (asm) eventually_sequentially_Suc[symmetric]) + by (auto elim:eventually_mono) + moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" + proof - + have "\\<^sub>F n in sequentially. 0 < c (Suc n)" + using c_pos by (subst eventually_sequentially_Suc) simp + moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" + using c_pos by (subst eventually_sequentially_Suc) simp + ultimately show ?thesis using c_pos by eventually_elim auto + qed + ultimately show ?thesis using ab_event abc_event + proof eventually_elim + case (elim n) + define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" + have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" using a_pos elim by (auto simp add: \\<^sub>0_def \\<^sub>1_def) + have "(\ - \\<^sub>1) * c n > 0" + apply (rule mult_pos_pos) + using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto + moreover have "\\<^sub>0 * (c (n+1) - \) > 0" + apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) + using elim(4) that(2) by linarith + ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto + moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith + moreover have "c n > 0" by (simp add: elim(4)) + ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" + by (auto simp add: field_simps) + also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" + proof - + have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" + by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ + divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) + moreover have "(a (n+1) / a n) \ 1" + using a_pos elim(5) by auto + ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) + qed + also have "... = (B * b (n+1)) / (B * b n)" + proof - + have "B * b n = c n * a n - c (n + 1)" + using elim by auto + also have "... = a n * (c n - \\<^sub>0)" + using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) + finally have "B * b n = a n * (c n - \\<^sub>0)" . + moreover have "B * b (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" + unfolding \\<^sub>1_def + using a_pos[rule_format,of "n+1"] + apply (subst \B * b (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) + by (auto simp:field_simps) + ultimately show ?thesis by (simp add: mult.commute) + qed + also have "... = b (n+1) / b n" + using \B>0\ by auto + finally show ?case . + qed + qed + + have eq_2_11:"\\<^sub>F n in sequentially. b (n+1) > b n + (1 - \)^2 * a n / B" + when "\>0" "\<1" "\ (\\<^sub>F n in sequentially. c (n+1) \ c n)" for \::real + proof - + have "\\<^sub>F x in sequentially. c x < c (Suc x) " using that(3) + by (simp add:not_eventually frequently_elim1) + moreover have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" + using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto + moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" + proof - + have "\\<^sub>F n in sequentially. 0 < c (Suc n)" + using c_pos by (subst eventually_sequentially_Suc) simp + then show ?thesis using c_pos by eventually_elim auto + qed + ultimately show ?thesis using ab_event abc_event bc_epsilon[OF \\>0\ \\<1\] + proof (elim frequently_rev_mp,eventually_elim) + case (elim n) + then have "c (n+1) / a n < \" + using a_pos[rule_format,of n] by auto + also have "... \ \ * c n" using elim(7) that(1) by auto + finally have "c (n+1) / a n < \ * c n" . + then have "c (n+1) / c n < \ * a n" + using a_pos[rule_format,of n] elim by (auto simp:field_simps) + then have "(1 - \) * a n < a n - c (n+1) / c n" + by (auto simp:algebra_simps) + then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" + apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) + using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) + then have "b n + (1 - \)^2 * a n / B < b n + (1 - \) * (a n - c (n+1) / c n) / B" + using \B>0\ by auto + also have "... = b n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" + using elim by (auto simp:field_simps) + also have "... = b n + (1 - \) * (b n / c n)" + proof - + have "B * b n = c n * a n - c (n + 1)" using elim by auto + from this[symmetric] show ?thesis + using \B>0\ by simp + qed + also have "... = (1+(1-\)/c n) * b n" + by (auto simp:algebra_simps) + also have "... = ((c n+1-\)/c n) * b n" + using elim by (auto simp:divide_simps) + also have "... \ ((c (n+1) -\)/c n) * b n" + proof - + define cp where "cp = c n+1" + have "c (n+1) \ cp" unfolding cp_def using \c n < c (Suc n)\ by auto + moreover have "c n>0" "b n>0" using elim by auto + ultimately show ?thesis + apply (fold cp_def) + by (auto simp:divide_simps) + qed + also have "... < b (n+1)" + using elim by (auto simp:divide_simps) + finally show ?case . + qed + qed + + have "\\<^sub>F n in sequentially. c (n+1) \ c n" + proof (rule ccontr) + assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" + from eq_2_11[OF _ _ this,of "1/2"] + have "\\<^sub>F n in sequentially. b (n+1) > b n + 1/4 * a n / B" + by (auto simp:algebra_simps power2_eq_square) + then have *:"\\<^sub>F n in sequentially. (b (n+1) - b n) / a n > 1 / (B * 4)" + apply (elim frequently_elim1) + subgoal for n + using a_pos[rule_format,of n] by (auto simp:field_simps) + done + define f where "f = (\n. (b (n+1) - b n) / a n)" + have "f \ lim f" + using convergent_LIMSEQ_iff ba_lim_exist unfolding f_def by auto + from this[unfolded tendsto_iff,rule_format, of "1 / (B*4)"] + have "\\<^sub>F x in sequentially. \f x - lim f\ < 1 / (B * 4)" + using \B>0\ by (auto simp:dist_real_def) + moreover have "\\<^sub>F n in sequentially. f n > 1 / (B * 4)" + using * unfolding f_def by auto + ultimately have "\\<^sub>F n in sequentially. f n > 1 / (B * 4) \ \f n - lim f\ < 1 / (B * 4)" + by (auto elim:frequently_eventually_frequently[rotated]) + from frequently_ex[OF this] + obtain n where "f n > 1 / (B * 4)" "\f n - lim f\ < 1 / (B * 4)" + by auto + moreover have "lim f \ 0" using ba_lim_leq unfolding f_def by auto + ultimately show False by linarith + qed + then obtain N where N_dec:"\n\N. c (n+1) \ c n" by (meson eventually_at_top_linorder) + define max_c where "max_c = (MAX n \ {..N}. c n)" + have max_c:"c n \ max_c" for n + proof (cases "n\N") + case True + then show ?thesis unfolding max_c_def by simp + next + case False + then have "n\N" by auto + then have "c n\c N" + proof (induct rule:nat_induct_at_least) + case base + then show ?case by simp + next + case (Suc n) + then have "c (n+1) \ c n" using N_dec by auto + then show ?case using \c n \ c N\ by auto + qed + moreover have "c N \ max_c" unfolding max_c_def by auto + ultimately show ?thesis by auto + qed + have "max_c > 0 " + proof - + obtain N where "\n\N. 0 < c n" + using c_pos[unfolded eventually_at_top_linorder] by auto + then have "c N > 0" by auto + then show ?thesis using max_c[of N] by simp + qed + have ba_limsup_bound:"1/(B*(B+1)) \ limsup (\n. b n/a n)" + "limsup (\n. b n/a n) \ max_c / B + 1 / (B+1)" + proof - + define f where "f = (\n. b n/a n)" + from tendsto_mult_right_zero[OF bac_close,of "1/B"] + have "(\n. f n - c n / B) \ 0" + unfolding f_def using \B>0\ by (auto simp:algebra_simps) + from this[unfolded tendsto_iff,rule_format,of "1/(B+1)"] + have "\\<^sub>F x in sequentially. \f x - c x / B\ < 1 / (B+1)" + using \B>0\ by auto + then have *:"\\<^sub>F n in sequentially. 1/(B*(B+1)) \ ereal (f n) \ ereal (f n) \ max_c / B + 1 / (B+1)" + using c_pos + proof eventually_elim + case (elim n) + then have "f n - c n / B < 1 / (B+1)" by auto + then have "f n < c n / B + 1 / (B+1)" by simp + also have "... \ max_c / B + 1 / (B+1)" + using max_c[of n] using \B>0\ by (auto simp:divide_simps) + finally have *:"f n < max_c / B + 1 / (B+1)" . + + have "1/(B*(B+1)) = 1/B - 1 / (B+1)" + using \B>0\ by (auto simp:divide_simps) + also have "... \ c n/B - 1 / (B+1)" + using \0 < c n\ \B>0\ by (auto,auto simp:divide_simps) + also have "... < f n" using elim by auto + finally have "1/(B*(B+1)) < f n" . + with * show ?case by simp + qed + show "limsup f \ max_c / B + 1 / (B+1)" + apply (rule Limsup_bounded) + using * by (auto elim:eventually_mono) + have "1/(B*(B+1)) \ liminf f" + apply (rule Liminf_bounded) + using * by (auto elim:eventually_mono) + also have "liminf f \ limsup f" by (simp add: Liminf_le_Limsup) + finally show "1/(B*(B+1)) \ limsup f" . + qed + + have "0 < inverse (ereal (max_c / B + 1 / (B+1)))" + using \max_c > 0\ \B>0\ + by (simp add: pos_add_strict) + also have "... \ inverse (limsup (\n. b n/a n))" + proof (rule ereal_inverse_antimono[OF _ ba_limsup_bound(2)]) + have "0<1/(B*(B+1))" using \B>0\ by auto + also have "... \ limsup (\n. b n/a n)" using ba_limsup_bound(1) . + finally show "0\limsup (\n. b n/a n)" using zero_ereal_def by auto + qed + also have "... = liminf (\n. inverse (ereal ( b n/a n)))" + apply (subst Liminf_inverse_ereal[symmetric]) + using a_pos ab_event by (auto elim!:eventually_mono simp:divide_simps) + also have "... = liminf (\n. ( a n/b n))" + apply (rule Liminf_eq) + using a_pos ab_event + apply (auto elim!:eventually_mono) + by (metis less_int_code(1)) + finally have "liminf (\n. ( a n/b n)) > 0" . + then show False using \liminf (\n. a n / b n) = 0\ by simp +qed + +end + +section\Some auxiliary results on the prime numbers. \ + +lemma nth_prime_nonzero[simp]:"nth_prime n \ 0" + by (simp add: prime_gt_0_nat prime_nth_prime) + +lemma nth_prime_gt_zero[simp]:"nth_prime n >0" + by (simp add: prime_gt_0_nat prime_nth_prime) + +lemma ratio_of_consecutive_primes: + "(\n. nth_prime (n+1)/nth_prime n) \1" +proof - + define f where "f=(\x. real (nth_prime (Suc x)) /real (nth_prime x))" + define g where "g=(\x. (real x * ln (real x)) + / (real (Suc x) * ln (real (Suc x))))" + have p_n:"(\x. real (nth_prime x) / (real x * ln (real x))) \ 1" + using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified] . + moreover have p_sn:"(\n. real (nth_prime (Suc n)) + / (real (Suc n) * ln (real (Suc n)))) \ 1" + using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified + ,THEN LIMSEQ_Suc] . + ultimately have "(\x. f x * g x) \ 1" + using tendsto_divide[OF p_sn p_n] + unfolding f_def g_def by (auto simp:algebra_simps) + moreover have "g \ 1" unfolding g_def + by real_asymp + ultimately have "(\x. if g x = 0 then 0 else f x) \ 1" + apply (drule_tac tendsto_divide[OF _ \g \ 1\]) + by auto + then have "f \ 1" + proof (elim filterlim_mono_eventually) + have "\\<^sub>F x in sequentially. (if g (x+3) = 0 then 0 + else f (x+3)) = f (x+3)" + unfolding g_def by auto + then show "\\<^sub>F x in sequentially. (if g x = 0 then 0 else f x) = f x" + apply (subst (asm) eventually_sequentially_seg) + by simp + qed auto + then show ?thesis unfolding f_def by auto +qed + +lemma nth_prime_double_sqrt_less: + assumes "\ > 0" + shows "\\<^sub>F n in sequentially. (nth_prime (2*n) - nth_prime n) + / sqrt (nth_prime n) < n powr (1/2+\)" +proof - + define pp ll where + "pp=(\n. (nth_prime (2*n) - nth_prime n) / sqrt (nth_prime n))" and + "ll=(\x::nat. x * ln x)" + have pp_pos:"pp (n+1) > 0" for n + unfolding pp_def by simp + + have "(\x. nth_prime (2 * x)) \[sequentially] (\x. (2 * x) * ln (2 * x))" + using nth_prime_asymptotics[THEN asymp_equiv_compose + ,of "(*) 2" sequentially,unfolded comp_def] + using mult_nat_left_at_top pos2 by blast + also have "... \[sequentially] (\x. 2 *x * ln x)" + by real_asymp + finally have "(\x. nth_prime (2 * x)) \[sequentially] (\x. 2 *x * ln x)" . + from this[unfolded asymp_equiv_def, THEN tendsto_mult_left,of 2] + have "(\x. nth_prime (2 * x) / (x * ln x)) \ 2" + unfolding asymp_equiv_def by auto + moreover have *:"(\x. nth_prime x / (x * ln x)) \ 1" + using nth_prime_asymptotics unfolding asymp_equiv_def by auto + ultimately + have "(\x. (nth_prime (2 * x) - nth_prime x) / ll x) \ 1" + unfolding ll_def + apply - + apply (drule (1) tendsto_diff) + apply (subst of_nat_diff,simp) + by (subst diff_divide_distrib,simp) + moreover have "(\x. sqrt (nth_prime x) / sqrt (ll x)) \ 1" + unfolding ll_def + using tendsto_real_sqrt[OF *] + by (auto simp: real_sqrt_divide) + ultimately have "(\x. pp x * (sqrt (ll x) / (ll x))) \ 1" + apply - + apply (drule (1) tendsto_divide,simp) + by (auto simp:field_simps of_nat_diff pp_def) + moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / ll x = 1/sqrt (ll x)" + apply (subst eventually_sequentially_Suc[symmetric]) + by (auto intro!:eventuallyI simp:ll_def divide_simps) + ultimately have "(\x. pp x / sqrt (ll x)) \ 1" + apply (elim filterlim_mono_eventually) + by (auto elim!:eventually_mono) (metis mult.right_neutral times_divide_eq_right) + moreover have "(\x. sqrt (ll x) / x powr (1/2+\)) \ 0" + unfolding ll_def using \\>0\ by real_asymp + ultimately have "(\x. pp x / x powr (1/2+\) * + (sqrt (ll x) / sqrt (ll x))) \ 0" + apply - + apply (drule (1) tendsto_mult) + by (auto elim:filterlim_mono_eventually) + moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / sqrt (ll x) = 1" + apply (subst eventually_sequentially_Suc[symmetric]) + by (auto intro!:eventuallyI simp:ll_def ) + ultimately have "(\x. pp x / x powr (1/2+\)) \ 0" + apply (elim filterlim_mono_eventually) + by (auto elim:eventually_mono) + from tendstoD[OF this, of 1,simplified] + show "\\<^sub>F x in sequentially. pp x < x powr (1 / 2 + \)" + apply (elim eventually_mono_sequentially[of _ 1]) + using pp_pos by auto +qed + + +section \Theorem 3.1\ + +text\Theorem 3.1 is an application of Theorem 2.1 with the sequences considered involving +the prime numbers.\ + +theorem theorem_3_10_Erdos_Straus: + fixes a::"nat \ int" + assumes a_pos:"\ n. a n >0" and "mono a" + and nth_1:"(\n. nth_prime n / (a n)^2) \ 0" + and nth_2:"liminf (\n. a n / nth_prime n) = 0" + shows "(\n. (nth_prime n / (\i \ n. a i))) \ \" +proof + assume asm:"(\n. (nth_prime n / (\i \ n. a i))) \ \" + + have a2_omega:"(\n. (a n)^2) \ \(\x. x * ln x)" + proof - + have "(\n. real (nth_prime n)) \ o(\n. real_of_int ((a n)\<^sup>2))" + apply (rule smalloI_tendsto[OF nth_1]) + using a_pos by (metis (mono_tags, lifting) less_int_code(1) + not_eventuallyD of_int_0_eq_iff zero_eq_power2) + moreover have "(\x. real (nth_prime x)) \ \(\x. real x * ln (real x))" + using nth_prime_bigtheta + by blast + ultimately show ?thesis + using landau_omega.small_big_trans smallo_imp_smallomega by blast + qed + + have a_gt_1:"\\<^sub>F n in sequentially. 1 < a n" + proof - + have "\\<^sub>F x in sequentially. \x * ln x\ \ (a x)\<^sup>2" + using a2_omega[unfolded smallomega_def,simplified,rule_format,of 1] + by auto + then have "\\<^sub>F x in sequentially. \(x+3) * ln (x+3)\ \ (a (x+3))\<^sup>2" + apply (subst (asm) eventually_sequentially_seg[symmetric, of _ 3]) + by simp + then have "\\<^sub>F n in sequentially. 1 < a ( n+3)" + proof (elim eventually_mono) + fix x + assume "\real (x + 3) * ln (real (x + 3))\ \ real_of_int ((a (x + 3))\<^sup>2)" + moreover have "\real (x + 3) * ln (real (x + 3))\ > 3" + proof - + have "ln (real (x + 3)) > 1" + apply simp using ln3_gt_1 ln_gt_1 by force + moreover have "real(x+3) \ 3" by simp + ultimately have "(x+3)*ln (real (x + 3)) > 3*1 " + apply (rule_tac mult_le_less_imp_less) + by auto + then show ?thesis by auto + qed + ultimately have "real_of_int ((a (x + 3))\<^sup>2) > 3" + by auto + then show "1 < a (x + 3)" + by (smt Suc3_eq_add_3 a_pos add.commute of_int_1 one_power2) + qed + then show ?thesis + apply (subst eventually_sequentially_seg[symmetric, of _ 3]) + by auto + qed + + obtain B::int and c where + "B>0" and Bc_large:"\\<^sub>F n in sequentially. B * nth_prime n + = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" + and ca_vanish: "(\n. c (Suc n) / real_of_int (a n)) \ 0" + proof - + note a_gt_1 + moreover have "(\n. real_of_int \int (nth_prime n)\ + / real_of_int (a (n - 1) * a n)) \ 0" + proof - + define f where "f=(\n. nth_prime (n+1) / (a n * a (n+1)))" + define g where "g=(\n. 2*nth_prime n / (a n)^2)" + have "\\<^sub>F x in sequentially. norm (f x) \ g x" + proof - + have "\\<^sub>F n in sequentially. nth_prime (n+1) < 2*nth_prime n" + using ratio_of_consecutive_primes[unfolded tendsto_iff + ,rule_format,of 1,simplified] + apply (elim eventually_mono) + by (auto simp :divide_simps dist_norm) + moreover have "\\<^sub>F n in sequentially. real_of_int (a n * a (n+1)) + \ (a n)^2" + apply (rule eventuallyI) + using \mono a\ by (auto simp:power2_eq_square a_pos incseq_SucD) + ultimately show ?thesis unfolding f_def g_def + apply eventually_elim + apply (subst norm_divide) + apply (rule_tac linordered_field_class.frac_le) + using a_pos[rule_format, THEN order.strict_implies_not_eq ] + by auto + qed + moreover have "g \ 0 " + using nth_1[THEN tendsto_mult_right_zero,of 2] unfolding g_def + by auto + ultimately have "f \ 0" + using Lim_null_comparison[of f g sequentially] + by auto + then show ?thesis + unfolding f_def + by (rule_tac LIMSEQ_imp_Suc) auto + qed + moreover have "(\n. real_of_int (int (nth_prime n)) + / real_of_int (prod a {..n})) \ \" + using asm by simp + ultimately have "\B>0. \c. (\\<^sub>F n in sequentially. + B * int (nth_prime n) = c n * a n - c (n + 1) \ + real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ + (\n. real_of_int (c (Suc n)) / real_of_int (a n)) \ 0" + using ab_rationality_imp[OF a_pos,of nth_prime] by fast + then show thesis + apply clarify + apply (rule_tac c=c and B=B in that) + by auto + qed + + have bac_close:"(\n. B * nth_prime n / a n - c n) \ 0" + proof - + have "\\<^sub>F n in sequentially. B * nth_prime n - c n * a n + c (n + 1) = 0" + using Bc_large by (auto elim!:eventually_mono) + then have "\\<^sub>F n in sequentially. (B * nth_prime n - c n * a n + c (n+1)) / a n = 0 " + apply eventually_elim + by auto + then have "\\<^sub>F n in sequentially. B * nth_prime n / a n - c n + c (n + 1) / a n = 0" + apply eventually_elim + using a_pos by (auto simp:divide_simps) (metis less_irrefl) + then have "(\n. B * nth_prime n / a n - c n + c (n + 1) / a n) \ 0" + by (simp add: eventually_mono tendsto_iff) + from tendsto_diff[OF this ca_vanish] + show ?thesis by auto + qed + + have c_pos:"\\<^sub>F n in sequentially. c n > 0" + proof - + from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" + apply (elim tendsto_of_int_diff_0) + using a_gt_1 apply (eventually_elim) + using \B>0\ by auto + show ?thesis + proof (rule ccontr) + assume "\ (\\<^sub>F n in sequentially. c n > 0)" + moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" + using * eventually_sequentially_Suc[of "\n. c n\0"] + by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) + ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" + using eventually_elim2 frequently_def by fastforce + moreover have "\\<^sub>F n in sequentially. nth_prime n > 0 + \ B * nth_prime n = c n * a n - c (n + 1)" + using Bc_large by eventually_elim auto + ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 + \ B * nth_prime n = c n * a n - c (n + 1)" + using frequently_eventually_frequently by fastforce + from frequently_ex[OF this] + obtain n where "c n = 0" "c (Suc n) \ 0" + "B * nth_prime n = c n * a n - c (n + 1)" + by auto + then have "B * nth_prime n \ 0" by auto + then show False using \B > 0\ + by (simp add: mult_le_0_iff) + qed + qed + + have B_nth_prime:"\\<^sub>F n in sequentially. nth_prime n > B" + proof - + have "\\<^sub>F x in sequentially. B+1 \ nth_prime x" + using nth_prime_at_top[unfolded filterlim_at_top_ge[where c="nat B+1"] + ,rule_format,of "nat B + 1",simplified] + + apply (elim eventually_mono) + using \B>0\ by auto + then show ?thesis + apply (elim eventually_mono) + by auto + qed + + have bc_epsilon:"\\<^sub>F n in sequentially. nth_prime (n+1) + / nth_prime n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real + proof - + have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" + using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto + moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" + apply (subst (asm) eventually_sequentially_Suc[symmetric]) + by simp + moreover have "\\<^sub>F n in sequentially. B * nth_prime (n+1) = c (n+1) * a (n+1) - c (n + 2)" + using Bc_large + apply (subst (asm) eventually_sequentially_Suc[symmetric]) + by (auto elim:eventually_mono) + moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" + proof - + have "\\<^sub>F n in sequentially. 0 < c (Suc n)" + using c_pos by (subst eventually_sequentially_Suc) simp + moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" + using c_pos by (subst eventually_sequentially_Suc) simp + ultimately show ?thesis using c_pos by eventually_elim auto + qed + ultimately show ?thesis using Bc_large + proof eventually_elim + case (elim n) + define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" + have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" + using a_pos elim \mono a\ + by (auto simp add: \\<^sub>0_def \\<^sub>1_def abs_of_pos) + have "(\ - \\<^sub>1) * c n > 0" + apply (rule mult_pos_pos) + using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto + moreover have "\\<^sub>0 * (c (n+1) - \) > 0" + apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) + using elim(4) that(2) by linarith + ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto + moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith + moreover have "c n > 0" by (simp add: elim(4)) + ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" + by (auto simp add:field_simps) + also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" + proof - + have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" + by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ + divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) + moreover have "(a (n+1) / a n) \ 1" + using a_pos \mono a\ by (simp add: mono_def) + ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) + qed + also have "... = (B * nth_prime (n+1)) / (B * nth_prime n)" + proof - + have "B * nth_prime n = c n * a n - c (n + 1)" + using elim by auto + also have "... = a n * (c n - \\<^sub>0)" + using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) + finally have "B * nth_prime n = a n * (c n - \\<^sub>0)" . + moreover have "B * nth_prime (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" + unfolding \\<^sub>1_def + using a_pos[rule_format,of "n+1"] + apply (subst \B * nth_prime (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) + by (auto simp:field_simps) + ultimately show ?thesis by (simp add: mult.commute) + qed + also have "... = nth_prime (n+1) / nth_prime n" + using \B>0\ by auto + finally show ?case . + qed + qed + + + have c_ubound:"\x. \n. c n > x" + proof (rule ccontr) + assume " \ (\x. \n. x < c n)" + then obtain ub where "\n. c n \ ub" "ub > 0" + by (meson dual_order.trans int_one_le_iff_zero_less le_cases not_le) + define pa where "pa = (\n. nth_prime n / a n)" + have pa_pos:"\n. pa n > 0" unfolding pa_def by (simp add: a_pos) + have "liminf (\n. 1 / pa n) = 0" + using nth_2 unfolding pa_def by auto + then have "(\y\<^sub>F x in sequentially. ereal (1 / pa x) \ y)" + apply (subst less_Liminf_iff[symmetric]) + using \0 < B\ \0 < ub\ by auto + then have "\\<^sub>F x in sequentially. 1 / pa x < B/(ub+1)" + by (meson frequently_mono le_less_trans less_ereal.simps(1)) + then have "\\<^sub>F x in sequentially. B*pa x > (ub+1)" + apply (elim frequently_elim1) + by (metis \0 < ub\ mult.left_neutral of_int_0_less_iff pa_pos pos_divide_less_eq + pos_less_divide_eq times_divide_eq_left zless_add1_eq) + moreover have "\\<^sub>F x in sequentially. c x \ ub" + using \\n. c n \ ub\ by simp + ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1" + apply (elim frequently_rev_mp eventually_mono) + by linarith + moreover have "(\n. B * pa n - c n) \0" + unfolding pa_def using bac_close by auto + from tendstoD[OF this,of 1] + have "\\<^sub>F n in sequentially. \B * pa n - c n\ < 1" + by auto + ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1 \ \B * pa x - c x\ < 1" + using frequently_eventually_frequently by blast + then show False + by (simp add: frequently_def) + qed + + have eq_2_11:"\\<^sub>F n in sequentially. c (n+1)>c n \ + nth_prime (n+1) > nth_prime n + (1 - \)^2 * a n / B" + when "\>0" "\<1" for \::real + proof - + have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" + using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto + moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" + proof - + have "\\<^sub>F n in sequentially. 0 < c (Suc n)" + using c_pos by (subst eventually_sequentially_Suc) simp + then show ?thesis using c_pos by eventually_elim auto + qed + ultimately show ?thesis using Bc_large bc_epsilon[OF \\>0\ \\<1\] + proof (eventually_elim, rule_tac impI) + case (elim n) + assume "c n < c (n + 1)" + have "c (n+1) / a n < \" + using a_pos[rule_format,of n] using elim(1,2) by auto + also have "... \ \ * c n" using elim(2) that(1) by auto + finally have "c (n+1) / a n < \ * c n" . + then have "c (n+1) / c n < \ * a n" + using a_pos[rule_format,of n] elim by (auto simp:field_simps) + then have "(1 - \) * a n < a n - c (n+1) / c n" + by (auto simp:algebra_simps) + then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" + apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) + using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) + then have "nth_prime n + (1 - \)^2 * a n / B < nth_prime n + (1 - \) * (a n - c (n+1) / c n) / B" + using \B>0\ by auto + also have "... = nth_prime n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" + using elim by (auto simp:field_simps) + also have "... = nth_prime n + (1 - \) * (nth_prime n / c n)" + proof - + have "B * nth_prime n = c n * a n - c (n + 1)" using elim by auto + from this[symmetric] show ?thesis + using \B>0\ by simp + qed + also have "... = (1+(1-\)/c n) * nth_prime n" + by (auto simp:algebra_simps) + also have "... = ((c n+1-\)/c n) * nth_prime n" + using elim by (auto simp:divide_simps) + also have "... \ ((c (n+1) -\)/c n) * nth_prime n" + proof - + define cp where "cp = c n+1" + have "c (n+1) \ cp" unfolding cp_def using \c n < c (n + 1)\ by auto + moreover have "c n>0" "nth_prime n>0" using elim by auto + ultimately show ?thesis + apply (fold cp_def) + by (auto simp:divide_simps) + qed + also have "... < nth_prime (n+1)" + using elim by (auto simp:divide_simps) + finally show "real (nth_prime n) + (1 - \)\<^sup>2 * real_of_int (a n) + / real_of_int B < real (nth_prime (n + 1))" . + qed + qed + + have c_neq_large:"\\<^sub>F n in sequentially. c (n+1) \ c n" + proof (rule ccontr) + assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" + then have that:"\\<^sub>F n in sequentially. c (n + 1) = c n" + unfolding frequently_def . + have "\\<^sub>F x in sequentially. (B * int (nth_prime x) = c x * a x - c (x + 1) + \ \real_of_int (c (x + 1))\ < real_of_int (a x) / 2) \ 0 < c x \ B < int (nth_prime x) + \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" + using Bc_large c_pos B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] + by eventually_elim (auto simp:divide_simps) + then have "\\<^sub>F m in sequentially. nth_prime (m+1) > (1+1/(2*B))*nth_prime m" + proof (elim frequently_eventually_at_top[OF that, THEN frequently_at_top_elim]) + fix n + assume "c (n + 1) = c n \ + (\y\n. (B * int (nth_prime y) = c y * a y - c (y + 1) \ + \real_of_int (c (y + 1))\ < real_of_int (a y) / 2) \ + 0 < c y \ B < int (nth_prime y) \ (c y < c (y + 1) \ + real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) + < real (nth_prime (y + 1))))" + then have "c (n + 1) = c n" + and Bc_eq:"\y\n. B * int (nth_prime y) = c y * a y - c (y + 1) \ 0 < c y + \ \real_of_int (c (y + 1))\ < real_of_int (a y) / 2 + \ B < int (nth_prime y) + \ (c y < c (y + 1) \ + real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) + < real (nth_prime (y + 1)))" + by auto + obtain m where "n c n" "c nN. N > n \ c N > c n" + using c_ubound[rule_format, of "MAX x\{..n}. c x"] + by (metis Max_ge atMost_iff dual_order.trans finite_atMost finite_imageI image_eqI + linorder_not_le order_refl) + then obtain N where "N>n" "c N>c n" by auto + define A m where "A={m. n (m+1)\N \ c (m+1) > c n}" and "m = Min A" + have "finite A" unfolding A_def + by (metis (no_types, lifting) A_def add_leE finite_nat_set_iff_bounded_le mem_Collect_eq) + moreover have "N-1\A" unfolding A_def + using \c n < c N\ \n < N\ \c (n + 1) = c n\ + by (smt Suc_diff_Suc Suc_eq_plus1 Suc_leI Suc_pred add.commute + add_diff_inverse_nat add_leD1 diff_is_0_eq' mem_Collect_eq nat_add_left_cancel_less + zero_less_one) + ultimately have "m\A" + using Min_in unfolding m_def by auto + then have "n0" + unfolding m_def A_def by auto + moreover have "c m \ c n" + proof (rule ccontr) + assume " \ c m \ c n" + then have "m-1\A" using \m\A\ \c (n + 1) = c n\ + unfolding A_def + by auto (smt One_nat_def Suc_eq_plus1 Suc_lessI less_diff_conv) + from Min_le[OF \finite A\ this,folded m_def] \m>0\ show False by auto + qed + ultimately show ?thesis using that[of m] by auto + qed + have "(1 + 1 / (2 * B)) * nth_prime m < nth_prime m + a m / (2*B)" + proof - + have "nth_prime m < a m" + proof - + have "B * int (nth_prime m) < c m * (a m - 1)" + using Bc_eq[rule_format,of m] \c m \ c n\ \c n < c (m + 1)\ \n < m\ + by (auto simp:algebra_simps) + also have "... \ c n * (a m - 1)" + by (simp add: \c m \ c n\ a_pos mult_right_mono) + finally have "B * int (nth_prime m) < c n * (a m - 1)" . + moreover have "c n\B" + proof - + have " B * int (nth_prime n) = c n * (a n - 1)" "B < int (nth_prime n)" + and c_a:"\real_of_int (c (n + 1))\ < real_of_int (a n) / 2" + using Bc_eq[rule_format,of n] \c (n + 1) = c n\ by (auto simp:algebra_simps) + from this(1) have " c n dvd (B * int (nth_prime n))" + by simp + moreover have "coprime (c n) (int (nth_prime n))" + proof - + have "c n < int (nth_prime n)" + proof (rule ccontr) + assume "\ c n < int (nth_prime n)" + then have asm:"c n \ int (nth_prime n)" by auto + then have "a n > 2 * nth_prime n" + using c_a \c (n + 1) = c n\ by auto + then have "a n -1 \ 2 * nth_prime n" + by simp + then have "a n - 1 > 2 * B" + using \B < int (nth_prime n)\ by auto + from mult_le_less_imp_less[OF asm this] \B>0\ + have "int (nth_prime n) * (2 * B) < c n * (a n - 1)" + by auto + then show False using \B * int (nth_prime n) = c n * (a n - 1)\ + by (smt \0 < B\ \B < int (nth_prime n)\ combine_common_factor + mult.commute mult_pos_pos) + qed + then have "\ nth_prime n dvd c n" + by (simp add: Bc_eq zdvd_not_zless) + then have "coprime (int (nth_prime n)) (c n)" + by (auto intro!:prime_imp_coprime_int) + then show ?thesis using coprime_commute by blast + qed + ultimately have "c n dvd B" + using coprime_dvd_mult_left_iff by auto + then show ?thesis using \0 < B\ zdvd_imp_le by blast + qed + moreover have "c n > 0 " using Bc_eq by blast + ultimately show ?thesis + using \B>0\ by (smt a_pos mult_mono) + qed + then show ?thesis using \B>0\ by (auto simp:field_simps) + qed + also have "... < nth_prime (m+1)" + using Bc_eq[rule_format, of m] \n \c m \ c n\ \c n < c (m+1)\ + by linarith + finally show "\j>n. (1 + 1 / real_of_int (2 * B)) * real (nth_prime j) + < real (nth_prime (j + 1))" using \m>n\ by auto + qed + then have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m > (1+1/(2*B))" + by (auto elim:frequently_elim1 simp:field_simps) + moreover have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m < (1+1/(2*B))" + using ratio_of_consecutive_primes[unfolded tendsto_iff,rule_format,of "1/(2*B)"] + \B>0\ + unfolding dist_real_def + by (auto elim!:eventually_mono simp:algebra_simps) + ultimately show False by (simp add: eventually_mono frequently_def) + qed + + have c_gt_half:"\\<^sub>F N in sequentially. card {n\{N..<2*N}. c n > c (n+1)} > N / 2" + proof - + define h where "h=(\n. (nth_prime (2*n) - nth_prime n) + / sqrt (nth_prime n))" + have "\\<^sub>F n in sequentially. h n < n / 2" + proof - + have "\\<^sub>F n in sequentially. h n < n powr (5/6)" + using nth_prime_double_sqrt_less[of "1/3"] + unfolding h_def by auto + moreover have "\\<^sub>F n in sequentially. n powr (5/6) < (n /2)" + by real_asymp + ultimately show ?thesis + by eventually_elim auto + qed + moreover have "\\<^sub>F n in sequentially. sqrt (nth_prime n) / a n < 1 / (2*B)" + using nth_1[THEN tendsto_real_sqrt,unfolded tendsto_iff + ,rule_format,of "1/(2*B)"] \B>0\ a_pos + by (auto simp:real_sqrt_divide abs_of_pos) + ultimately have "\\<^sub>F x in sequentially. c (x+1) \ c x + \ sqrt (nth_prime x) / a x < 1 / (2*B) + \ h x < x / 2 + \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" + using c_neq_large B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] + by eventually_elim (auto simp:divide_simps) + then show ?thesis + proof (elim eventually_at_top_mono) + fix N assume "N\1" and N_asm:"\y\N. c (y + 1) \ c y \ + sqrt (real (nth_prime y)) / real_of_int (a y) + < 1 / real_of_int (2 * B) \ h y < y / 2 \ + (c y < c (y + 1) \ + real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) + < real (nth_prime (y + 1)))" + + define S where "S={n \ {N..<2 * N}. c n < c (n + 1)}" + define g where "g=(\n. (nth_prime (n+1) - nth_prime n) + / sqrt (nth_prime n))" + define f where "f=(\n. nth_prime (n+1) - nth_prime n)" + have g_gt_1:"g n>1" when "n\N" "c n < c (n + 1)" for n + proof - + have "nth_prime n + sqrt (nth_prime n) < nth_prime (n+1)" + proof - + have "nth_prime n + sqrt (nth_prime n) < nth_prime n + a n / (2*B)" + using N_asm[rule_format,OF \n\N\] a_pos + by (auto simp:field_simps) + also have "... < nth_prime (n+1)" + using N_asm[rule_format,OF \n\N\] \c n < c (n + 1)\ by auto + finally show ?thesis . + qed + then show ?thesis unfolding g_def + using \c n < c (n + 1)\ by auto + qed + have g_geq_0:"g n \ 0" for n + unfolding g_def by auto + + have "finite S" "\x\S. x\N \ c x sum g S" + proof (induct S) + case empty + then show ?case by auto + next + case (insert x F) + moreover have "g x>1" + proof - + have "c x < c (x+1)" "x\N" using insert(4) by auto + then show ?thesis using g_gt_1 by auto + qed + ultimately show ?case by simp + qed + also have "... \ sum g {N..<2*N}" + apply (rule sum_mono2) + unfolding S_def using g_geq_0 by auto + also have "... \ sum (\n. f n/sqrt (nth_prime N)) {N..<2*N}" + apply (rule sum_mono) + unfolding f_def g_def by (auto intro!:divide_left_mono) + also have "... = sum f {N..<2*N} / sqrt (nth_prime N)" + unfolding sum_divide_distrib[symmetric] by auto + also have "... = (nth_prime (2*N) - nth_prime N) / sqrt (nth_prime N)" + proof - + have "sum f {N..<2 * N} = nth_prime (2 * N) - nth_prime N" + proof (induct N) + case 0 + then show ?case by simp + next + case (Suc N) + have ?case if "N=0" + proof - + have "sum f {Suc N..<2 * Suc N} = sum f {1}" + using that by (simp add: numeral_2_eq_2) + also have "... = nth_prime 2 - nth_prime 1" + unfolding f_def by (simp add:numeral_2_eq_2) + also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" + using that by auto + finally show ?thesis . + qed + moreover have ?case if "N\0" + proof - + have "sum f {Suc N..<2 * Suc N} = sum f {N..<2 * Suc N} - f N" + apply (subst (2) sum.atLeast_Suc_lessThan) + using that by auto + also have "... = sum f {N..<2 * N}+ f (2*N) + f(2*N+1) - f N" + by auto + also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" + using Suc unfolding f_def by auto + finally show ?thesis . + qed + ultimately show ?case by blast + qed + then show ?thesis by auto + qed + also have "... = h N" + unfolding h_def by auto + also have "... < N/2" + using N_asm by auto + finally have "card S < N/2" . + + define T where "T={n \ {N..<2 * N}. c n > c (n + 1)}" + have "T \ S = {N..<2 * N}" "T \ S = {}" "finite T" + unfolding T_def S_def using N_asm by fastforce+ + + then have "card T + card S = card {N..<2 * N}" + using card_Un_disjoint \finite S\ by metis + also have "... = N" + by simp + finally have "card T + card S = N" . + with \card S < N/2\ + show "card T > N/2" by linarith + qed + qed + + text\Inequality (3.5) in the original paper required a slight modification: \ + + have a_gt_plus:"\\<^sub>F n in sequentially. c n > c (n+1) \a (n+1) > a n + (a n - c(n+1) - 1) / c (n+1)" + proof - + note a_gt_1[THEN eventually_all_ge_at_top] c_pos[THEN eventually_all_ge_at_top] + moreover have "\\<^sub>F n in sequentially. + B * int (nth_prime (n+1)) = c (n+1) * a (n+1) - c (n + 2)" + using Bc_large + apply (subst (asm) eventually_sequentially_Suc[symmetric]) + by (auto elim:eventually_mono) + moreover have "\\<^sub>F n in sequentially. + B * int (nth_prime n) = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" + using Bc_large by (auto elim:eventually_mono) + ultimately show ?thesis + apply (eventually_elim) + proof (rule impI) + fix n + assume "\y\n. 1 < a y" "\y\n. 0 < c y" + and + Suc_n_eq:"B * int (nth_prime (n + 1)) = c (n + 1) * a (n + 1) - c (n + 2)" and + "B * int (nth_prime n) = c n * a n - c (n + 1) \ + real_of_int \c (n + 1)\ < real_of_int (a n) / 2" + and "c (n + 1) < c n" + then have n_eq:"B * int (nth_prime n) = c n * a n - c (n + 1)" and + c_less_a: "real_of_int \c (n + 1)\ < real_of_int (a n) / 2" + by auto + from \\y\n. 1 < a y\ \\y\n. 0 < c y\ + have *:"a n>1" "a (n+1) > 1" "c n > 0" + "c (n+1) > 0" "c (n+2) > 0" + by auto + then have "(1+1/c (n+1))* (a n - 1)/a (n+1) = (c (n+1)+1) * ((a n - 1) / (c (n+1) * a (n+1)))" + by (auto simp:field_simps) + also have "... \ c n * ((a n - 1) / (c (n+1) * a (n+1)))" + apply (rule mult_right_mono) + subgoal using \c (n + 1) < c n\ by auto + subgoal by (smt \0 < c (n + 1)\ a_pos divide_nonneg_pos mult_pos_pos of_int_0_le_iff + of_int_0_less_iff) + done + also have "... = (c n * (a n - 1)) / (c (n+1) * a (n+1))" by auto + also have "... < (c n * (a n - 1)) / (c (n+1) * a (n+1) - c (n+2))" + apply (rule divide_strict_left_mono) + subgoal using \c (n+2) > 0\ by auto + unfolding Suc_n_eq[symmetric] using * \B>0\ by auto + also have "... < (c n * a n - c (n+1)) / (c (n+1) * a (n+1) - c (n+2))" + apply (rule frac_less) + unfolding Suc_n_eq[symmetric] using * \B>0\ \c (n + 1) < c n\ + by (auto simp:algebra_simps) + also have "... = nth_prime n / nth_prime (n+1)" + unfolding Suc_n_eq[symmetric] n_eq[symmetric] using \B>0\ by auto + also have "... < 1" by auto + finally have "(1 + 1 / real_of_int (c (n + 1))) * real_of_int (a n - 1) + / real_of_int (a (n + 1)) < 1 " . + then show "a n + (a n - c (n + 1) - 1) / (c (n + 1)) < (a (n + 1))" + using * by (auto simp:field_simps) + qed + qed + have a_gt_1:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + 1" + using Bc_large a_gt_plus c_pos[THEN eventually_all_ge_at_top] + apply eventually_elim + proof (rule impI) + fix n assume + "c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" + "c (n + 1) < c n" and B_eq:"B * int (nth_prime n) = c n * a n - c (n + 1) \ + \real_of_int (c (n + 1))\ < real_of_int (a n) / 2" and c_pos:"\y\n. 0 < c y" + from this(1,2) + have "a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" by auto + moreover have "a n - 2 * c (n+1) > 0" + using B_eq c_pos[rule_format,of "n+1"] by auto + then have "a n - 2 * c (n+1) \ 1" by simp + then have "(a n - c (n + 1) - 1) / c (n + 1) \ 1" + using c_pos[rule_format,of "n+1"] by (auto simp:field_simps) + ultimately show "a n + 1 < a (n + 1)" by auto + qed + + text\The following corresponds to inequality (3.6) in the paper, which had to be + slightly corrected: \ + + have a_gt_sqrt:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + (sqrt n - 2)" + proof - + have a_2N:"\\<^sub>F N in sequentially. a (2*N) \ N /2 +1" + using c_gt_half a_gt_1[THEN eventually_all_ge_at_top] + proof eventually_elim + case (elim N) + define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" + define f where "f = (\n. a (Suc n) - a n)" + + have f_1:"\x\S. f x\1" and f_0:"\x. f x\0" + subgoal using elim unfolding S_def f_def by auto + subgoal using \mono a\[THEN incseq_SucD] unfolding f_def by auto + done + have "N / 2 < card S" + using elim unfolding S_def by auto + also have "... \ sum f S" + unfolding of_int_sum + apply (rule sum_bounded_below[of _ 1,simplified]) + using f_1 by auto + also have "... \ sum f {N..<2 * N}" + unfolding of_int_sum + apply (rule sum_mono2) + unfolding S_def using f_0 by auto + also have "... = a (2*N) - a N" + unfolding of_int_sum f_def of_int_diff + apply (rule sum_Suc_diff') + by auto + finally have "N / 2 < a (2*N) - a N" . + then show ?case using a_pos[rule_format,of N] by linarith + qed + + have a_n4:"\\<^sub>F n in sequentially. a n > n/4" + proof - + obtain N where a_N:"\n\N. a (2*n) \ n /2+1" + using a_2N unfolding eventually_at_top_linorder by auto + have "a n>n/4" when "n\2*N" for n + proof - + define n' where "n'=n div 2" + have "n'\N" unfolding n'_def using that by auto + have "n/4 < n' /2+1" + unfolding n'_def by auto + also have "... \ a (2*n')" + using a_N \n'\N\ by auto + also have "... \a n" unfolding n'_def + apply (cases "even n") + subgoal by simp + subgoal by (simp add: assms(2) incseqD) + done + finally show ?thesis . + qed + then show ?thesis + unfolding eventually_at_top_linorder by auto + qed + + have c_sqrt:"\\<^sub>F n in sequentially. c n < sqrt n / 4" + proof - + have "\\<^sub>F x in sequentially. x>1" by simp + moreover have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" + using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] + by simp + ultimately have "\\<^sub>F n in sequentially. c n < B*8 *ln n + 1" using a_n4 Bc_large + proof eventually_elim + case (elim n) + from this(4) have "c n=(B*nth_prime n+c (n+1))/a n" + using a_pos[rule_format,of n] + by (auto simp:divide_simps) + also have "... = (B*nth_prime n)/a n+c (n+1)/a n" + by (auto simp:divide_simps) + also have "... < (B*nth_prime n)/a n + 1" + proof - + have "c (n+1)/a n < 1" using elim(4) by auto + then show ?thesis by auto + qed + also have "... < B*8 * ln n + 1" + proof - + have "B*nth_prime n < 2*B*n*ln n" + using \real (nth_prime n) / (real n * ln (real n)) < 2\ \B>0\ \ 1 < n\ + by (auto simp:divide_simps) + moreover have "real n / 4 < real_of_int (a n)" by fact + ultimately have "(B*nth_prime n) / a n < (2*B*n*ln n) / (n/4)" + apply (rule_tac frac_less) + using \B>0\ \ 1 < n\ by auto + also have "... = B*8 * ln n" + using \ 1 < n\ by auto + finally show ?thesis by auto + qed + finally show ?case . + qed + moreover have "\\<^sub>F n in sequentially. B*8 *ln n + 1 < sqrt n / 4" + by real_asymp + ultimately show ?thesis + by eventually_elim auto + qed + + have + "\\<^sub>F n in sequentially. 0 < c (n+1)" + "\\<^sub>F n in sequentially. c (n+1) < sqrt (n+1) / 4" + "\\<^sub>F n in sequentially. n > 4" + "\\<^sub>F n in sequentially. (n - 4) / sqrt (n + 1) + 1 > sqrt n" + subgoal using c_pos[THEN eventually_all_ge_at_top] + by eventually_elim auto + subgoal using c_sqrt[THEN eventually_all_ge_at_top] + by eventually_elim (use le_add1 in blast) + subgoal by simp + subgoal + by real_asymp + done + then show ?thesis using a_gt_plus a_n4 + apply eventually_elim + proof (rule impI) + fix n assume asm:"0 < c (n + 1)" "c (n + 1) < sqrt (real (n + 1)) / 4" and + a_ineq:"c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" + "c (n + 1) < c n" and "n / 4 < a n" "n > 4" + and n_neq:" sqrt (real n) < real (n - 4) / sqrt (real (n + 1)) + 1" + + have "(n-4) / sqrt(n+1) = (n/4 - 1)/ (sqrt (real (n + 1)) / 4)" + using \n>4\ by (auto simp:divide_simps) + also have "... < (a n - 1) / c (n + 1)" + apply (rule frac_less) + using \n > 4\ \n / 4 < a n\ \0 < c (n + 1)\ \c (n + 1) < sqrt (real (n + 1)) / 4\ + by auto + also have "... - 1 = (a n - c (n + 1) - 1) / c (n + 1)" + using \0 < c (n + 1)\ by (auto simp:field_simps) + also have "a n + ... < a (n+1)" + using a_ineq by auto + finally have "a n + ((n - 4) / sqrt (n + 1) - 1) < a (n + 1)" by simp + moreover have "(n - 4) / sqrt (n + 1) - 1 > sqrt n - 2" + using n_neq[THEN diff_strict_right_mono,of 2] \n>4\ + by (auto simp:algebra_simps of_nat_diff) + ultimately show "real_of_int (a n) + (sqrt (real n) - 2) < real_of_int (a (n + 1))" + by argo + qed + qed + + text\The following corresponds to inequality $ a_{2N} > N^{3/2}/2$ in the paper, + which had to be slightly corrected: \ + + have a_2N_sqrt:"\\<^sub>F N in sequentially. a (2*N) > real N * (sqrt (real N)/2 - 1)" + using c_gt_half a_gt_sqrt[THEN eventually_all_ge_at_top] eventually_gt_at_top[of 4] + proof eventually_elim + case (elim N) + define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" + define f where "f = (\n. a (Suc n) - a n)" + + have f_N:"\x\S. f x\sqrt N - 2" + proof + fix x assume "x\S" + then have "sqrt (real x) - 2 < f x" "x\N" + using elim unfolding S_def f_def by auto + moreover have "sqrt x - 2 \ sqrt N - 2" + using \x\N\ by simp + ultimately show "sqrt (real N) - 2 \ real_of_int (f x)" by argo + qed + have f_0:"\x. f x\0" + using \mono a\[THEN incseq_SucD] unfolding f_def by auto + + have "(N / 2) * (sqrt N - 2) < card S * (sqrt N - 2)" + apply (rule mult_strict_right_mono) + subgoal using elim unfolding S_def by auto + subgoal using \N>4\ + by (metis diff_gt_0_iff_gt numeral_less_real_of_nat_iff real_sqrt_four real_sqrt_less_iff) + done + also have "... \ sum f S" + unfolding of_int_sum + apply (rule sum_bounded_below) + using f_N by auto + also have "... \ sum f {N..<2 * N}" + unfolding of_int_sum + apply (rule sum_mono2) + unfolding S_def using f_0 by auto + also have "... = a (2*N) - a N" + unfolding of_int_sum f_def of_int_diff + apply (rule sum_Suc_diff') + by auto + finally have "real N / 2 * (sqrt (real N) - 2) < real_of_int (a (2 * N) - a N)" + . + then have "real N / 2 * (sqrt (real N) - 2) < a (2 * N)" + using a_pos[rule_format,of N] by linarith + then show ?case by (auto simp:field_simps) + qed + + text\The following part is required to derive the final contradiction of the proof.\ + + have a_n_sqrt:"\\<^sub>F n in sequentially. a n > (((n-1)/2) powr (3/2) - (n-1)) /2" + proof (rule sequentially_even_odd_imp) + define f where "f=(\N. ((real (2 * N - 1) / 2) powr (3 / 2) - real (2 * N - 1)) / 2)" + define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" + have "\\<^sub>F N in sequentially. g N > f N" + unfolding f_def g_def + by real_asymp + moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" + unfolding g_def using a_2N_sqrt . + ultimately show "\\<^sub>F N in sequentially. f N < a (2 * N)" + by eventually_elim auto + next + define f where "f=(\N. ((real (2 * N + 1 - 1) / 2) powr (3 / 2) + - real (2 * N + 1 - 1)) / 2)" + define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" + have "\\<^sub>F N in sequentially. g N = f N" + using eventually_gt_at_top[of 0] + apply eventually_elim + unfolding f_def g_def + by (auto simp:algebra_simps powr_half_sqrt[symmetric] powr_mult_base) + moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" + unfolding g_def using a_2N_sqrt . + moreover have "\\<^sub>F N in sequentially. a (2 * N + 1) \ a (2*N)" + apply (rule eventuallyI) + using \mono a\ by (simp add: incseqD) + ultimately show "\\<^sub>F N in sequentially. f N < (a (2 * N + 1))" + apply eventually_elim + by auto + qed + + have a_nth_prime_gt:"\\<^sub>F n in sequentially. a n / nth_prime n > 1" + proof - + define f where "f=(\n::nat. (((n-1)/2) powr (3/2) - (n-1)) /2)" + have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" + using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] + by simp + from this[] eventually_gt_at_top[of 1] + have "\\<^sub>F n in sequentially. real (nth_prime n) < 2*(real n * ln n)" + apply eventually_elim + by (auto simp:field_simps) + moreover have *:"\\<^sub>F N in sequentially. f N >0 " + unfolding f_def + by real_asymp + moreover have " \\<^sub>F n in sequentially. f n < a n" + using a_n_sqrt unfolding f_def . + ultimately have "\\<^sub>F n in sequentially. a n / nth_prime n + > f n / (2*(real n * ln n))" + apply eventually_elim + apply (rule frac_less2) + by auto + moreover have "\\<^sub>F n in sequentially. + (f n)/ (2*(real n * ln n)) > 1" + unfolding f_def + by real_asymp + ultimately show ?thesis + by eventually_elim argo + qed + + have a_nth_prime_lt:"\\<^sub>F n in sequentially. a n / nth_prime n < 1" + proof - + have "liminf (\x. a x / nth_prime x) < 1" + using nth_2 by auto + from this[unfolded less_Liminf_iff] + show ?thesis + apply (auto elim!:frequently_elim1) + by (meson divide_less_eq_1 ereal_less_eq(7) leD leI + nth_prime_nonzero of_nat_eq_0_iff of_nat_less_0_iff order.trans) + qed + + from a_nth_prime_gt a_nth_prime_lt show False + by (simp add: eventually_mono frequently_def) +qed + +section\Acknowledgements\ + +text\A.K.-A. and W.L. were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) + funded by the European Research Council and led by Professor Lawrence Paulson + at the University of Cambridge, UK.\ + end \ No newline at end of file diff --git a/thys/Jordan_Normal_Form/document/root.tex b/thys/Jordan_Normal_Form/document/root.tex --- a/thys/Jordan_Normal_Form/document/root.tex +++ b/thys/Jordan_Normal_Form/document/root.tex @@ -1,133 +1,133 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym} - -% this should be the last package used -\usepackage{pdfsetup} - -% urls in roman style, theory text in math-similar italics -\urlstyle{rm} -\isabellestyle{it} - -\newcommand\isafor{\textsf{IsaFoR}} -\newcommand\ceta{\textsf{Ce\kern-.18emT\kern-.18emA}} - -\begin{document} - -\title{Matrices, Jordan Normal Forms, and Spectral Radius Theory\footnote{Supported by FWF (Austrian Science Fund) project Y757.}} -\author{Ren\'e Thiemann and Akihisa Yamada} -\maketitle - -\begin{abstract} - Matrix interpretations are useful as measure functions in termination proving. - In order to use these interpretations also for complexity analysis, - the growth rate of matrix powers has to examined. Here, we formalized - an important result of spectral radius theory, namely that the growth rate - is polynomially bounded if and only if the spectral radius of a matrix is at most one. - - To formally prove this result we first studied the growth rates of matrices - in Jordan normal form, and prove the result that every - complex matrix has a Jordan normal form by means of two algorithms: - we first convert matrices into similar ones via Schur decomposition, and - then apply a second algorithm which converts an upper-triangular matrix into - Jordan normal form. We further showed uniqueness of Jordan normal forms which - then gives rise to a modular algorithm to compute individual blocks of a Jordan - normal form. - - The whole development is based on a new abstract type for matrices, which is - also executable by a suitable setup of the code generator. It - completely subsumes our former AFP-entry on executable matrices - \cite{Matrix-AFP}, and its main advantage is its close connection to the - HMA-representation which allowed us to easily adapt existing proofs on - determinants. - - All the results have been applied to improve \ceta\ \cite{CeTA,CeTAcomplexity}, - our certifier to validate termination and complexity proof certificates. -\end{abstract} - -\tableofcontents - -\section{Introduction} - -The spectral radius of a square, complex valued matrix $A$ is defined as the -largest norm of some eigenvalue $c$ with eigenvector $v$. -It is a central notion to estimate how -the values in $A^n$ for increasing $n$. If the spectral radius is larger -than $1$, clearly the values grow exponentially, since then -$A^n \cdot v = c^n \cdot v$ becomes exponentially large. - -The other results, namely -that the values in $A^n$ are bounded by a constant, -if the spectral radius is smaller -than $1$, and that there is a polynomial bound if the spectral radius -is exactly $1$ are only immediate for matrices which have an eigenbasis, -a precondition which is not satisfied by every matrix. - -However, these results are derivable via Jordan normal forms (JNFs): -If $J$ is a JNF of $A$, then the growth rates of $A^n$ and $J^n$ are related -by a constant as $A$ and $J$ are similar matrices. And for the values in $J^n$ -there is a closed formula which gives the desired complexity bounds. -To be more precise, the values in $J^n$ are bounded by -${\cal O}(|c|^n \cdot n^{k-1})$ where $k$ is the size of the largest -block of an eigenvalue $c$ which has maximal norm w.r.t.\ the set of all -eigenvalues. And since every complex matrix has a JNF, we can derive the -polynomial (resp.\ constant bounds), if the spectral radius is 1 (resp.\ smaller -than 1). - -These results are already applied in current complexity tools, and the motivation -of this development was to extend our certifier \ceta\ to be able to validate -corresponding complexity proofs. To this end, we formalized -the following main results: -\begin{itemize} -\item an algorithm to compute the characteristic polynomial, since - the eigenvalues are exactly the roots of this polynomial; -\item the complexity bounds for JNFs; and -\item an algorithm which computes JNFs for every matrix, provided that the - list of eigenvalues is given. With the help of the fundamental theorem - of algebra this shows that every complex matrix has a JNF. -\end{itemize} - -Since \ceta\ is generated from Isabelle/HOL via code-generation, all the -algorithms and results need to be available at code-generation time. Especially -there is no possibility to create types on the fly which are chosen to fit -the matrix dimensions of the input. To this end, we cannot use the -matrix-representation of HOL multivariate analysis (HMA). - -Instead, we provide a new matrix library which is based on HOL-algebra with -its explicit carriers. In contrast to our earlier development \cite{Matrix-AFP}, -we do not immediately formalize everything as lists of lists, -but use a more mathematical -notion as triples of the form (dimension, dimension, characteristic-function). -This makes reasoning very similar to HMA, and a suitable implementation type -can be chosen afterwards: we provide one via immutable arrays (we use IArray's from -the HOL library), -but one can also think of an implementation for sparse matrices, etc. -Even the infinite carrier itself is executable where we rely upon Lochbihler's -container framework \cite{Containers-AFP} to have different set representations -at the same time. - -As a consequence of not using HMA, we could not directly reuse existing -algorithms which -have been formalized for this representation. For instance, we formalized our -own version of Gauss-Jordan elimination which is not very different to the -one of Divas\'on and Aransay in \cite{Gauss_Jordan-AFP}: both define row-echelon -form and apply elementary row transformations. Whereas Gauss-Jordan elimination -has been developed from scratch as a case-study to see how suitable our matrix -representation is, in other cases we often just -copied and adjusted existing proofs from HMA. For instance, most of the library for -determinants has been copied from the Isabelle distribution and adapted to our -matrix representation. - - -As a result of our formalization, \ceta\ is now able to check -polynomial bounds for matrix interpretations \cite{MatrixJAR}. - - -% include generated text of all theories -\input{session} - - - -\bibliographystyle{abbrv} -\bibliography{root} - -\end{document} +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\newcommand\isafor{\textsf{IsaFoR}} +\newcommand\ceta{\textsf{Ce\kern-.18emT\kern-.18emA}} + +\begin{document} + +\title{Matrices, Jordan Normal Forms, and Spectral Radius Theory\footnote{Supported by FWF (Austrian Science Fund) project Y757.}} +\author{Ren\'e Thiemann and Akihisa Yamada} +\maketitle + +\begin{abstract} + Matrix interpretations are useful as measure functions in termination proving. + In order to use these interpretations also for complexity analysis, + the growth rate of matrix powers has to examined. Here, we formalized + an important result of spectral radius theory, namely that the growth rate + is polynomially bounded if and only if the spectral radius of a matrix is at most one. + + To formally prove this result we first studied the growth rates of matrices + in Jordan normal form, and prove the result that every + complex matrix has a Jordan normal form by means of two algorithms: + we first convert matrices into similar ones via Schur decomposition, and + then apply a second algorithm which converts an upper-triangular matrix into + Jordan normal form. We further showed uniqueness of Jordan normal forms which + then gives rise to a modular algorithm to compute individual blocks of a Jordan + normal form. + + The whole development is based on a new abstract type for matrices, which is + also executable by a suitable setup of the code generator. It + completely subsumes our former AFP-entry on executable matrices + \cite{Matrix-AFP}, and its main advantage is its close connection to the + HMA-representation which allowed us to easily adapt existing proofs on + determinants. + + All the results have been applied to improve \ceta\ \cite{CeTA,CeTAcomplexity}, + our certifier to validate termination and complexity proof certificates. +\end{abstract} + +\tableofcontents + +\section{Introduction} + +The spectral radius of a square, complex valued matrix $A$ is defined as the +largest norm of some eigenvalue $c$ with eigenvector $v$. +It is a central notion to estimate how +the values in $A^n$ for increasing $n$. If the spectral radius is larger +than $1$, clearly the values grow exponentially, since then +$A^n \cdot v = c^n \cdot v$ becomes exponentially large. + +The other results, namely +that the values in $A^n$ are bounded by a constant, +if the spectral radius is smaller +than $1$, and that there is a polynomial bound if the spectral radius +is exactly $1$ are only immediate for matrices which have an eigenbasis, +a precondition which is not satisfied by every matrix. + +However, these results are derivable via Jordan normal forms (JNFs): +If $J$ is a JNF of $A$, then the growth rates of $A^n$ and $J^n$ are related +by a constant as $A$ and $J$ are similar matrices. And for the values in $J^n$ +there is a closed formula which gives the desired complexity bounds. +To be more precise, the values in $J^n$ are bounded by +${\cal O}(|c|^n \cdot n^{k-1})$ where $k$ is the size of the largest +block of an eigenvalue $c$ which has maximal norm w.r.t.\ the set of all +eigenvalues. And since every complex matrix has a JNF, we can derive the +polynomial (resp.\ constant bounds), if the spectral radius is 1 (resp.\ smaller +than 1). + +These results are already applied in current complexity tools, and the motivation +of this development was to extend our certifier \ceta\ to be able to validate +corresponding complexity proofs. To this end, we formalized +the following main results: +\begin{itemize} +\item an algorithm to compute the characteristic polynomial, since + the eigenvalues are exactly the roots of this polynomial; +\item the complexity bounds for JNFs; and +\item an algorithm which computes JNFs for every matrix, provided that the + list of eigenvalues is given. With the help of the fundamental theorem + of algebra this shows that every complex matrix has a JNF. +\end{itemize} + +Since \ceta\ is generated from Isabelle/HOL via code-generation, all the +algorithms and results need to be available at code-generation time. Especially +there is no possibility to create types on the fly which are chosen to fit +the matrix dimensions of the input. To this end, we cannot use the +matrix-representation of HOL multivariate analysis (HMA). + +Instead, we provide a new matrix library which is based on HOL-algebra with +its explicit carriers. In contrast to our earlier development \cite{Matrix-AFP}, +we do not immediately formalize everything as lists of lists, +but use a more mathematical +notion as triples of the form (dimension, dimension, characteristic-function). +This makes reasoning very similar to HMA, and a suitable implementation type +can be chosen afterwards: we provide one via immutable arrays (we use IArray's from +the HOL library), +but one can also think of an implementation for sparse matrices, etc. +Even the infinite carrier itself is executable where we rely upon Lochbihler's +container framework \cite{Containers-AFP} to have different set representations +at the same time. + +As a consequence of not using HMA, we could not directly reuse existing +algorithms which +have been formalized for this representation. For instance, we formalized our +own version of Gauss-Jordan elimination which is not very different to the +one of Divas\'on and Aransay in \cite{Gauss_Jordan-AFP}: both define row-echelon +form and apply elementary row transformations. Whereas Gauss-Jordan elimination +has been developed from scratch as a case-study to see how suitable our matrix +representation is, in other cases we often just +copied and adjusted existing proofs from HMA. For instance, most of the library for +determinants has been copied from the Isabelle distribution and adapted to our +matrix representation. + + +As a result of our formalization, \ceta\ is now able to check +polynomial bounds for matrix interpretations \cite{MatrixJAR}. + + +% include generated text of all theories +\input{session} + + + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/Paraconsistency/document/root.tex b/thys/Paraconsistency/document/root.tex --- a/thys/Paraconsistency/document/root.tex +++ b/thys/Paraconsistency/document/root.tex @@ -1,124 +1,124 @@ -\documentclass[11pt,a4paper]{article} - -\title{Paraconsistency} - -\author{Anders Schlichtkrull \&\ J{\o}rgen Villadsen, DTU Compute, Denmark} - -\date{\isadate\today} - -\usepackage{datetime,isabelle,isabellesym,parskip,underscore} - -\newdateformat{isadate}{\THEDAY\ \monthname[\THEMONTH] \THEYEAR} - -\usepackage[cm]{fullpage} - -\usepackage{pdfsetup} - -\isabellestyle{tt} - -\urlstyle{rm} - -\renewcommand{\isachardoublequote}{} -\renewcommand{\isachardoublequoteopen}{} -\renewcommand{\isachardoublequoteclose}{} - -\renewcommand{\isamarkupchapter}[1]{\clearpage\isamarkupsection{#1}} -\renewcommand{\isamarkupsection}[1]{\section*{#1}\addcontentsline{toc}{section}{#1}} -\renewcommand{\isamarkupsubsection}[1]{\medskip\subsection*{#1}} -\renewcommand{\isamarkupsubsubsection}[1]{\medskip\subsubsection*{#1}} - -\renewcommand{\isabeginpar}{\par\ifisamarkup\relax\else\bigskip\fi} -\renewcommand{\isaendpar}{\par\bigskip} - -\begin{document} - -\makeatletter -\parbox[t]{\textwidth}{\centering\Huge\bfseries\@title}\par\kern5mm -\parbox[t]{\textwidth}{\centering\Large\bfseries\@author}\par\kern3mm -\parbox[t]{\textwidth}{\centering\bfseries\@date}\par\kern8mm -\makeatother - -\begin{abstract}\normalsize\noindent -Paraconsistency is about handling inconsistency in a coherent way. In classical and intuitionistic -logic everything follows from an inconsistent theory. A paraconsistent logic avoids the explosion. -Quite a few applications in computer science and engineering are discussed in the Intelligent -Systems Reference Library Volume 110: Towards Paraconsistent Engineering (Springer 2016). We -formalize a paraconsistent many-valued logic that we motivated and described in a special issue on -logical approaches to paraconsistency (Journal of Applied Non-Classical Logics 2005). We limit -ourselves to the propositional fragment of the higher-order logic. The logic is based on so-called -key equalities and has a countably infinite number of truth values. We prove theorems in the logic -using the definition of validity. We verify truth tables and also counterexamples for non-theorems. -We prove meta-theorems about the logic and finally we investigate a case study. -\end{abstract} - -\tableofcontents - -\isamarkupsection{Preface} - -The present formalization in Isabelle essentially follows our extended abstract \cite{Jensen+12}. -The Stanford Encyclopedia of Philosophy has a comprehensive overview of logical approaches to -paraconsistency \cite{Priest+15}. We have elsewhere explained the rationale for our paraconsistent -many-valued logic and considered applications in multi-agent systems and natural language semantics -\cite{Villadsen05-JANCL,Villadsen09,Villadsen10,Villadsen14}. - -It is a revised and extended version of our formalization \url{https://github.com/logic-tools/mvl} -that accompany our chapter in a book on partiality published by Cambridge Scholars Press. The GitHub -link provides more information. We are grateful to the editors --- Henning Christiansen, M. Dolores -Jim\'{e}nez L\'{o}pez, Roussanka Loukanova and Larry Moss --- for the opportunity to contribute to -the book. - -\input{session} - -\clearpage\addcontentsline{toc}{section}{References} - -\begin{thebibliography}{0} - -\bibitem{Jensen+12} -A.~S. Jensen and J.~Villadsen. -\newblock -\emph{Paraconsistent Computational Logic}. -\newblock -In P.~Blackburn, K.~F.~J{\o}rgensen, N.~Jones, and E.~Palmgren, editors, 8th Scandinavian Logic -Symposium: Abstracts, pages 59--61, Roskilde University, 2012. - -\bibitem{Priest+15} -G.~Priest, K.~Tanaka and Z.~Weber. -\newblock -\emph{Paraconsistent Logic}. -\newblock -In E.~N. Zalta et~al., editors, Stanford Encyclopedia of Philosophy, Online Entry -\url{http://plato.stanford.edu/entries/logic-paraconsistent/} Spring Edition, 2015. - -\bibitem{Villadsen05-JANCL} -J.~Villadsen. -\newblock -\emph{Supra-logic: Using Transfinite Type Theory with Type Variables for Paraconsistency}. -\newblock -Logical Approaches to Paraconsistency, Journal of Applied Non-Classical Logics, 15(1):45--58, 2005. - -\bibitem{Villadsen09} -J.~Villadsen. -\newblock -\emph{Infinite-Valued Propositional Type Theory for Semantics}. -\newblock -In J.-Y.~B\'{e}ziau and A.~Costa-Leite, editors, Dimensions of Logical Concepts, pages 277--297, -Unicamp Cole\c{c}.~CLE 54, 2009. - -\bibitem{Villadsen10} -J.~Villadsen. -\newblock -\emph{Nabla: A Linguistic System Based on Type Theory}. -\newblock -Foundations of Communication and Cognition (New Series), LIT Verlag, 2010. - -\bibitem{Villadsen14} -J.~Villadsen. -\newblock -\emph{Multi-dimensional Type Theory: Rules, Categories and Combinators for Syntax and Semantics}. -\newblock -In P.~Blache, H.~Christiansen, V.~Dahl, D.~Duchier, and J.~Villadsen, editors, Constraints and -Language, pages 167--189, Cambridge Scholars Press, 2014. - -\end{thebibliography} - -\end{document} +\documentclass[11pt,a4paper]{article} + +\title{Paraconsistency} + +\author{Anders Schlichtkrull \&\ J{\o}rgen Villadsen, DTU Compute, Denmark} + +\date{\isadate\today} + +\usepackage{datetime,isabelle,isabellesym,parskip,underscore} + +\newdateformat{isadate}{\THEDAY\ \monthname[\THEMONTH] \THEYEAR} + +\usepackage[cm]{fullpage} + +\usepackage{pdfsetup} + +\isabellestyle{tt} + +\urlstyle{rm} + +\renewcommand{\isachardoublequote}{} +\renewcommand{\isachardoublequoteopen}{} +\renewcommand{\isachardoublequoteclose}{} + +\renewcommand{\isamarkupchapter}[1]{\clearpage\isamarkupsection{#1}} +\renewcommand{\isamarkupsection}[1]{\section*{#1}\addcontentsline{toc}{section}{#1}} +\renewcommand{\isamarkupsubsection}[1]{\medskip\subsection*{#1}} +\renewcommand{\isamarkupsubsubsection}[1]{\medskip\subsubsection*{#1}} + +\renewcommand{\isabeginpar}{\par\ifisamarkup\relax\else\bigskip\fi} +\renewcommand{\isaendpar}{\par\bigskip} + +\begin{document} + +\makeatletter +\parbox[t]{\textwidth}{\centering\Huge\bfseries\@title}\par\kern5mm +\parbox[t]{\textwidth}{\centering\Large\bfseries\@author}\par\kern3mm +\parbox[t]{\textwidth}{\centering\bfseries\@date}\par\kern8mm +\makeatother + +\begin{abstract}\normalsize\noindent +Paraconsistency is about handling inconsistency in a coherent way. In classical and intuitionistic +logic everything follows from an inconsistent theory. A paraconsistent logic avoids the explosion. +Quite a few applications in computer science and engineering are discussed in the Intelligent +Systems Reference Library Volume 110: Towards Paraconsistent Engineering (Springer 2016). We +formalize a paraconsistent many-valued logic that we motivated and described in a special issue on +logical approaches to paraconsistency (Journal of Applied Non-Classical Logics 2005). We limit +ourselves to the propositional fragment of the higher-order logic. The logic is based on so-called +key equalities and has a countably infinite number of truth values. We prove theorems in the logic +using the definition of validity. We verify truth tables and also counterexamples for non-theorems. +We prove meta-theorems about the logic and finally we investigate a case study. +\end{abstract} + +\tableofcontents + +\isamarkupsection{Preface} + +The present formalization in Isabelle essentially follows our extended abstract \cite{Jensen+12}. +The Stanford Encyclopedia of Philosophy has a comprehensive overview of logical approaches to +paraconsistency \cite{Priest+15}. We have elsewhere explained the rationale for our paraconsistent +many-valued logic and considered applications in multi-agent systems and natural language semantics +\cite{Villadsen05-JANCL,Villadsen09,Villadsen10,Villadsen14}. + +It is a revised and extended version of our formalization \url{https://github.com/logic-tools/mvl} +that accompany our chapter in a book on partiality published by Cambridge Scholars Press. The GitHub +link provides more information. We are grateful to the editors --- Henning Christiansen, M. Dolores +Jim\'{e}nez L\'{o}pez, Roussanka Loukanova and Larry Moss --- for the opportunity to contribute to +the book. + +\input{session} + +\clearpage\addcontentsline{toc}{section}{References} + +\begin{thebibliography}{0} + +\bibitem{Jensen+12} +A.~S. Jensen and J.~Villadsen. +\newblock +\emph{Paraconsistent Computational Logic}. +\newblock +In P.~Blackburn, K.~F.~J{\o}rgensen, N.~Jones, and E.~Palmgren, editors, 8th Scandinavian Logic +Symposium: Abstracts, pages 59--61, Roskilde University, 2012. + +\bibitem{Priest+15} +G.~Priest, K.~Tanaka and Z.~Weber. +\newblock +\emph{Paraconsistent Logic}. +\newblock +In E.~N. Zalta et~al., editors, Stanford Encyclopedia of Philosophy, Online Entry +\url{http://plato.stanford.edu/entries/logic-paraconsistent/} Spring Edition, 2015. + +\bibitem{Villadsen05-JANCL} +J.~Villadsen. +\newblock +\emph{Supra-logic: Using Transfinite Type Theory with Type Variables for Paraconsistency}. +\newblock +Logical Approaches to Paraconsistency, Journal of Applied Non-Classical Logics, 15(1):45--58, 2005. + +\bibitem{Villadsen09} +J.~Villadsen. +\newblock +\emph{Infinite-Valued Propositional Type Theory for Semantics}. +\newblock +In J.-Y.~B\'{e}ziau and A.~Costa-Leite, editors, Dimensions of Logical Concepts, pages 277--297, +Unicamp Cole\c{c}.~CLE 54, 2009. + +\bibitem{Villadsen10} +J.~Villadsen. +\newblock +\emph{Nabla: A Linguistic System Based on Type Theory}. +\newblock +Foundations of Communication and Cognition (New Series), LIT Verlag, 2010. + +\bibitem{Villadsen14} +J.~Villadsen. +\newblock +\emph{Multi-dimensional Type Theory: Rules, Categories and Combinators for Syntax and Semantics}. +\newblock +In P.~Blache, H.~Christiansen, V.~Dahl, D.~Duchier, and J.~Villadsen, editors, Constraints and +Language, pages 167--189, Cambridge Scholars Press, 2014. + +\end{thebibliography} + +\end{document} diff --git a/thys/Recursion-Addition/recursion.thy b/thys/Recursion-Addition/recursion.thy --- a/thys/Recursion-Addition/recursion.thy +++ b/thys/Recursion-Addition/recursion.thy @@ -1,1525 +1,1525 @@ -(* Title: Recursion theorem - Author: Georgy Dunaev , 2020 - Maintainer: Georgy Dunaev -*) -section "Recursion Submission" - -text \Recursion Theorem is proved in the following document. -It also contains the addition on natural numbers. -The development is done in the context of Zermelo-Fraenkel set theory.\ - -theory recursion - imports ZF -begin - -section \Basic Set Theory\ -text \Useful lemmas about sets, functions and natural numbers\ -lemma pisubsig : \Pi(A,P)\Pow(Sigma(A,P))\ -proof - fix x - assume \x \ Pi(A,P)\ - hence \x \ {f\Pow(Sigma(A,P)). A\domain(f) & function(f)}\ - by (unfold Pi_def) - thus \x \ Pow(Sigma(A, P))\ - by (rule CollectD1) -qed - -lemma apparg: - fixes f A B - assumes T0:\f:A\B\ - assumes T1:\f ` a = b\ - assumes T2:\a \ A\ - shows \\a, b\ \ f\ -proof(rule iffD2[OF func.apply_iff], rule T0) - show T:\a \ A \ f ` a = b\ - by (rule conjI[OF T2 T1]) -qed - -theorem nat_induct_bound : - assumes H0:\P(0)\ - assumes H1:\!!x. x\nat \ P(x) \ P(succ(x))\ - shows \\n\nat. P(n)\ -proof(rule ballI) - fix n - assume H2:\n\nat\ - show \P(n)\ - proof(rule nat_induct[of n]) - from H2 show \n\nat\ by assumption - next - show \P(0)\ by (rule H0) - next - fix x - assume H3:\x\nat\ - assume H4:\P(x)\ - show \P(succ(x))\ by (rule H1[OF H3 H4]) - qed -qed - -theorem nat_Tr : \\n\nat. m\n \ m\nat\ -proof(rule nat_induct_bound) - show \m \ 0 \ m \ nat\ by auto -next - fix x - assume H0:\x \ nat\ - assume H1:\m \ x \ m \ nat\ - show \m \ succ(x) \ m \ nat\ - proof(rule impI) - assume H2:\m\succ(x)\ - show \m \ nat\ - proof(rule succE[OF H2]) - assume H3:\m = x\ - from H0 and H3 show \m \ nat\ - by auto - next - assume H4:\m \ x\ - show \m \ nat\ - by(rule mp[OF H1 H4]) - qed - qed -qed - -(* Natural numbers are linearly ordered. *) -theorem zeroleq : \\n\nat. 0\n \ 0=n\ -proof(rule ballI) - fix n - assume H1:\n\nat\ - show \0\n\0=n\ - proof(rule nat_induct[of n]) - from H1 show \n \ nat\ by assumption - next - show \0 \ 0 \ 0 = 0\ by (rule disjI2, rule refl) - next - fix x - assume H2:\x\nat\ - assume H3:\ 0 \ x \ 0 = x\ - show \0 \ succ(x) \ 0 = succ(x)\ - proof(rule disjE[OF H3]) - assume H4:\0\x\ - show \0 \ succ(x) \ 0 = succ(x)\ - proof(rule disjI1) - show \0 \ succ(x)\ - by (rule succI2[OF H4]) - qed - next - assume H4:\0=x\ - show \0 \ succ(x) \ 0 = succ(x)\ - proof(rule disjI1) - have q:\x \ succ(x)\ by auto - from q and H4 show \0 \ succ(x)\ by auto - qed - qed - qed -qed - -theorem JH2_1ii : \m\succ(n) \ m\n\m=n\ - by auto - -theorem nat_transitive:\\n\nat. \k. \m. k \ m \ m \ n \ k \ n\ -proof(rule nat_induct_bound) - show \\k. \m. k \ m \ m \ 0 \ k \ 0\ - proof(rule allI, rule allI, rule impI) - fix k m - assume H:\k \ m \ m \ 0\ - then have H:\m \ 0\ by auto - then show \k \ 0\ by auto - qed -next - fix n - assume H0:\n \ nat\ - assume H1:\\k. - \m. - k \ m \ m \ n \ - k \ n\ - show \\k. \m. - k \ m \ - m \ succ(n) \ - k \ succ(n)\ - proof(rule allI, rule allI, rule impI) - fix k m - assume H4:\k \ m \ m \ succ(n)\ - hence H4':\m \ succ(n)\ by (rule conjunct2) - hence H4'':\m\n \ m=n\ by (rule succE, auto) - from H4 have Q:\k \ m\ by (rule conjunct1) - have H1S:\\m. k \ m \ m \ n \ k \ n\ - by (rule spec[OF H1]) - have H1S:\k \ m \ m \ n \ k \ n\ - by (rule spec[OF H1S]) - show \k \ succ(n)\ - proof(rule disjE[OF H4'']) - assume L:\m\n\ - from Q and L have QL:\k \ m \ m \ n\ by auto - have G:\k \ n\ by (rule mp [OF H1S QL]) - show \k \ succ(n)\ - by (rule succI2[OF G]) - next - assume L:\m=n\ - from Q have F:\k \ succ(m)\ by auto - from L and Q show \k \ succ(n)\ by auto - qed - qed -qed - -theorem nat_xninx : \\n\nat. \(n\n)\ -proof(rule nat_induct_bound) - show \0\0\ - by auto -next - fix x - assume H0:\x\nat\ - assume H1:\x\x\ - show \succ(x) \ succ(x)\ - proof(rule contrapos[OF H1]) - assume Q:\succ(x) \ succ(x)\ - have D:\succ(x)\x \ succ(x)=x\ - by (rule JH2_1ii[OF Q]) - show \x\x\ - proof(rule disjE[OF D]) - assume Y1:\succ(x)\x\ - have U:\x\succ(x)\ by (rule succI1) - have T:\x \ succ(x) \ succ(x) \ x \ x \ x\ - by (rule spec[OF spec[OF bspec[OF nat_transitive H0]]]) - have R:\x \ succ(x) \ succ(x) \ x\ - by (rule conjI[OF U Y1]) - show \x\x\ - by (rule mp[OF T R]) - next - assume Y1:\succ(x)=x\ - show \x\x\ - by (rule subst[OF Y1], rule Q) - qed - qed -qed - -theorem nat_asym : \\n\nat. \m. \(n\m \ m\n)\ -proof(rule ballI, rule allI) - fix n m - assume H0:\n \ nat\ - have Q:\\(n\n)\ - by(rule bspec[OF nat_xninx H0]) - show \\ (n \ m \ m \ n)\ - proof(rule contrapos[OF Q]) - assume W:\(n \ m \ m \ n)\ - show \n\n\ - by (rule mp[OF spec[OF spec[OF bspec[OF nat_transitive H0]]] W]) - qed -qed - -theorem zerolesucc :\\n\nat. 0 \ succ(n)\ -proof(rule nat_induct_bound) - show \0\1\ - by auto -next - fix x - assume H0:\x\nat\ - assume H1:\0\succ(x)\ - show \0\succ(succ(x))\ - proof - assume J:\0 \ succ(x)\ - show \0 = succ(x)\ - by(rule notE[OF J H1]) - qed -qed - -theorem succ_le : \\n\nat. succ(m)\succ(n) \ m\n\ -proof(rule nat_induct_bound) - show \ succ(m) \ 1 \ m \ 0\ - by blast -next - fix x - assume H0:\x \ nat\ - assume H1:\succ(m) \ succ(x) \ m \ x\ - show \ succ(m) \ - succ(succ(x)) \ - m \ succ(x)\ - proof(rule impI) - assume J0:\succ(m) \ succ(succ(x))\ - show \m \ succ(x)\ - proof(rule succE[OF J0]) - assume R:\succ(m) = succ(x)\ - hence R:\m=x\ by (rule upair.succ_inject) - from R and succI1 show \m \ succ(x)\ by auto - next - assume R:\succ(m) \ succ(x)\ - have R:\m\x\ by (rule mp[OF H1 R]) - then show \m \ succ(x)\ by auto - qed - qed -qed - -theorem succ_le2 : \\n\nat. \m. succ(m)\succ(n) \ m\n\ -proof - fix n - assume H:\n\nat\ - show \\m. succ(m) \ succ(n) \ m \ n\ - proof - fix m - from succ_le and H show \succ(m) \ succ(n) \ m \ n\ by auto - qed -qed - -theorem le_succ : \\n\nat. m\n \ succ(m)\succ(n)\ -proof(rule nat_induct_bound) - show \m \ 0 \ succ(m) \ 1\ - by auto -next - fix x - assume H0:\x\nat\ - assume H1:\m \ x \ succ(m) \ succ(x)\ - show \m \ succ(x) \ - succ(m) \ succ(succ(x))\ - proof(rule impI) - assume HR1:\m\succ(x)\ - show \succ(m) \ succ(succ(x))\ - proof(rule succE[OF HR1]) - assume Q:\m = x\ - from Q show \succ(m) \ succ(succ(x))\ - by auto - next - assume Q:\m \ x\ - have Q:\succ(m) \ succ(x)\ - by (rule mp[OF H1 Q]) - from Q show \succ(m) \ succ(succ(x))\ - by (rule succI2) - qed - qed -qed - -theorem nat_linord:\\n\nat. \m\nat. m\n\m=n\n\m\ -proof(rule ballI) - fix n - assume H1:\n\nat\ - show \\m\nat. m \ n \ m = n \ n \ m\ - proof(rule nat_induct[of n]) - from H1 show \n\nat\ by assumption - next - show \\m\nat. m \ 0 \ m = 0 \ 0 \ m\ - proof - fix m - assume J:\m\nat\ - show \ m \ 0 \ m = 0 \ 0 \ m\ - proof(rule disjI2) - have Q:\0\m\0=m\ by (rule bspec[OF zeroleq J]) - show \m = 0 \ 0 \ m\ - by (rule disjE[OF Q], auto) - qed - qed - next - fix x - assume K:\x\nat\ - assume M:\\m\nat. m \ x \ m = x \ x \ m\ - show \\m\nat. - m \ succ(x) \ - m = succ(x) \ - succ(x) \ m\ - proof(rule nat_induct_bound) - show \0 \ succ(x) \ 0 = succ(x) \ succ(x) \ 0\ - proof(rule disjI1) - show \0 \ succ(x)\ - by (rule bspec[OF zerolesucc K]) - qed - next - fix y - assume H0:\y \ nat\ - assume H1:\y \ succ(x) \ y = succ(x) \ succ(x) \ y\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - proof(rule disjE[OF H1]) - assume W:\y\succ(x)\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - proof(rule succE[OF W]) - assume G:\y=x\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - by (rule disjI2, rule disjI1, rule subst[OF G], rule refl) - next - assume G:\y \ x\ - have R:\succ(y) \ succ(x)\ - by (rule mp[OF bspec[OF le_succ K] G]) - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - by(rule disjI1, rule R) - qed - next - assume W:\y = succ(x) \ succ(x) \ y\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - proof(rule disjE[OF W]) - assume W:\y=succ(x)\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - by (rule disjI2, rule disjI2, rule subst[OF W], rule succI1) - next - assume W:\succ(x)\y\ - show \succ(y) \ succ(x) \ - succ(y) = succ(x) \ - succ(x) \ succ(y)\ - by (rule disjI2, rule disjI2, rule succI2[OF W]) - qed - qed - qed - qed -qed - -lemma tgb: - assumes knat: \k\nat\ - assumes D: \t \ k \ A\ - shows \t \ Pow(nat \ A)\ -proof - - from D - have q:\t\{t\Pow(Sigma(k,%_.A)). k\domain(t) & function(t)}\ - by(unfold Pi_def) - have J:\t \ Pow(k \ A)\ - by (rule CollectD1[OF q]) - have G:\k \ A \ nat \ A\ - proof(rule func.Sigma_mono) - from knat - show \k\nat\ - by (rule QUniv.naturals_subset_nat) - next - show \\x. x \ k \ A \ A\ - by auto - qed - show \t \ Pow(nat \ A)\ - by (rule subsetD, rule func.Pow_mono[OF G], rule J) -qed - -section \Compatible set\ -text \Union of compatible set of functions is a function.\ - -definition compat :: \[i,i]\o\ - where "compat(f1,f2) == \x.\y1.\y2.\x,y1\ \ f1 \ \x,y2\ \ f2 \ y1=y2" - -lemma compatI [intro]: - assumes H:\\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2\ - shows \compat(f1,f2)\ -proof(unfold compat_def) - show \\x y1 y2. \x, y1\ \ f1 \ \x, y2\ \ f2 \ y1 = y2\ - proof(rule allI | rule impI)+ - fix x y1 y2 - assume K:\\x, y1\ \ f1 \ \x, y2\ \ f2\ - have K1:\\x, y1\ \ f1\ by (rule conjunct1[OF K]) - have K2:\\x, y2\ \ f2\ by (rule conjunct2[OF K]) - show \y1 = y2\ by (rule H[OF K1 K2]) - qed -qed - -lemma compatD: - assumes H: \compat(f1,f2)\ - shows \\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2\ -proof - - fix x y1 y2 - assume Q1:\\x, y1\ \ f1\ - assume Q2:\\x, y2\ \ f2\ - from H have H:\\x y1 y2. \x, y1\ \ f1 \ \x, y2\ \ f2 \ y1 = y2\ - by (unfold compat_def) - show \y1=y2\ - proof(rule mp[OF spec[OF spec[OF spec[OF H]]]]) - show \\x, y1\ \ f1 \ \x, y2\ \ f2\ - by(rule conjI[OF Q1 Q2]) - qed -qed - -lemma compatE: - assumes H: \compat(f1,f2)\ - and W:\(\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2) \ E\ -shows \E\ - by (rule W, rule compatD[OF H], assumption+) - - -definition compatset :: \i\o\ - where "compatset(S) == \f1\S.\f2\S. compat(f1,f2)" - -lemma compatsetI [intro] : - assumes 1:\\f1 f2. \f1\S;f2\S\ \ compat(f1,f2)\ - shows \compatset(S)\ - by (unfold compatset_def, rule ballI, rule ballI, rule 1, assumption+) - -lemma compatsetD: - assumes H: \compatset(S)\ - shows \\f1 f2.\f1\S; f2\S\\compat(f1,f2)\ -proof - - fix f1 f2 - assume H1:\f1\S\ - assume H2:\f2\S\ - from H have H:\\f1\S.\f2\S. compat(f1,f2)\ - by (unfold compatset_def) - show \compat(f1,f2)\ - by (rule bspec[OF bspec[OF H H1] H2]) -qed - -lemma compatsetE: - assumes H: \compatset(S)\ - and W:\(\f1 f2.\f1\S; f2\S\\compat(f1,f2)) \ E\ -shows \E\ - by (rule W, rule compatsetD[OF H], assumption+) - -theorem upairI1 : \a \ {a, b}\ -proof - assume \a \ {b}\ - show \a = a\ by (rule refl) -qed - -theorem upairI2 : \b \ {a, b}\ -proof - assume H:\b \ {b}\ - have Y:\b \ {b}\ by (rule upair.singletonI) - show \b = a\ by (rule notE[OF H Y]) -qed - -theorem sinup : \{x} \ \x, xa\\ -proof (unfold Pair_def) - show \{x} \ {{x, x}, {x, xa}}\ - proof (rule IFOL.subst) - show \{x} \ {{x},{x,xa}}\ - by (rule upairI1) - next - show \{{x}, {x, xa}} = {{x, x}, {x, xa}}\ - by blast - qed -qed - -theorem compatsetunionfun : - fixes S - assumes H0:\compatset(S)\ - shows \function(\S)\ -proof(unfold function_def) - show \ \x y1. \x, y1\ \ \S \ - (\y2. \x, y2\ \ \S \ y1 = y2)\ - proof(rule allI, rule allI, rule impI, rule allI, rule impI) - fix x y1 y2 - assume F1:\\x, y1\ \ \S\ - assume F2:\\x, y2\ \ \S\ - show \y1=y2\ - proof(rule UnionE[OF F1], rule UnionE[OF F2]) - fix f1 f2 - assume J1:\\x, y1\ \ f1\ - assume J2:\\x, y2\ \ f2\ - assume K1:\f1 \ S\ - assume K2:\f2 \ S\ - have R:\compat(f1,f2)\ - by (rule compatsetD[OF H0 K1 K2]) - show \y1=y2\ - by(rule compatD[OF R J1 J2]) - qed - qed -qed - -theorem mkel : - assumes 1:\A\ - assumes 2:\A\B\ - shows \B\ - by (rule 2, rule 1) - -theorem valofunion : - fixes S - assumes H0:\compatset(S)\ - assumes W:\f\S\ - assumes Q:\f:A\B\ - assumes T:\a\A\ - assumes P:\f ` a = v\ - shows N:\(\S)`a = v\ -proof - - have K:\\a, v\ \ f\ - by (rule apparg[OF Q P T]) - show N:\(\S)`a = v\ - proof(rule function_apply_equality) - show \function(\S)\ - by(rule compatsetunionfun[OF H0]) - next - show \\a, v\ \ \S\ - by(rule UnionI[OF W K ]) - qed -qed - -section "Partial computation" - -definition satpc :: \[i,i,i] \ o \ - where \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ - -text \$m$-step computation based on $a$ and $g$\ -definition partcomp :: \[i,i,i,i,i]\o\ - where \partcomp(A,t,m,a,g) == (t:succ(m)\A) \ (t`0=a) \ satpc(t,m,g)\ - -lemma partcompI [intro]: - assumes H1:\(t:succ(m)\A)\ - assumes H2:\(t`0=a)\ - assumes H3:\satpc(t,m,g)\ - shows \partcomp(A,t,m,a,g)\ -proof (unfold partcomp_def, auto) - show \t \ succ(m) \ A\ by (rule H1) - show \(t`0=a)\ by (rule H2) - show \satpc(t,m,g)\ by (rule H3) -qed - -lemma partcompD1: \partcomp(A,t,m,a,g) \ t \ succ(m) \ A\ - by (unfold partcomp_def, auto) - -lemma partcompD2: \partcomp(A,t,m,a,g) \ (t`0=a)\ - by (unfold partcomp_def, auto) - -lemma partcompD3: \partcomp(A,t,m,a,g) \ satpc(t,m,g)\ - by (unfold partcomp_def, auto) - -lemma partcompE [elim] : - assumes 1:\partcomp(A,t,m,a,g)\ - and 2:\\(t:succ(m)\A) ; (t`0=a) ; satpc(t,m,g)\ \ E\ - shows \E\ - by (rule 2, rule partcompD1[OF 1], rule partcompD2[OF 1], rule partcompD3[OF 1]) - -text \If we add ordered pair in the middle of partial computation then -it will not change.\ -lemma addmiddle: -(* fixes t m a g*) - assumes mnat:\m\nat\ - assumes F:\partcomp(A,t,m,a,g)\ - assumes xinm:\x\m\ - shows \cons(\succ(x), g ` \t ` x, x\\, t) = t\ -proof(rule partcompE[OF F]) - assume F1:\t \ succ(m) \ A\ - assume F2:\t ` 0 = a\ - assume F3:\satpc(t, m, g)\ - from F3 - have W:\\n\m. t ` succ(n) = g ` \t ` n, n\\ - by (unfold satpc_def) - have U:\t ` succ(x) = g ` \t ` x, x\\ - by (rule bspec[OF W xinm]) - have E:\\succ(x), (g ` \t ` x, x\)\ \ t\ - proof(rule apparg[OF F1 U]) - show \succ(x) \ succ(m)\ - by(rule mp[OF bspec[OF le_succ mnat] xinm]) - qed - show ?thesis - by (rule equalities.cons_absorb[OF E]) -qed - - -section \Set of functions \ -text \It is denoted as $F$ on page 48 in "Introduction to Set Theory".\ -definition pcs :: \[i,i,i]\i\ - where \pcs(A,a,g) == {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ - -lemma pcs_uniq : - assumes F1:\m1\nat\ - assumes F2:\m2\nat\ - assumes H1: \partcomp(A,f1,m1,a,g)\ - assumes H2: \partcomp(A,f2,m2,a,g)\ - shows \\n\nat. n\succ(m1) \ n\succ(m2) \ f1`n = f2`n\ -proof(rule partcompE[OF H1], rule partcompE[OF H2]) - assume H11:\f1 \ succ(m1) \ A\ - assume H12:\f1 ` 0 = a \ - assume H13:\satpc(f1, m1, g)\ - assume H21:\f2 \ succ(m2) \ A\ - assume H22:\f2 ` 0 = a\ - assume H23:\satpc(f2, m2, g)\ - show \\n\nat. n\succ(m1) \ n\succ(m2) \ f1`n = f2`n\ -proof(rule nat_induct_bound) - from H12 and H22 - show \0\succ(m1) \ 0\succ(m2) \ f1 ` 0 = f2 ` 0\ - by auto -next - fix x - assume J0:\x\nat\ - assume J1:\x \ succ(m1) \ x \ succ(m2) \ f1 ` x = f2 ` x\ - from H13 have G1:\\n \ m1 . f1`succ(n) = g ` \ - by (unfold satpc_def, auto) - from H23 have G2:\\n \ m2 . f2`succ(n) = g ` \ - by (unfold satpc_def, auto) - show \succ(x) \ succ(m1) \ succ(x) \ succ(m2) \ - f1 ` succ(x) = f2 ` succ(x)\ - proof - assume K:\succ(x) \ succ(m1) \ succ(x) \ succ(m2)\ - from K have K1:\succ(x) \ succ(m1)\ by auto - from K have K2:\succ(x) \ succ(m2)\ by auto - have K1':\x \ m1\ by (rule mp[OF bspec[OF succ_le F1] K1]) - have K2':\x \ m2\ by (rule mp[OF bspec[OF succ_le F2] K2]) - have U1:\x\succ(m1)\ - by (rule Nat.succ_in_naturalD[OF K1 Nat.nat_succI[OF F1]]) - have U2:\x\succ(m2)\ - by (rule Nat.succ_in_naturalD[OF K2 Nat.nat_succI[OF F2]]) - have Y1:\f1`succ(x) = g ` \ - by (rule bspec[OF G1 K1']) - have Y2:\f2`succ(x) = g ` \ - by (rule bspec[OF G2 K2']) - have \f1 ` x = f2 ` x\ - by(rule mp[OF J1 conjI[OF U1 U2]]) - then have Y:\g ` = g ` \ by auto - from Y1 and Y2 and Y - show \f1 ` succ(x) = f2 ` succ(x)\ - by auto - qed -qed -qed - -lemma domainsubsetfunc : - assumes Q:\f1\f2\ - shows \domain(f1)\domain(f2)\ -proof - fix x - assume H:\x \ domain(f1)\ - show \x \ domain(f2)\ - proof(rule domainE[OF H]) - fix y - assume W:\\x, y\ \ f1\ - have \\x, y\ \ f2\ - by(rule subsetD[OF Q W]) - then show \x \ domain(f2)\ - by(rule domainI) - qed -qed - -lemma natdomfunc: - assumes 1:\q\A\ - assumes J0:\f1 \ Pow(nat \ A)\ - assumes U:\m1 \ domain(f1)\ - shows \m1\nat\ -proof - - from J0 have J0 : \f1 \ nat \ A\ - by auto - have J0:\domain(f1) \ domain(nat \ A)\ - by(rule func.domain_mono[OF J0]) - have F:\m1 \ domain(nat \ A)\ - by(rule subsetD[OF J0 U]) - have R:\domain(nat \ A) = nat\ - by (rule equalities.domain_of_prod[OF 1]) - show \m1 \ nat\ - by(rule subst[OF R], rule F) -qed - -lemma pcs_lem : - assumes 1:\q\A\ - shows \compatset(pcs(A, a, g))\ -proof (*(rule compatsetI)*) - fix f1 f2 - assume H1:\f1 \ pcs(A, a, g)\ - then have H1':\f1 \ {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ by (unfold pcs_def) - hence H1'A:\f1 \ Pow(nat*A)\ by auto - hence H1'A:\f1 \ (nat*A)\ by auto - assume H2:\f2 \ pcs(A, a, g)\ - then have H2':\f2 \ {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ by (unfold pcs_def) - show \compat(f1, f2)\ - proof(rule compatI) - fix x y1 y2 - assume P1:\\x, y1\ \ f1\ - assume P2:\\x, y2\ \ f2\ - show \y1 = y2\ - proof(rule CollectE[OF H1'], rule CollectE[OF H2']) - assume J0:\f1 \ Pow(nat \ A)\ - assume J1:\f2 \ Pow(nat \ A)\ - assume J2:\\m\nat. partcomp(A, f1, m, a, g)\ - assume J3:\\m\nat. partcomp(A, f2, m, a, g)\ - show \y1 = y2\ - proof(rule bexE[OF J2], rule bexE[OF J3]) - fix m1 m2 - assume K1:\partcomp(A, f1, m1, a, g)\ - assume K2:\partcomp(A, f2, m2, a, g)\ - hence K2':\(f2:succ(m2)\A) \ (f2`0=a) \ satpc(f2,m2,g)\ - by (unfold partcomp_def) - from K1 have K1'A:\(f1:succ(m1)\A)\ by (rule partcompD1) - from K2' have K2'A:\(f2:succ(m2)\A)\ by auto - from K1'A have K1'AD:\domain(f1) = succ(m1)\ - by(rule domain_of_fun) - from K2'A have K2'AD:\domain(f2) = succ(m2)\ - by(rule domain_of_fun) - have L1:\f1`x=y1\ - by (rule func.apply_equality[OF P1], rule K1'A) - have L2:\f2`x=y2\ - by(rule func.apply_equality[OF P2], rule K2'A) - have m1nat:\m1\nat\ - proof(rule natdomfunc[OF 1 J0]) - show \m1 \ domain(f1)\ - by (rule ssubst[OF K1'AD], auto) - qed - have m2nat:\m2\nat\ - proof(rule natdomfunc[OF 1 J1]) - show \m2 \ domain(f2)\ - by (rule ssubst[OF K2'AD], auto) - qed - have G1:\\x, y1\ \ (nat*A)\ - by(rule subsetD[OF H1'A P1]) - have KK:\x\nat\ - by(rule SigmaE[OF G1], auto) - (*x is in the domain of f1 i.e. succ(m1) -so we can have both x \ ?m1.2 \ x \ ?m2.2 -how to prove that m1 \ nat ? from J0 ! f1 is a subset of nat \ A*) - have W:\f1`x=f2`x\ - proof(rule mp[OF bspec[OF pcs_uniq KK] ]) - show \m1 \ nat\ - by (rule m1nat) - next - show \m2 \ nat\ - by (rule m2nat) - next - show \partcomp(A, f1, m1, a, g)\ - by (rule K1) - next - show \partcomp(A, f2, m2, a, g)\ - by (rule K2) - next - (* P1:\\x, y1\ \ f1\ - K1'A:\(f1:succ(m1)\A)\ - *) - have U1:\x \ succ(m1)\ - by (rule func.domain_type[OF P1 K1'A]) - have U2:\x \ succ(m2)\ - by (rule func.domain_type[OF P2 K2'A]) - show \x \ succ(m1) \ x \ succ(m2)\ - by (rule conjI[OF U1 U2]) - qed - from L1 and W and L2 - show \y1 = y2\ by auto - qed - qed - qed -qed - -theorem fuissu : \f \ X -> Y \ f \ X\Y\ -proof - fix w - assume H1 : \f \ X -> Y\ - then have J1:\f \ {q\Pow(Sigma(X,\_.Y)). X\domain(q) & function(q)}\ - by (unfold Pi_def) - then have J2:\f \ Pow(Sigma(X,\_.Y))\ - by auto - then have J3:\f \ Sigma(X,\_.Y)\ - by auto - assume H2 : \w \ f\ - from J3 and H2 have \w\Sigma(X,\_.Y)\ - by auto - then have J4:\w \ (\x\X. (\y\Y. {\x,y\}))\ - by auto - show \w \ X*Y\ - proof (rule UN_E[OF J4]) - fix x - assume V1:\x \ X\ - assume V2:\w \ (\y\Y. {\x, y\})\ - show \w \ X \ Y\ - proof (rule UN_E[OF V2]) - fix y - assume V3:\y \ Y\ - assume V4:\w \ {\x, y\}\ - then have V4:\w = \x, y\\ - by auto - have v5:\\x, y\ \ Sigma(X,\_.Y)\ - proof(rule SigmaI) - show \x \ X\ by (rule V1) - next - show \y \ Y\ by (rule V3) - qed - then have V5:\\x, y\ \ X*Y\ - by auto - from V4 and V5 show \w \ X \ Y\ by auto - qed - qed -qed - -theorem recuniq : - fixes f - assumes H0:\f \ nat -> A \ f ` 0 = a \ satpc(f, nat, g)\ - fixes t - assumes H1:\t \ nat -> A \ t ` 0 = a \ satpc(t, nat, g)\ - fixes x - shows \f=t\ -proof - - from H0 have H02:\\n \ nat. f`succ(n) = g ` <(f`n), n>\ by (unfold satpc_def, auto) - from H0 have H01:\f ` 0 = a\ by auto - from H0 have H00:\f \ nat -> A\ by auto - from H1 have H12:\\n \ nat. t`succ(n) = g ` <(t`n), n>\ by (unfold satpc_def, auto) - from H1 have H11:\t ` 0 = a\ by auto - from H1 have H10:\t \ nat -> A\ by auto - show \f=t\ - proof (rule fun_extension[OF H00 H10]) - fix x - assume K: \x \ nat\ - show \(f ` x) = (t ` x)\ - proof(rule nat_induct[of x]) - show \x \ nat\ by (rule K) - next - from H01 and H11 show \f ` 0 = t ` 0\ - by auto - next - fix x - assume A:\x\nat\ - assume B:\f`x = t`x\ - show \f ` succ(x) = t ` succ(x)\ - proof - - from H02 and A have H02':\f`succ(x) = g ` <(f`x), x>\ - by (rule bspec) - from H12 and A have H12':\t`succ(x) = g ` <(t`x), x>\ - by (rule bspec) - from B and H12' have H12'':\t`succ(x) = g ` <(f`x), x>\ by auto - from H12'' and H02' show \f ` succ(x) = t ` succ(x)\ by auto - qed - qed - qed -qed - -section \Lemmas for recursion theorem\ - -locale recthm = - fixes A :: "i" - and a :: "i" - and g :: "i" - assumes hyp1 : \a \ A\ - and hyp2 : \g : ((A*nat)\A)\ -begin - -lemma l3:\function(\pcs(A, a, g))\ - by (rule compatsetunionfun, rule pcs_lem, rule hyp1) - -lemma l1 : \\pcs(A, a, g) \ nat \ A\ -proof - fix x - assume H:\x \ \pcs(A, a, g)\ - hence H:\x \ \{t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ - by (unfold pcs_def) - show \x \ nat \ A\ - proof(rule UnionE[OF H]) - fix B - assume J1:\x\B\ - assume J2:\B \ {t \ Pow(nat \ A) . - \m\nat. partcomp(A, t, m, a, g)}\ - hence J2:\B \ Pow(nat \ A)\ by auto - hence J2:\B \ nat \ A\ by auto - from J1 and J2 show \x \ nat \ A\ - by auto - qed -qed - -lemma le1: - assumes H:\x\1\ - shows \x=0\ -proof - show \x \ 0\ - proof - fix z - assume J:\z\x\ - show \z\0\ - proof(rule succE[OF H]) - assume J:\x\0\ - show \z\0\ - by (rule notE[OF not_mem_empty J]) - next - assume K:\x=0\ - from J and K show \z\0\ - by auto - qed - qed -next - show \0 \ x\ by auto -qed - -lemma lsinglfun : \function({\0, a\})\ -proof(unfold function_def) - show \ \x y. \x, y\ \ {\0, a\} \ - (\y'. \x, y'\ \ {\0, a\} \ - y = y')\ - proof(rule allI,rule allI,rule impI,rule allI,rule impI) - fix x y y' - assume H0:\\x, y\ \ {\0, a\}\ - assume H1:\\x, y'\ \ {\0, a\}\ - show \y = y'\ - proof(rule upair.singletonE[OF H0],rule upair.singletonE[OF H1]) - assume H0:\\x, y\ = \0, a\\ - assume H1:\\x, y'\ = \0, a\\ - from H0 and H1 have H:\\x, y\ = \x, y'\\ by auto - then show \y = y'\ by auto - qed - qed -qed - -lemma singlsatpc:\satpc({\0, a\}, 0, g)\ -proof(unfold satpc_def) - show \\n\0. {\0, a\} ` succ(n) = - g ` \{\0, a\} ` n, n\\ - by auto -qed - -lemma zerostep : - shows \partcomp(A, {\0, a\}, 0, a, g)\ -proof(unfold partcomp_def) - show \{\0, a\} \ 1 -> A \ {\0, a\} ` 0 = a \ satpc({\0, a\}, 0, g)\ - proof - show \{\0, a\} \ 1 -> A\ - proof (unfold Pi_def) - show \{\0, a\} \ {f \ Pow(1 \ A) . 1 \ domain(f) \ function(f)}\ - proof - show \{\0, a\} \ Pow(1 \ A)\ - proof(rule PowI, rule equalities.singleton_subsetI) - show \\0, a\ \ 1 \ A\ - proof - show \0 \ 1\ by auto - next - show \a \ A\ by (rule hyp1) - qed - qed - next - show \1 \ domain({\0, a\}) \ function({\0, a\})\ - proof - show \1 \ domain({\0, a\})\ - proof - fix x - assume W:\x\1\ - from W have W:\x=0\ by (rule le1) - have Y:\0\domain({\0, a\})\ - by auto - from W and Y - show \x\domain({\0, a\})\ - by auto - qed - next - show \function({\0, a\})\ - by (rule lsinglfun) - qed - qed - qed - show \{\0, a\} ` 0 = a \ satpc({\0, a\}, 0, g)\ - proof - show \{\0, a\} ` 0 = a\ - by (rule func.singleton_apply) - next - show \satpc({\0, a\}, 0, g)\ - by (rule singlsatpc) - qed - qed -qed - -lemma zainupcs : \\0, a\ \ \pcs(A, a, g)\ -proof - show \\0, a\ \ {\0, a\}\ - by auto -next - (* {\0, a\} is a 0-step computation *) - show \{\0, a\} \ pcs(A, a, g)\ - proof(unfold pcs_def) - show \{\0, a\} \ {t \ Pow(nat \ A) . \m\nat. partcomp(A, t, m, a, g)}\ - proof - show \{\0, a\} \ Pow(nat \ A)\ - proof(rule PowI, rule equalities.singleton_subsetI) - show \\0, a\ \ nat \ A\ - proof - show \0 \ nat\ by auto - next - show \a \ A\ by (rule hyp1) - qed - qed - next - show \\m\nat. partcomp(A, {\0, a\}, m, a, g)\ - proof - show \partcomp(A, {\0, a\}, 0, a, g)\ - by (rule zerostep) - next - show \0 \ nat\ by auto - qed - qed - qed -qed - -lemma l2': \0 \ domain(\pcs(A, a, g))\ -proof - show \\0, a\ \ \pcs(A, a, g)\ - by (rule zainupcs) -qed - -text \Push an ordered pair to the end of partial computation t -and obtain another partial computation.\ -lemma shortlem : - assumes mnat:\m\nat\ - assumes F:\partcomp(A,t,m,a,g)\ - shows \partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ -proof(rule partcompE[OF F]) - assume F1:\t \ succ(m) \ A\ - assume F2:\t ` 0 = a\ - assume F3:\satpc(t, m, g)\ - show ?thesis (*\partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ *) - proof - have ljk:\cons(\succ(m), g ` \t ` m, m\\, t) \ (cons(succ(m),succ(m)) \ A)\ - proof(rule func.fun_extend3[OF F1]) - show \succ(m) \ succ(m)\ - by (rule upair.mem_not_refl) - have tmA:\t ` m \ A\ - by (rule func.apply_funtype[OF F1], auto) - show \g ` \t ` m, m\ \ A\ - by(rule func.apply_funtype[OF hyp2], auto, rule tmA, rule mnat) - qed - have \cons(\succ(m), g ` \t ` m, m\\, t) \ (cons(succ(m),succ(m)) \ A)\ - by (rule ljk) - then have \cons(\cons(m, m), g ` \t ` m, m\\, t) \ cons(cons(m, m), cons(m, m)) \ A\ - by (unfold succ_def) - then show \cons(\succ(m), g ` \t ` m, m\\, t) \ succ(succ(m)) \ A\ - by (unfold succ_def, assumption) - show \cons(\succ(m), g ` \t ` m, m\\, t) ` 0 = a\ - proof(rule trans, rule func.fun_extend_apply[OF F1]) - show \succ(m) \ succ(m)\ by (rule upair.mem_not_refl) - show \(if 0 = succ(m) then g ` \t ` m, m\ else t ` 0) = a\ - by(rule trans, rule upair.if_not_P, auto, rule F2) - qed - show \satpc(cons(\succ(m), g ` \t ` m, m\\, t), succ(m), g)\ - proof(unfold satpc_def, rule ballI) - fix n - assume Q:\n \ succ(m)\ - show \cons(\succ(m), g ` \t ` m, m\\, t) ` succ(n) -= g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - proof(rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) - show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = - g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - proof(rule upair.succE[OF Q]) - assume Y:\n=m\ - show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = - g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - proof(rule trans, rule upair.if_P) - from Y show \succ(n) = succ(m)\ by auto - next - have L1:\t ` m = cons(\succ(m), g ` \t ` m, m\\, t) ` n\ - proof(rule sym, rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) - show \ (if n = succ(m) then g ` \t ` m, m\ else t ` n) = t ` m\ - proof(rule trans, rule upair.if_not_P) - from Y show \t ` n = t ` m\ by auto - show \n \ succ(m)\ - proof(rule not_sym) - show \succ(m) \ n\ - by(rule subst, rule sym, rule Y, rule upair.succ_neq_self) - qed - qed - qed - from Y - have L2:\m = n\ - by auto - have L:\ \t ` m, m\ = \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - by(rule subst_context2[OF L1 L2]) - show \ g ` \t ` m, m\ = g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - by(rule subst_context[OF L]) - qed - next - assume Y:\n \ m\ - show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = - g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - proof(rule trans, rule upair.if_not_P) - show \succ(n) \ succ(m)\ - by(rule contrapos, rule upair.mem_imp_not_eq, rule Y, rule upair.succ_inject, assumption) - next - have X:\cons(\succ(m), g ` \t ` m, m\\, t) ` n = t ` n\ - proof(rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) - show \(if n = succ(m) then g ` \t ` m, m\ else t ` n) = t ` n\ - proof(rule upair.if_not_P) - show \n \ succ(m)\ - proof(rule contrapos) - assume q:"n=succ(m)" - from q and Y have M:\succ(m)\m\ - by auto - show \m\m\ - by(rule Nat.succ_in_naturalD[OF M mnat]) - next - show \m \ m\ by (rule upair.mem_not_refl) - qed - qed - qed - from F3 - have W:\\n\m. t ` succ(n) = g ` \t ` n, n\\ - by (unfold satpc_def) - have U:\t ` succ(n) = g ` \t ` n, n\\ - by (rule bspec[OF W Y]) - show \t ` succ(n) = g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ - by (rule trans, rule U, rule sym, rule subst_context[OF X]) - qed - qed - qed - qed - qed -qed - -lemma l2:\nat \ domain(\pcs(A, a, g))\ -proof - fix x - assume G:\x\nat\ - show \x \ domain(\pcs(A, a, g))\ - proof(rule nat_induct[of x]) - show \x\nat\ by (rule G) - next - fix x - assume Q1:\x\nat\ - assume Q2:\x\domain(\pcs(A, a, g))\ - show \succ(x)\domain(\pcs(A, a, g))\ - proof(rule domainE[OF Q2]) - fix y - assume W1:\\x, y\ \ (\pcs(A, a, g))\ - show \succ(x)\domain(\pcs(A, a, g))\ - proof(rule UnionE[OF W1]) - fix t - assume E1:\\x, y\ \ t\ - assume E2:\t \ pcs(A, a, g)\ - hence E2:\t\{t\Pow(nat*A). \m \ nat. partcomp(A,t,m,a,g)}\ - by(unfold pcs_def) - have E21:\t\Pow(nat*A)\ - by(rule CollectD1[OF E2]) - have E22m:\\m\nat. partcomp(A,t,m,a,g)\ - by(rule CollectD2[OF E2]) - show \succ(x)\domain(\pcs(A, a, g))\ - proof(rule bexE[OF E22m]) - fix m - assume mnat:\m\nat\ - assume E22P:\partcomp(A,t,m,a,g)\ - hence E22:\((t:succ(m)\A) \ (t`0=a)) \ satpc(t,m,g)\ - by(unfold partcomp_def, auto) - hence E223:\satpc(t,m,g)\ by auto - hence E223:\\n \ m . t`succ(n) = g ` \ - by(unfold satpc_def, auto) - from E22 have E221:\(t:succ(m)\A)\ - by auto - from E221 have domt:\domain(t) = succ(m)\ - by (rule func.domain_of_fun) - from E1 have xind:\x \ domain(t)\ - by (rule equalities.domainI) - from xind and domt have xinsm:\x \ succ(m)\ - by auto - show \succ(x)\domain(\pcs(A, a, g))\ - proof - (*proof(rule exE[OF E22])*) - show \ \succ(x), g ` \ \ (\pcs(A, a, g))\ (*?*) - proof - (*t\{\succ(x), g ` \}*) - show \cons(\succ(x), g ` \, t) \ pcs(A, a, g)\ - proof(unfold pcs_def, rule CollectI) - from E21 - have L1:\t \ nat \ A\ - by auto - from Q1 have J1:\succ(x)\nat\ - by auto(*Nat.nat_succI*) - have txA: \t ` x \ A\ - by (rule func.apply_type[OF E221 xinsm]) - from txA and Q1 have txx:\\t ` x, x\ \ A \ nat\ - by auto - have secp: \g ` \t ` x, x\ \ A\ - by(rule func.apply_type[OF hyp2 txx]) - from J1 and secp - have L2:\\succ(x),g ` \t ` x, x\\ \ nat \ A\ - by auto - show \ cons(\succ(x),g ` \t ` x, x\\,t) \ Pow(nat \ A)\ - proof(rule PowI) - show \ cons(\succ(x), g ` \t ` x, x\\, t) \ nat \ A\ - proof - show \\succ(x), g ` \t ` x, x\\ \ nat \ A \ t \ nat \ A\ - by (rule conjI[OF L2 L1]) - qed - qed - next - show \\m \ nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ - proof(rule succE[OF xinsm]) - assume xeqm:\x=m\ - show \\m \ nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ - proof - show \partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), succ(x), a, g)\ - proof(rule shortlem[OF Q1]) - show \partcomp(A, t, x, a, g)\ - proof(rule subst[of m x], rule sym, rule xeqm) - show \partcomp(A, t, m, a, g)\ - by (rule E22P) - qed - qed - next - from Q1 show \succ(x) \ nat\ by auto - qed - next - assume xinm:\x\m\ - have lmm:\cons(\succ(x), g ` \t ` x, x\\, t) = t\ - by (rule addmiddle[OF mnat E22P xinm]) - show \\m\nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ - by(rule subst[of t], rule sym, rule lmm, rule E22m) - qed - qed - next - show \\succ(x), g ` \t ` x, x\\ \ cons(\succ(x), g ` \t ` x, x\\, t)\ - by auto - qed - qed - qed - qed - qed - next - show \0 \ domain(\pcs(A, a, g))\ - by (rule l2') - qed -qed - -lemma useful : \\m\nat. \t. partcomp(A,t,m,a,g)\ -proof(rule nat_induct_bound) - show \\t. partcomp(A, t, 0, a, g)\ - proof - show \partcomp(A, {\0, a\}, 0, a, g)\ - by (rule zerostep) - qed -next - fix m - assume mnat:\m\nat\ - assume G:\\t. partcomp(A,t,m,a,g)\ - show \\t. partcomp(A,t,succ(m),a,g)\ - proof(rule exE[OF G]) - fix t - assume G:\partcomp(A,t,m,a,g)\ - show \\t. partcomp(A,t,succ(m),a,g)\ - proof - show \partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ - by(rule shortlem[OF mnat G]) - qed - qed -qed - -lemma l4 : \(\pcs(A,a,g)) \ nat -> A\ -proof(unfold Pi_def) - show \ \pcs(A, a, g) \ {f \ Pow(nat \ A) . nat \ domain(f) \ function(f)}\ - proof - show \\pcs(A, a, g) \ Pow(nat \ A)\ - proof - show \\pcs(A, a, g) \ nat \ A\ - by (rule l1) - qed - next - show \nat \ domain(\pcs(A, a, g)) \ function(\pcs(A, a, g))\ - proof - show \nat \ domain(\pcs(A, a, g))\ - by (rule l2) - next - show \function(\pcs(A, a, g))\ - by (rule l3) - qed - qed -qed - -lemma l5: \(\pcs(A, a, g)) ` 0 = a\ -proof(rule func.function_apply_equality) - show \function(\pcs(A, a, g))\ - by (rule l3) -next - show \\0, a\ \ \pcs(A, a, g)\ - by (rule zainupcs) -qed - -lemma ballE2: - assumes \\x\AA. P(x)\ - assumes \x\AA\ - assumes \P(x) ==> Q\ - shows Q - by (rule assms(3), rule bspec, rule assms(1), rule assms(2)) - -text \ Recall that - \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ - \partcomp(A,t,m,a,g) == (t:succ(m)\A) \ (t`0=a) \ satpc(t,m,g)\ - \pcs(A,a,g) == {t\Pow(nat*A). \m. partcomp(A,t,m,a,g)}\ -\ - -lemma l6new: \satpc(\pcs(A, a, g), nat, g)\ -proof (unfold satpc_def, rule ballI) - fix n - assume nnat:\n\nat\ - hence snnat:\succ(n)\nat\ by auto - (* l2:\nat \ domain(\pcs(A, a, g))\ *) - show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ - proof(rule ballE2[OF useful snnat], erule exE) - fix t - assume Y:\partcomp(A, t, succ(n), a, g)\ - show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ - proof(rule partcompE[OF Y]) - assume Y1:\t \ succ(succ(n)) \ A\ - assume Y2:\t ` 0 = a\ - assume Y3:\satpc(t, succ(n), g)\ - hence Y3:\\x \ succ(n) . t`succ(x) = g ` \ - by (unfold satpc_def) - hence Y3:\t`succ(n) = g ` \ - by (rule bspec, auto) - have e1:\(\pcs(A, a, g)) ` succ(n) = t ` succ(n)\ - proof(rule valofunion, rule pcs_lem, rule hyp1) - show \t \ pcs(A, a, g)\ - proof(unfold pcs_def, rule CollectI) - show \t \ Pow(nat \ A)\ - proof(rule tgb) - show \t \ succ(succ(n)) \ A\ by (rule Y1) - next - from snnat - show \succ(succ(n)) \ nat\ by auto - qed - next - show \\m\nat. partcomp(A, t, m, a, g)\ - by(rule bexI, rule Y, rule snnat) - qed - next - show \t \ succ(succ(n)) \ A\ by (rule Y1) - next - show \succ(n) \ succ(succ(n))\ by auto - next - show \t ` succ(n) = t ` succ(n)\ by (rule refl) - qed - have e2:\(\pcs(A, a, g)) ` n = t ` n\ - proof(rule valofunion, rule pcs_lem, rule hyp1) - show \t \ pcs(A, a, g)\ - proof(unfold pcs_def, rule CollectI) - show \t \ Pow(nat \ A)\ - proof(rule tgb) - show \t \ succ(succ(n)) \ A\ by (rule Y1) - next - from snnat - show \succ(succ(n)) \ nat\ by auto - qed - next - show \\m\nat. partcomp(A, t, m, a, g)\ - by(rule bexI, rule Y, rule snnat) - qed - next - show \t \ succ(succ(n)) \ A\ by (rule Y1) - next - show \n \ succ(succ(n))\ by auto - next - show \t ` n = t ` n\ by (rule refl) - qed - have e3:\g ` \(\pcs(A, a, g)) ` n, n\ = g ` \t ` n, n\\ - by (rule subst[OF e2], rule refl) - show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ - by (rule trans, rule e1,rule trans, rule Y3, rule sym, rule e3) - qed - qed -qed - -section "Recursion theorem" - -theorem recursionthm: - shows \\!f. ((f \ (nat\A)) \ ((f`0) = a) \ satpc(f,nat,g))\ -(* where \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ *) -proof - show \\f. f \ nat -> A \ f ` 0 = a \ satpc(f, nat, g)\ - proof - show \(\pcs(A,a,g)) \ nat -> A \ (\pcs(A,a,g)) ` 0 = a \ satpc(\pcs(A,a,g), nat, g)\ - proof - show \\pcs(A, a, g) \ nat -> A\ - by (rule l4) - next - show \(\pcs(A, a, g)) ` 0 = a \ satpc(\pcs(A, a, g), nat, g)\ - proof - show \(\pcs(A, a, g)) ` 0 = a\ - by (rule l5) - next - show \satpc(\pcs(A, a, g), nat, g)\ - by (rule l6new) - qed - qed - qed -next - show \\f y. f \ nat -> A \ - f ` 0 = a \ - satpc(f, nat, g) \ - y \ nat -> A \ - y ` 0 = a \ - satpc(y, nat, g) \ - f = y\ - by (rule recuniq) -qed - -end - -section "Lemmas for addition" - -text \ -Let's define function t(x) = (a+x). -Firstly we need to define a function \g:nat \ nat \ nat\, such that -\g`\t`n, n\ = t`succ(n) = a + (n + 1) = (a + n) + 1 = (t`n) + 1\ -So \g`\a, b\ = a + 1\ and \g(p) = succ(pr1(p))\ -and \satpc(t,\,g) \ \n \ \ . t`succ(n) = succ(t`n)\. -\ - -definition addg :: \i\ - where addg_def : \addg == \x\(nat*nat). succ(fst(x))\ - -lemma addgfun: \function(addg)\ - by (unfold addg_def, rule func.function_lam) - -lemma addgsubpow : \addg \ Pow((nat \ nat) \ nat)\ -proof (unfold addg_def, rule subsetD) - show \(\x\nat \ nat. succ(fst(x))) \ nat \ nat \ nat\ - proof(rule func.lam_type) - fix x - assume \x\nat \ nat\ - hence \fst(x)\nat\ by auto - thus \succ(fst(x)) \ nat\ by auto - qed -next - show \nat \ nat \ nat \ Pow((nat \ nat) \ nat)\ - by (rule pisubsig) -qed - -lemma addgdom : \nat \ nat \ domain(addg)\ -proof(unfold addg_def) - have e:\domain(\x\nat \ nat. succ(fst(x))) = nat \ nat\ - by (rule domain_lam) (* "domain(Lambda(A,b)) = A" *) - show \nat \ nat \ - domain(\x\nat \ nat. succ(fst(x)))\ - by (rule subst, rule sym, rule e, auto) -qed - -lemma plussucc: - assumes F:\f \ (nat\nat)\ - assumes H:\satpc(f,nat,addg)\ - shows \\n \ nat . f`succ(n) = succ(f`n)\ -proof - fix n - assume J:\n\nat\ - from H - have H:\\n \ nat . f`succ(n) = (\x\(nat*nat). succ(fst(x)))` \ - by (unfold satpc_def, unfold addg_def) - have H:\f`succ(n) = (\x\(nat*nat). succ(fst(x)))` \ - by (rule bspec[OF H J]) - have Q:\(\x\(nat*nat). succ(fst(x)))` = succ(fst())\ - proof(rule func.beta) - show \\f ` n, n\ \ nat \ nat\ - proof - show \f ` n \ nat\ - by (rule func.apply_funtype[OF F J]) - show \n \ nat\ - by (rule J) - qed - qed - have HQ:\f`succ(n) = succ(fst())\ - by (rule trans[OF H Q]) - have K:\fst() = f`n\ - by auto - hence K:\succ(fst()) = succ(f`n)\ - by (rule subst_context) - show \f`succ(n) = succ(f`n)\ - by (rule trans[OF HQ K]) -qed - -section "Definition of addition" - -text \Theorem that addition of natural numbers exists -and unique in some sense. Due to theorem 'plussucc' the term - \satpc(f,nat,addg)\ - can be replaced here with - \\n \ nat . f`succ(n) = succ(f`n)\.\ -theorem addition: - assumes \a\nat\ - shows - \\!f. ((f \ (nat\nat)) \ ((f`0) = a) \ satpc(f,nat,addg))\ -proof(rule recthm.recursionthm, unfold recthm_def) - show \a \ nat \ addg \ nat \ nat \ nat\ - proof - show \a\nat\ by (rule assms(1)) - next - show \addg \ nat \ nat \ nat\ - proof(unfold Pi_def, rule CollectI) - show \addg \ Pow((nat \ nat) \ nat)\ - by (rule addgsubpow) - next - have A2: \nat \ nat \ domain(addg)\ - by(rule addgdom) - have A3: \function(addg)\ - by (rule addgfun) - show \nat \ nat \ domain(addg) \ function(addg)\ - by(rule conjI[OF A2 A3]) - qed - qed -qed - -end +(* Title: Recursion theorem + Author: Georgy Dunaev , 2020 + Maintainer: Georgy Dunaev +*) +section "Recursion Submission" + +text \Recursion Theorem is proved in the following document. +It also contains the addition on natural numbers. +The development is done in the context of Zermelo-Fraenkel set theory.\ + +theory recursion + imports ZF +begin + +section \Basic Set Theory\ +text \Useful lemmas about sets, functions and natural numbers\ +lemma pisubsig : \Pi(A,P)\Pow(Sigma(A,P))\ +proof + fix x + assume \x \ Pi(A,P)\ + hence \x \ {f\Pow(Sigma(A,P)). A\domain(f) & function(f)}\ + by (unfold Pi_def) + thus \x \ Pow(Sigma(A, P))\ + by (rule CollectD1) +qed + +lemma apparg: + fixes f A B + assumes T0:\f:A\B\ + assumes T1:\f ` a = b\ + assumes T2:\a \ A\ + shows \\a, b\ \ f\ +proof(rule iffD2[OF func.apply_iff], rule T0) + show T:\a \ A \ f ` a = b\ + by (rule conjI[OF T2 T1]) +qed + +theorem nat_induct_bound : + assumes H0:\P(0)\ + assumes H1:\!!x. x\nat \ P(x) \ P(succ(x))\ + shows \\n\nat. P(n)\ +proof(rule ballI) + fix n + assume H2:\n\nat\ + show \P(n)\ + proof(rule nat_induct[of n]) + from H2 show \n\nat\ by assumption + next + show \P(0)\ by (rule H0) + next + fix x + assume H3:\x\nat\ + assume H4:\P(x)\ + show \P(succ(x))\ by (rule H1[OF H3 H4]) + qed +qed + +theorem nat_Tr : \\n\nat. m\n \ m\nat\ +proof(rule nat_induct_bound) + show \m \ 0 \ m \ nat\ by auto +next + fix x + assume H0:\x \ nat\ + assume H1:\m \ x \ m \ nat\ + show \m \ succ(x) \ m \ nat\ + proof(rule impI) + assume H2:\m\succ(x)\ + show \m \ nat\ + proof(rule succE[OF H2]) + assume H3:\m = x\ + from H0 and H3 show \m \ nat\ + by auto + next + assume H4:\m \ x\ + show \m \ nat\ + by(rule mp[OF H1 H4]) + qed + qed +qed + +(* Natural numbers are linearly ordered. *) +theorem zeroleq : \\n\nat. 0\n \ 0=n\ +proof(rule ballI) + fix n + assume H1:\n\nat\ + show \0\n\0=n\ + proof(rule nat_induct[of n]) + from H1 show \n \ nat\ by assumption + next + show \0 \ 0 \ 0 = 0\ by (rule disjI2, rule refl) + next + fix x + assume H2:\x\nat\ + assume H3:\ 0 \ x \ 0 = x\ + show \0 \ succ(x) \ 0 = succ(x)\ + proof(rule disjE[OF H3]) + assume H4:\0\x\ + show \0 \ succ(x) \ 0 = succ(x)\ + proof(rule disjI1) + show \0 \ succ(x)\ + by (rule succI2[OF H4]) + qed + next + assume H4:\0=x\ + show \0 \ succ(x) \ 0 = succ(x)\ + proof(rule disjI1) + have q:\x \ succ(x)\ by auto + from q and H4 show \0 \ succ(x)\ by auto + qed + qed + qed +qed + +theorem JH2_1ii : \m\succ(n) \ m\n\m=n\ + by auto + +theorem nat_transitive:\\n\nat. \k. \m. k \ m \ m \ n \ k \ n\ +proof(rule nat_induct_bound) + show \\k. \m. k \ m \ m \ 0 \ k \ 0\ + proof(rule allI, rule allI, rule impI) + fix k m + assume H:\k \ m \ m \ 0\ + then have H:\m \ 0\ by auto + then show \k \ 0\ by auto + qed +next + fix n + assume H0:\n \ nat\ + assume H1:\\k. + \m. + k \ m \ m \ n \ + k \ n\ + show \\k. \m. + k \ m \ + m \ succ(n) \ + k \ succ(n)\ + proof(rule allI, rule allI, rule impI) + fix k m + assume H4:\k \ m \ m \ succ(n)\ + hence H4':\m \ succ(n)\ by (rule conjunct2) + hence H4'':\m\n \ m=n\ by (rule succE, auto) + from H4 have Q:\k \ m\ by (rule conjunct1) + have H1S:\\m. k \ m \ m \ n \ k \ n\ + by (rule spec[OF H1]) + have H1S:\k \ m \ m \ n \ k \ n\ + by (rule spec[OF H1S]) + show \k \ succ(n)\ + proof(rule disjE[OF H4'']) + assume L:\m\n\ + from Q and L have QL:\k \ m \ m \ n\ by auto + have G:\k \ n\ by (rule mp [OF H1S QL]) + show \k \ succ(n)\ + by (rule succI2[OF G]) + next + assume L:\m=n\ + from Q have F:\k \ succ(m)\ by auto + from L and Q show \k \ succ(n)\ by auto + qed + qed +qed + +theorem nat_xninx : \\n\nat. \(n\n)\ +proof(rule nat_induct_bound) + show \0\0\ + by auto +next + fix x + assume H0:\x\nat\ + assume H1:\x\x\ + show \succ(x) \ succ(x)\ + proof(rule contrapos[OF H1]) + assume Q:\succ(x) \ succ(x)\ + have D:\succ(x)\x \ succ(x)=x\ + by (rule JH2_1ii[OF Q]) + show \x\x\ + proof(rule disjE[OF D]) + assume Y1:\succ(x)\x\ + have U:\x\succ(x)\ by (rule succI1) + have T:\x \ succ(x) \ succ(x) \ x \ x \ x\ + by (rule spec[OF spec[OF bspec[OF nat_transitive H0]]]) + have R:\x \ succ(x) \ succ(x) \ x\ + by (rule conjI[OF U Y1]) + show \x\x\ + by (rule mp[OF T R]) + next + assume Y1:\succ(x)=x\ + show \x\x\ + by (rule subst[OF Y1], rule Q) + qed + qed +qed + +theorem nat_asym : \\n\nat. \m. \(n\m \ m\n)\ +proof(rule ballI, rule allI) + fix n m + assume H0:\n \ nat\ + have Q:\\(n\n)\ + by(rule bspec[OF nat_xninx H0]) + show \\ (n \ m \ m \ n)\ + proof(rule contrapos[OF Q]) + assume W:\(n \ m \ m \ n)\ + show \n\n\ + by (rule mp[OF spec[OF spec[OF bspec[OF nat_transitive H0]]] W]) + qed +qed + +theorem zerolesucc :\\n\nat. 0 \ succ(n)\ +proof(rule nat_induct_bound) + show \0\1\ + by auto +next + fix x + assume H0:\x\nat\ + assume H1:\0\succ(x)\ + show \0\succ(succ(x))\ + proof + assume J:\0 \ succ(x)\ + show \0 = succ(x)\ + by(rule notE[OF J H1]) + qed +qed + +theorem succ_le : \\n\nat. succ(m)\succ(n) \ m\n\ +proof(rule nat_induct_bound) + show \ succ(m) \ 1 \ m \ 0\ + by blast +next + fix x + assume H0:\x \ nat\ + assume H1:\succ(m) \ succ(x) \ m \ x\ + show \ succ(m) \ + succ(succ(x)) \ + m \ succ(x)\ + proof(rule impI) + assume J0:\succ(m) \ succ(succ(x))\ + show \m \ succ(x)\ + proof(rule succE[OF J0]) + assume R:\succ(m) = succ(x)\ + hence R:\m=x\ by (rule upair.succ_inject) + from R and succI1 show \m \ succ(x)\ by auto + next + assume R:\succ(m) \ succ(x)\ + have R:\m\x\ by (rule mp[OF H1 R]) + then show \m \ succ(x)\ by auto + qed + qed +qed + +theorem succ_le2 : \\n\nat. \m. succ(m)\succ(n) \ m\n\ +proof + fix n + assume H:\n\nat\ + show \\m. succ(m) \ succ(n) \ m \ n\ + proof + fix m + from succ_le and H show \succ(m) \ succ(n) \ m \ n\ by auto + qed +qed + +theorem le_succ : \\n\nat. m\n \ succ(m)\succ(n)\ +proof(rule nat_induct_bound) + show \m \ 0 \ succ(m) \ 1\ + by auto +next + fix x + assume H0:\x\nat\ + assume H1:\m \ x \ succ(m) \ succ(x)\ + show \m \ succ(x) \ + succ(m) \ succ(succ(x))\ + proof(rule impI) + assume HR1:\m\succ(x)\ + show \succ(m) \ succ(succ(x))\ + proof(rule succE[OF HR1]) + assume Q:\m = x\ + from Q show \succ(m) \ succ(succ(x))\ + by auto + next + assume Q:\m \ x\ + have Q:\succ(m) \ succ(x)\ + by (rule mp[OF H1 Q]) + from Q show \succ(m) \ succ(succ(x))\ + by (rule succI2) + qed + qed +qed + +theorem nat_linord:\\n\nat. \m\nat. m\n\m=n\n\m\ +proof(rule ballI) + fix n + assume H1:\n\nat\ + show \\m\nat. m \ n \ m = n \ n \ m\ + proof(rule nat_induct[of n]) + from H1 show \n\nat\ by assumption + next + show \\m\nat. m \ 0 \ m = 0 \ 0 \ m\ + proof + fix m + assume J:\m\nat\ + show \ m \ 0 \ m = 0 \ 0 \ m\ + proof(rule disjI2) + have Q:\0\m\0=m\ by (rule bspec[OF zeroleq J]) + show \m = 0 \ 0 \ m\ + by (rule disjE[OF Q], auto) + qed + qed + next + fix x + assume K:\x\nat\ + assume M:\\m\nat. m \ x \ m = x \ x \ m\ + show \\m\nat. + m \ succ(x) \ + m = succ(x) \ + succ(x) \ m\ + proof(rule nat_induct_bound) + show \0 \ succ(x) \ 0 = succ(x) \ succ(x) \ 0\ + proof(rule disjI1) + show \0 \ succ(x)\ + by (rule bspec[OF zerolesucc K]) + qed + next + fix y + assume H0:\y \ nat\ + assume H1:\y \ succ(x) \ y = succ(x) \ succ(x) \ y\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + proof(rule disjE[OF H1]) + assume W:\y\succ(x)\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + proof(rule succE[OF W]) + assume G:\y=x\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + by (rule disjI2, rule disjI1, rule subst[OF G], rule refl) + next + assume G:\y \ x\ + have R:\succ(y) \ succ(x)\ + by (rule mp[OF bspec[OF le_succ K] G]) + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + by(rule disjI1, rule R) + qed + next + assume W:\y = succ(x) \ succ(x) \ y\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + proof(rule disjE[OF W]) + assume W:\y=succ(x)\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + by (rule disjI2, rule disjI2, rule subst[OF W], rule succI1) + next + assume W:\succ(x)\y\ + show \succ(y) \ succ(x) \ + succ(y) = succ(x) \ + succ(x) \ succ(y)\ + by (rule disjI2, rule disjI2, rule succI2[OF W]) + qed + qed + qed + qed +qed + +lemma tgb: + assumes knat: \k\nat\ + assumes D: \t \ k \ A\ + shows \t \ Pow(nat \ A)\ +proof - + from D + have q:\t\{t\Pow(Sigma(k,%_.A)). k\domain(t) & function(t)}\ + by(unfold Pi_def) + have J:\t \ Pow(k \ A)\ + by (rule CollectD1[OF q]) + have G:\k \ A \ nat \ A\ + proof(rule func.Sigma_mono) + from knat + show \k\nat\ + by (rule QUniv.naturals_subset_nat) + next + show \\x. x \ k \ A \ A\ + by auto + qed + show \t \ Pow(nat \ A)\ + by (rule subsetD, rule func.Pow_mono[OF G], rule J) +qed + +section \Compatible set\ +text \Union of compatible set of functions is a function.\ + +definition compat :: \[i,i]\o\ + where "compat(f1,f2) == \x.\y1.\y2.\x,y1\ \ f1 \ \x,y2\ \ f2 \ y1=y2" + +lemma compatI [intro]: + assumes H:\\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2\ + shows \compat(f1,f2)\ +proof(unfold compat_def) + show \\x y1 y2. \x, y1\ \ f1 \ \x, y2\ \ f2 \ y1 = y2\ + proof(rule allI | rule impI)+ + fix x y1 y2 + assume K:\\x, y1\ \ f1 \ \x, y2\ \ f2\ + have K1:\\x, y1\ \ f1\ by (rule conjunct1[OF K]) + have K2:\\x, y2\ \ f2\ by (rule conjunct2[OF K]) + show \y1 = y2\ by (rule H[OF K1 K2]) + qed +qed + +lemma compatD: + assumes H: \compat(f1,f2)\ + shows \\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2\ +proof - + fix x y1 y2 + assume Q1:\\x, y1\ \ f1\ + assume Q2:\\x, y2\ \ f2\ + from H have H:\\x y1 y2. \x, y1\ \ f1 \ \x, y2\ \ f2 \ y1 = y2\ + by (unfold compat_def) + show \y1=y2\ + proof(rule mp[OF spec[OF spec[OF spec[OF H]]]]) + show \\x, y1\ \ f1 \ \x, y2\ \ f2\ + by(rule conjI[OF Q1 Q2]) + qed +qed + +lemma compatE: + assumes H: \compat(f1,f2)\ + and W:\(\x y1 y2.\\x,y1\ \ f1; \x,y2\ \ f2\\y1=y2) \ E\ +shows \E\ + by (rule W, rule compatD[OF H], assumption+) + + +definition compatset :: \i\o\ + where "compatset(S) == \f1\S.\f2\S. compat(f1,f2)" + +lemma compatsetI [intro] : + assumes 1:\\f1 f2. \f1\S;f2\S\ \ compat(f1,f2)\ + shows \compatset(S)\ + by (unfold compatset_def, rule ballI, rule ballI, rule 1, assumption+) + +lemma compatsetD: + assumes H: \compatset(S)\ + shows \\f1 f2.\f1\S; f2\S\\compat(f1,f2)\ +proof - + fix f1 f2 + assume H1:\f1\S\ + assume H2:\f2\S\ + from H have H:\\f1\S.\f2\S. compat(f1,f2)\ + by (unfold compatset_def) + show \compat(f1,f2)\ + by (rule bspec[OF bspec[OF H H1] H2]) +qed + +lemma compatsetE: + assumes H: \compatset(S)\ + and W:\(\f1 f2.\f1\S; f2\S\\compat(f1,f2)) \ E\ +shows \E\ + by (rule W, rule compatsetD[OF H], assumption+) + +theorem upairI1 : \a \ {a, b}\ +proof + assume \a \ {b}\ + show \a = a\ by (rule refl) +qed + +theorem upairI2 : \b \ {a, b}\ +proof + assume H:\b \ {b}\ + have Y:\b \ {b}\ by (rule upair.singletonI) + show \b = a\ by (rule notE[OF H Y]) +qed + +theorem sinup : \{x} \ \x, xa\\ +proof (unfold Pair_def) + show \{x} \ {{x, x}, {x, xa}}\ + proof (rule IFOL.subst) + show \{x} \ {{x},{x,xa}}\ + by (rule upairI1) + next + show \{{x}, {x, xa}} = {{x, x}, {x, xa}}\ + by blast + qed +qed + +theorem compatsetunionfun : + fixes S + assumes H0:\compatset(S)\ + shows \function(\S)\ +proof(unfold function_def) + show \ \x y1. \x, y1\ \ \S \ + (\y2. \x, y2\ \ \S \ y1 = y2)\ + proof(rule allI, rule allI, rule impI, rule allI, rule impI) + fix x y1 y2 + assume F1:\\x, y1\ \ \S\ + assume F2:\\x, y2\ \ \S\ + show \y1=y2\ + proof(rule UnionE[OF F1], rule UnionE[OF F2]) + fix f1 f2 + assume J1:\\x, y1\ \ f1\ + assume J2:\\x, y2\ \ f2\ + assume K1:\f1 \ S\ + assume K2:\f2 \ S\ + have R:\compat(f1,f2)\ + by (rule compatsetD[OF H0 K1 K2]) + show \y1=y2\ + by(rule compatD[OF R J1 J2]) + qed + qed +qed + +theorem mkel : + assumes 1:\A\ + assumes 2:\A\B\ + shows \B\ + by (rule 2, rule 1) + +theorem valofunion : + fixes S + assumes H0:\compatset(S)\ + assumes W:\f\S\ + assumes Q:\f:A\B\ + assumes T:\a\A\ + assumes P:\f ` a = v\ + shows N:\(\S)`a = v\ +proof - + have K:\\a, v\ \ f\ + by (rule apparg[OF Q P T]) + show N:\(\S)`a = v\ + proof(rule function_apply_equality) + show \function(\S)\ + by(rule compatsetunionfun[OF H0]) + next + show \\a, v\ \ \S\ + by(rule UnionI[OF W K ]) + qed +qed + +section "Partial computation" + +definition satpc :: \[i,i,i] \ o \ + where \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ + +text \$m$-step computation based on $a$ and $g$\ +definition partcomp :: \[i,i,i,i,i]\o\ + where \partcomp(A,t,m,a,g) == (t:succ(m)\A) \ (t`0=a) \ satpc(t,m,g)\ + +lemma partcompI [intro]: + assumes H1:\(t:succ(m)\A)\ + assumes H2:\(t`0=a)\ + assumes H3:\satpc(t,m,g)\ + shows \partcomp(A,t,m,a,g)\ +proof (unfold partcomp_def, auto) + show \t \ succ(m) \ A\ by (rule H1) + show \(t`0=a)\ by (rule H2) + show \satpc(t,m,g)\ by (rule H3) +qed + +lemma partcompD1: \partcomp(A,t,m,a,g) \ t \ succ(m) \ A\ + by (unfold partcomp_def, auto) + +lemma partcompD2: \partcomp(A,t,m,a,g) \ (t`0=a)\ + by (unfold partcomp_def, auto) + +lemma partcompD3: \partcomp(A,t,m,a,g) \ satpc(t,m,g)\ + by (unfold partcomp_def, auto) + +lemma partcompE [elim] : + assumes 1:\partcomp(A,t,m,a,g)\ + and 2:\\(t:succ(m)\A) ; (t`0=a) ; satpc(t,m,g)\ \ E\ + shows \E\ + by (rule 2, rule partcompD1[OF 1], rule partcompD2[OF 1], rule partcompD3[OF 1]) + +text \If we add ordered pair in the middle of partial computation then +it will not change.\ +lemma addmiddle: +(* fixes t m a g*) + assumes mnat:\m\nat\ + assumes F:\partcomp(A,t,m,a,g)\ + assumes xinm:\x\m\ + shows \cons(\succ(x), g ` \t ` x, x\\, t) = t\ +proof(rule partcompE[OF F]) + assume F1:\t \ succ(m) \ A\ + assume F2:\t ` 0 = a\ + assume F3:\satpc(t, m, g)\ + from F3 + have W:\\n\m. t ` succ(n) = g ` \t ` n, n\\ + by (unfold satpc_def) + have U:\t ` succ(x) = g ` \t ` x, x\\ + by (rule bspec[OF W xinm]) + have E:\\succ(x), (g ` \t ` x, x\)\ \ t\ + proof(rule apparg[OF F1 U]) + show \succ(x) \ succ(m)\ + by(rule mp[OF bspec[OF le_succ mnat] xinm]) + qed + show ?thesis + by (rule equalities.cons_absorb[OF E]) +qed + + +section \Set of functions \ +text \It is denoted as $F$ on page 48 in "Introduction to Set Theory".\ +definition pcs :: \[i,i,i]\i\ + where \pcs(A,a,g) == {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ + +lemma pcs_uniq : + assumes F1:\m1\nat\ + assumes F2:\m2\nat\ + assumes H1: \partcomp(A,f1,m1,a,g)\ + assumes H2: \partcomp(A,f2,m2,a,g)\ + shows \\n\nat. n\succ(m1) \ n\succ(m2) \ f1`n = f2`n\ +proof(rule partcompE[OF H1], rule partcompE[OF H2]) + assume H11:\f1 \ succ(m1) \ A\ + assume H12:\f1 ` 0 = a \ + assume H13:\satpc(f1, m1, g)\ + assume H21:\f2 \ succ(m2) \ A\ + assume H22:\f2 ` 0 = a\ + assume H23:\satpc(f2, m2, g)\ + show \\n\nat. n\succ(m1) \ n\succ(m2) \ f1`n = f2`n\ +proof(rule nat_induct_bound) + from H12 and H22 + show \0\succ(m1) \ 0\succ(m2) \ f1 ` 0 = f2 ` 0\ + by auto +next + fix x + assume J0:\x\nat\ + assume J1:\x \ succ(m1) \ x \ succ(m2) \ f1 ` x = f2 ` x\ + from H13 have G1:\\n \ m1 . f1`succ(n) = g ` \ + by (unfold satpc_def, auto) + from H23 have G2:\\n \ m2 . f2`succ(n) = g ` \ + by (unfold satpc_def, auto) + show \succ(x) \ succ(m1) \ succ(x) \ succ(m2) \ + f1 ` succ(x) = f2 ` succ(x)\ + proof + assume K:\succ(x) \ succ(m1) \ succ(x) \ succ(m2)\ + from K have K1:\succ(x) \ succ(m1)\ by auto + from K have K2:\succ(x) \ succ(m2)\ by auto + have K1':\x \ m1\ by (rule mp[OF bspec[OF succ_le F1] K1]) + have K2':\x \ m2\ by (rule mp[OF bspec[OF succ_le F2] K2]) + have U1:\x\succ(m1)\ + by (rule Nat.succ_in_naturalD[OF K1 Nat.nat_succI[OF F1]]) + have U2:\x\succ(m2)\ + by (rule Nat.succ_in_naturalD[OF K2 Nat.nat_succI[OF F2]]) + have Y1:\f1`succ(x) = g ` \ + by (rule bspec[OF G1 K1']) + have Y2:\f2`succ(x) = g ` \ + by (rule bspec[OF G2 K2']) + have \f1 ` x = f2 ` x\ + by(rule mp[OF J1 conjI[OF U1 U2]]) + then have Y:\g ` = g ` \ by auto + from Y1 and Y2 and Y + show \f1 ` succ(x) = f2 ` succ(x)\ + by auto + qed +qed +qed + +lemma domainsubsetfunc : + assumes Q:\f1\f2\ + shows \domain(f1)\domain(f2)\ +proof + fix x + assume H:\x \ domain(f1)\ + show \x \ domain(f2)\ + proof(rule domainE[OF H]) + fix y + assume W:\\x, y\ \ f1\ + have \\x, y\ \ f2\ + by(rule subsetD[OF Q W]) + then show \x \ domain(f2)\ + by(rule domainI) + qed +qed + +lemma natdomfunc: + assumes 1:\q\A\ + assumes J0:\f1 \ Pow(nat \ A)\ + assumes U:\m1 \ domain(f1)\ + shows \m1\nat\ +proof - + from J0 have J0 : \f1 \ nat \ A\ + by auto + have J0:\domain(f1) \ domain(nat \ A)\ + by(rule func.domain_mono[OF J0]) + have F:\m1 \ domain(nat \ A)\ + by(rule subsetD[OF J0 U]) + have R:\domain(nat \ A) = nat\ + by (rule equalities.domain_of_prod[OF 1]) + show \m1 \ nat\ + by(rule subst[OF R], rule F) +qed + +lemma pcs_lem : + assumes 1:\q\A\ + shows \compatset(pcs(A, a, g))\ +proof (*(rule compatsetI)*) + fix f1 f2 + assume H1:\f1 \ pcs(A, a, g)\ + then have H1':\f1 \ {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ by (unfold pcs_def) + hence H1'A:\f1 \ Pow(nat*A)\ by auto + hence H1'A:\f1 \ (nat*A)\ by auto + assume H2:\f2 \ pcs(A, a, g)\ + then have H2':\f2 \ {t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ by (unfold pcs_def) + show \compat(f1, f2)\ + proof(rule compatI) + fix x y1 y2 + assume P1:\\x, y1\ \ f1\ + assume P2:\\x, y2\ \ f2\ + show \y1 = y2\ + proof(rule CollectE[OF H1'], rule CollectE[OF H2']) + assume J0:\f1 \ Pow(nat \ A)\ + assume J1:\f2 \ Pow(nat \ A)\ + assume J2:\\m\nat. partcomp(A, f1, m, a, g)\ + assume J3:\\m\nat. partcomp(A, f2, m, a, g)\ + show \y1 = y2\ + proof(rule bexE[OF J2], rule bexE[OF J3]) + fix m1 m2 + assume K1:\partcomp(A, f1, m1, a, g)\ + assume K2:\partcomp(A, f2, m2, a, g)\ + hence K2':\(f2:succ(m2)\A) \ (f2`0=a) \ satpc(f2,m2,g)\ + by (unfold partcomp_def) + from K1 have K1'A:\(f1:succ(m1)\A)\ by (rule partcompD1) + from K2' have K2'A:\(f2:succ(m2)\A)\ by auto + from K1'A have K1'AD:\domain(f1) = succ(m1)\ + by(rule domain_of_fun) + from K2'A have K2'AD:\domain(f2) = succ(m2)\ + by(rule domain_of_fun) + have L1:\f1`x=y1\ + by (rule func.apply_equality[OF P1], rule K1'A) + have L2:\f2`x=y2\ + by(rule func.apply_equality[OF P2], rule K2'A) + have m1nat:\m1\nat\ + proof(rule natdomfunc[OF 1 J0]) + show \m1 \ domain(f1)\ + by (rule ssubst[OF K1'AD], auto) + qed + have m2nat:\m2\nat\ + proof(rule natdomfunc[OF 1 J1]) + show \m2 \ domain(f2)\ + by (rule ssubst[OF K2'AD], auto) + qed + have G1:\\x, y1\ \ (nat*A)\ + by(rule subsetD[OF H1'A P1]) + have KK:\x\nat\ + by(rule SigmaE[OF G1], auto) + (*x is in the domain of f1 i.e. succ(m1) +so we can have both x \ ?m1.2 \ x \ ?m2.2 +how to prove that m1 \ nat ? from J0 ! f1 is a subset of nat \ A*) + have W:\f1`x=f2`x\ + proof(rule mp[OF bspec[OF pcs_uniq KK] ]) + show \m1 \ nat\ + by (rule m1nat) + next + show \m2 \ nat\ + by (rule m2nat) + next + show \partcomp(A, f1, m1, a, g)\ + by (rule K1) + next + show \partcomp(A, f2, m2, a, g)\ + by (rule K2) + next + (* P1:\\x, y1\ \ f1\ + K1'A:\(f1:succ(m1)\A)\ + *) + have U1:\x \ succ(m1)\ + by (rule func.domain_type[OF P1 K1'A]) + have U2:\x \ succ(m2)\ + by (rule func.domain_type[OF P2 K2'A]) + show \x \ succ(m1) \ x \ succ(m2)\ + by (rule conjI[OF U1 U2]) + qed + from L1 and W and L2 + show \y1 = y2\ by auto + qed + qed + qed +qed + +theorem fuissu : \f \ X -> Y \ f \ X\Y\ +proof + fix w + assume H1 : \f \ X -> Y\ + then have J1:\f \ {q\Pow(Sigma(X,\_.Y)). X\domain(q) & function(q)}\ + by (unfold Pi_def) + then have J2:\f \ Pow(Sigma(X,\_.Y))\ + by auto + then have J3:\f \ Sigma(X,\_.Y)\ + by auto + assume H2 : \w \ f\ + from J3 and H2 have \w\Sigma(X,\_.Y)\ + by auto + then have J4:\w \ (\x\X. (\y\Y. {\x,y\}))\ + by auto + show \w \ X*Y\ + proof (rule UN_E[OF J4]) + fix x + assume V1:\x \ X\ + assume V2:\w \ (\y\Y. {\x, y\})\ + show \w \ X \ Y\ + proof (rule UN_E[OF V2]) + fix y + assume V3:\y \ Y\ + assume V4:\w \ {\x, y\}\ + then have V4:\w = \x, y\\ + by auto + have v5:\\x, y\ \ Sigma(X,\_.Y)\ + proof(rule SigmaI) + show \x \ X\ by (rule V1) + next + show \y \ Y\ by (rule V3) + qed + then have V5:\\x, y\ \ X*Y\ + by auto + from V4 and V5 show \w \ X \ Y\ by auto + qed + qed +qed + +theorem recuniq : + fixes f + assumes H0:\f \ nat -> A \ f ` 0 = a \ satpc(f, nat, g)\ + fixes t + assumes H1:\t \ nat -> A \ t ` 0 = a \ satpc(t, nat, g)\ + fixes x + shows \f=t\ +proof - + from H0 have H02:\\n \ nat. f`succ(n) = g ` <(f`n), n>\ by (unfold satpc_def, auto) + from H0 have H01:\f ` 0 = a\ by auto + from H0 have H00:\f \ nat -> A\ by auto + from H1 have H12:\\n \ nat. t`succ(n) = g ` <(t`n), n>\ by (unfold satpc_def, auto) + from H1 have H11:\t ` 0 = a\ by auto + from H1 have H10:\t \ nat -> A\ by auto + show \f=t\ + proof (rule fun_extension[OF H00 H10]) + fix x + assume K: \x \ nat\ + show \(f ` x) = (t ` x)\ + proof(rule nat_induct[of x]) + show \x \ nat\ by (rule K) + next + from H01 and H11 show \f ` 0 = t ` 0\ + by auto + next + fix x + assume A:\x\nat\ + assume B:\f`x = t`x\ + show \f ` succ(x) = t ` succ(x)\ + proof - + from H02 and A have H02':\f`succ(x) = g ` <(f`x), x>\ + by (rule bspec) + from H12 and A have H12':\t`succ(x) = g ` <(t`x), x>\ + by (rule bspec) + from B and H12' have H12'':\t`succ(x) = g ` <(f`x), x>\ by auto + from H12'' and H02' show \f ` succ(x) = t ` succ(x)\ by auto + qed + qed + qed +qed + +section \Lemmas for recursion theorem\ + +locale recthm = + fixes A :: "i" + and a :: "i" + and g :: "i" + assumes hyp1 : \a \ A\ + and hyp2 : \g : ((A*nat)\A)\ +begin + +lemma l3:\function(\pcs(A, a, g))\ + by (rule compatsetunionfun, rule pcs_lem, rule hyp1) + +lemma l1 : \\pcs(A, a, g) \ nat \ A\ +proof + fix x + assume H:\x \ \pcs(A, a, g)\ + hence H:\x \ \{t\Pow(nat*A). \m\nat. partcomp(A,t,m,a,g)}\ + by (unfold pcs_def) + show \x \ nat \ A\ + proof(rule UnionE[OF H]) + fix B + assume J1:\x\B\ + assume J2:\B \ {t \ Pow(nat \ A) . + \m\nat. partcomp(A, t, m, a, g)}\ + hence J2:\B \ Pow(nat \ A)\ by auto + hence J2:\B \ nat \ A\ by auto + from J1 and J2 show \x \ nat \ A\ + by auto + qed +qed + +lemma le1: + assumes H:\x\1\ + shows \x=0\ +proof + show \x \ 0\ + proof + fix z + assume J:\z\x\ + show \z\0\ + proof(rule succE[OF H]) + assume J:\x\0\ + show \z\0\ + by (rule notE[OF not_mem_empty J]) + next + assume K:\x=0\ + from J and K show \z\0\ + by auto + qed + qed +next + show \0 \ x\ by auto +qed + +lemma lsinglfun : \function({\0, a\})\ +proof(unfold function_def) + show \ \x y. \x, y\ \ {\0, a\} \ + (\y'. \x, y'\ \ {\0, a\} \ + y = y')\ + proof(rule allI,rule allI,rule impI,rule allI,rule impI) + fix x y y' + assume H0:\\x, y\ \ {\0, a\}\ + assume H1:\\x, y'\ \ {\0, a\}\ + show \y = y'\ + proof(rule upair.singletonE[OF H0],rule upair.singletonE[OF H1]) + assume H0:\\x, y\ = \0, a\\ + assume H1:\\x, y'\ = \0, a\\ + from H0 and H1 have H:\\x, y\ = \x, y'\\ by auto + then show \y = y'\ by auto + qed + qed +qed + +lemma singlsatpc:\satpc({\0, a\}, 0, g)\ +proof(unfold satpc_def) + show \\n\0. {\0, a\} ` succ(n) = + g ` \{\0, a\} ` n, n\\ + by auto +qed + +lemma zerostep : + shows \partcomp(A, {\0, a\}, 0, a, g)\ +proof(unfold partcomp_def) + show \{\0, a\} \ 1 -> A \ {\0, a\} ` 0 = a \ satpc({\0, a\}, 0, g)\ + proof + show \{\0, a\} \ 1 -> A\ + proof (unfold Pi_def) + show \{\0, a\} \ {f \ Pow(1 \ A) . 1 \ domain(f) \ function(f)}\ + proof + show \{\0, a\} \ Pow(1 \ A)\ + proof(rule PowI, rule equalities.singleton_subsetI) + show \\0, a\ \ 1 \ A\ + proof + show \0 \ 1\ by auto + next + show \a \ A\ by (rule hyp1) + qed + qed + next + show \1 \ domain({\0, a\}) \ function({\0, a\})\ + proof + show \1 \ domain({\0, a\})\ + proof + fix x + assume W:\x\1\ + from W have W:\x=0\ by (rule le1) + have Y:\0\domain({\0, a\})\ + by auto + from W and Y + show \x\domain({\0, a\})\ + by auto + qed + next + show \function({\0, a\})\ + by (rule lsinglfun) + qed + qed + qed + show \{\0, a\} ` 0 = a \ satpc({\0, a\}, 0, g)\ + proof + show \{\0, a\} ` 0 = a\ + by (rule func.singleton_apply) + next + show \satpc({\0, a\}, 0, g)\ + by (rule singlsatpc) + qed + qed +qed + +lemma zainupcs : \\0, a\ \ \pcs(A, a, g)\ +proof + show \\0, a\ \ {\0, a\}\ + by auto +next + (* {\0, a\} is a 0-step computation *) + show \{\0, a\} \ pcs(A, a, g)\ + proof(unfold pcs_def) + show \{\0, a\} \ {t \ Pow(nat \ A) . \m\nat. partcomp(A, t, m, a, g)}\ + proof + show \{\0, a\} \ Pow(nat \ A)\ + proof(rule PowI, rule equalities.singleton_subsetI) + show \\0, a\ \ nat \ A\ + proof + show \0 \ nat\ by auto + next + show \a \ A\ by (rule hyp1) + qed + qed + next + show \\m\nat. partcomp(A, {\0, a\}, m, a, g)\ + proof + show \partcomp(A, {\0, a\}, 0, a, g)\ + by (rule zerostep) + next + show \0 \ nat\ by auto + qed + qed + qed +qed + +lemma l2': \0 \ domain(\pcs(A, a, g))\ +proof + show \\0, a\ \ \pcs(A, a, g)\ + by (rule zainupcs) +qed + +text \Push an ordered pair to the end of partial computation t +and obtain another partial computation.\ +lemma shortlem : + assumes mnat:\m\nat\ + assumes F:\partcomp(A,t,m,a,g)\ + shows \partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ +proof(rule partcompE[OF F]) + assume F1:\t \ succ(m) \ A\ + assume F2:\t ` 0 = a\ + assume F3:\satpc(t, m, g)\ + show ?thesis (*\partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ *) + proof + have ljk:\cons(\succ(m), g ` \t ` m, m\\, t) \ (cons(succ(m),succ(m)) \ A)\ + proof(rule func.fun_extend3[OF F1]) + show \succ(m) \ succ(m)\ + by (rule upair.mem_not_refl) + have tmA:\t ` m \ A\ + by (rule func.apply_funtype[OF F1], auto) + show \g ` \t ` m, m\ \ A\ + by(rule func.apply_funtype[OF hyp2], auto, rule tmA, rule mnat) + qed + have \cons(\succ(m), g ` \t ` m, m\\, t) \ (cons(succ(m),succ(m)) \ A)\ + by (rule ljk) + then have \cons(\cons(m, m), g ` \t ` m, m\\, t) \ cons(cons(m, m), cons(m, m)) \ A\ + by (unfold succ_def) + then show \cons(\succ(m), g ` \t ` m, m\\, t) \ succ(succ(m)) \ A\ + by (unfold succ_def, assumption) + show \cons(\succ(m), g ` \t ` m, m\\, t) ` 0 = a\ + proof(rule trans, rule func.fun_extend_apply[OF F1]) + show \succ(m) \ succ(m)\ by (rule upair.mem_not_refl) + show \(if 0 = succ(m) then g ` \t ` m, m\ else t ` 0) = a\ + by(rule trans, rule upair.if_not_P, auto, rule F2) + qed + show \satpc(cons(\succ(m), g ` \t ` m, m\\, t), succ(m), g)\ + proof(unfold satpc_def, rule ballI) + fix n + assume Q:\n \ succ(m)\ + show \cons(\succ(m), g ` \t ` m, m\\, t) ` succ(n) += g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + proof(rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) + show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = + g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + proof(rule upair.succE[OF Q]) + assume Y:\n=m\ + show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = + g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + proof(rule trans, rule upair.if_P) + from Y show \succ(n) = succ(m)\ by auto + next + have L1:\t ` m = cons(\succ(m), g ` \t ` m, m\\, t) ` n\ + proof(rule sym, rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) + show \ (if n = succ(m) then g ` \t ` m, m\ else t ` n) = t ` m\ + proof(rule trans, rule upair.if_not_P) + from Y show \t ` n = t ` m\ by auto + show \n \ succ(m)\ + proof(rule not_sym) + show \succ(m) \ n\ + by(rule subst, rule sym, rule Y, rule upair.succ_neq_self) + qed + qed + qed + from Y + have L2:\m = n\ + by auto + have L:\ \t ` m, m\ = \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + by(rule subst_context2[OF L1 L2]) + show \ g ` \t ` m, m\ = g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + by(rule subst_context[OF L]) + qed + next + assume Y:\n \ m\ + show \(if succ(n) = succ(m) then g ` \t ` m, m\ else t ` succ(n)) = + g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + proof(rule trans, rule upair.if_not_P) + show \succ(n) \ succ(m)\ + by(rule contrapos, rule upair.mem_imp_not_eq, rule Y, rule upair.succ_inject, assumption) + next + have X:\cons(\succ(m), g ` \t ` m, m\\, t) ` n = t ` n\ + proof(rule trans, rule func.fun_extend_apply[OF F1], rule upair.mem_not_refl) + show \(if n = succ(m) then g ` \t ` m, m\ else t ` n) = t ` n\ + proof(rule upair.if_not_P) + show \n \ succ(m)\ + proof(rule contrapos) + assume q:"n=succ(m)" + from q and Y have M:\succ(m)\m\ + by auto + show \m\m\ + by(rule Nat.succ_in_naturalD[OF M mnat]) + next + show \m \ m\ by (rule upair.mem_not_refl) + qed + qed + qed + from F3 + have W:\\n\m. t ` succ(n) = g ` \t ` n, n\\ + by (unfold satpc_def) + have U:\t ` succ(n) = g ` \t ` n, n\\ + by (rule bspec[OF W Y]) + show \t ` succ(n) = g ` \cons(\succ(m), g ` \t ` m, m\\, t) ` n, n\\ + by (rule trans, rule U, rule sym, rule subst_context[OF X]) + qed + qed + qed + qed + qed +qed + +lemma l2:\nat \ domain(\pcs(A, a, g))\ +proof + fix x + assume G:\x\nat\ + show \x \ domain(\pcs(A, a, g))\ + proof(rule nat_induct[of x]) + show \x\nat\ by (rule G) + next + fix x + assume Q1:\x\nat\ + assume Q2:\x\domain(\pcs(A, a, g))\ + show \succ(x)\domain(\pcs(A, a, g))\ + proof(rule domainE[OF Q2]) + fix y + assume W1:\\x, y\ \ (\pcs(A, a, g))\ + show \succ(x)\domain(\pcs(A, a, g))\ + proof(rule UnionE[OF W1]) + fix t + assume E1:\\x, y\ \ t\ + assume E2:\t \ pcs(A, a, g)\ + hence E2:\t\{t\Pow(nat*A). \m \ nat. partcomp(A,t,m,a,g)}\ + by(unfold pcs_def) + have E21:\t\Pow(nat*A)\ + by(rule CollectD1[OF E2]) + have E22m:\\m\nat. partcomp(A,t,m,a,g)\ + by(rule CollectD2[OF E2]) + show \succ(x)\domain(\pcs(A, a, g))\ + proof(rule bexE[OF E22m]) + fix m + assume mnat:\m\nat\ + assume E22P:\partcomp(A,t,m,a,g)\ + hence E22:\((t:succ(m)\A) \ (t`0=a)) \ satpc(t,m,g)\ + by(unfold partcomp_def, auto) + hence E223:\satpc(t,m,g)\ by auto + hence E223:\\n \ m . t`succ(n) = g ` \ + by(unfold satpc_def, auto) + from E22 have E221:\(t:succ(m)\A)\ + by auto + from E221 have domt:\domain(t) = succ(m)\ + by (rule func.domain_of_fun) + from E1 have xind:\x \ domain(t)\ + by (rule equalities.domainI) + from xind and domt have xinsm:\x \ succ(m)\ + by auto + show \succ(x)\domain(\pcs(A, a, g))\ + proof + (*proof(rule exE[OF E22])*) + show \ \succ(x), g ` \ \ (\pcs(A, a, g))\ (*?*) + proof + (*t\{\succ(x), g ` \}*) + show \cons(\succ(x), g ` \, t) \ pcs(A, a, g)\ + proof(unfold pcs_def, rule CollectI) + from E21 + have L1:\t \ nat \ A\ + by auto + from Q1 have J1:\succ(x)\nat\ + by auto(*Nat.nat_succI*) + have txA: \t ` x \ A\ + by (rule func.apply_type[OF E221 xinsm]) + from txA and Q1 have txx:\\t ` x, x\ \ A \ nat\ + by auto + have secp: \g ` \t ` x, x\ \ A\ + by(rule func.apply_type[OF hyp2 txx]) + from J1 and secp + have L2:\\succ(x),g ` \t ` x, x\\ \ nat \ A\ + by auto + show \ cons(\succ(x),g ` \t ` x, x\\,t) \ Pow(nat \ A)\ + proof(rule PowI) + show \ cons(\succ(x), g ` \t ` x, x\\, t) \ nat \ A\ + proof + show \\succ(x), g ` \t ` x, x\\ \ nat \ A \ t \ nat \ A\ + by (rule conjI[OF L2 L1]) + qed + qed + next + show \\m \ nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ + proof(rule succE[OF xinsm]) + assume xeqm:\x=m\ + show \\m \ nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ + proof + show \partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), succ(x), a, g)\ + proof(rule shortlem[OF Q1]) + show \partcomp(A, t, x, a, g)\ + proof(rule subst[of m x], rule sym, rule xeqm) + show \partcomp(A, t, m, a, g)\ + by (rule E22P) + qed + qed + next + from Q1 show \succ(x) \ nat\ by auto + qed + next + assume xinm:\x\m\ + have lmm:\cons(\succ(x), g ` \t ` x, x\\, t) = t\ + by (rule addmiddle[OF mnat E22P xinm]) + show \\m\nat. partcomp(A, cons(\succ(x), g ` \t ` x, x\\, t), m, a, g)\ + by(rule subst[of t], rule sym, rule lmm, rule E22m) + qed + qed + next + show \\succ(x), g ` \t ` x, x\\ \ cons(\succ(x), g ` \t ` x, x\\, t)\ + by auto + qed + qed + qed + qed + qed + next + show \0 \ domain(\pcs(A, a, g))\ + by (rule l2') + qed +qed + +lemma useful : \\m\nat. \t. partcomp(A,t,m,a,g)\ +proof(rule nat_induct_bound) + show \\t. partcomp(A, t, 0, a, g)\ + proof + show \partcomp(A, {\0, a\}, 0, a, g)\ + by (rule zerostep) + qed +next + fix m + assume mnat:\m\nat\ + assume G:\\t. partcomp(A,t,m,a,g)\ + show \\t. partcomp(A,t,succ(m),a,g)\ + proof(rule exE[OF G]) + fix t + assume G:\partcomp(A,t,m,a,g)\ + show \\t. partcomp(A,t,succ(m),a,g)\ + proof + show \partcomp(A,cons(\succ(m), g ` \, t),succ(m),a,g)\ + by(rule shortlem[OF mnat G]) + qed + qed +qed + +lemma l4 : \(\pcs(A,a,g)) \ nat -> A\ +proof(unfold Pi_def) + show \ \pcs(A, a, g) \ {f \ Pow(nat \ A) . nat \ domain(f) \ function(f)}\ + proof + show \\pcs(A, a, g) \ Pow(nat \ A)\ + proof + show \\pcs(A, a, g) \ nat \ A\ + by (rule l1) + qed + next + show \nat \ domain(\pcs(A, a, g)) \ function(\pcs(A, a, g))\ + proof + show \nat \ domain(\pcs(A, a, g))\ + by (rule l2) + next + show \function(\pcs(A, a, g))\ + by (rule l3) + qed + qed +qed + +lemma l5: \(\pcs(A, a, g)) ` 0 = a\ +proof(rule func.function_apply_equality) + show \function(\pcs(A, a, g))\ + by (rule l3) +next + show \\0, a\ \ \pcs(A, a, g)\ + by (rule zainupcs) +qed + +lemma ballE2: + assumes \\x\AA. P(x)\ + assumes \x\AA\ + assumes \P(x) ==> Q\ + shows Q + by (rule assms(3), rule bspec, rule assms(1), rule assms(2)) + +text \ Recall that + \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ + \partcomp(A,t,m,a,g) == (t:succ(m)\A) \ (t`0=a) \ satpc(t,m,g)\ + \pcs(A,a,g) == {t\Pow(nat*A). \m. partcomp(A,t,m,a,g)}\ +\ + +lemma l6new: \satpc(\pcs(A, a, g), nat, g)\ +proof (unfold satpc_def, rule ballI) + fix n + assume nnat:\n\nat\ + hence snnat:\succ(n)\nat\ by auto + (* l2:\nat \ domain(\pcs(A, a, g))\ *) + show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ + proof(rule ballE2[OF useful snnat], erule exE) + fix t + assume Y:\partcomp(A, t, succ(n), a, g)\ + show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ + proof(rule partcompE[OF Y]) + assume Y1:\t \ succ(succ(n)) \ A\ + assume Y2:\t ` 0 = a\ + assume Y3:\satpc(t, succ(n), g)\ + hence Y3:\\x \ succ(n) . t`succ(x) = g ` \ + by (unfold satpc_def) + hence Y3:\t`succ(n) = g ` \ + by (rule bspec, auto) + have e1:\(\pcs(A, a, g)) ` succ(n) = t ` succ(n)\ + proof(rule valofunion, rule pcs_lem, rule hyp1) + show \t \ pcs(A, a, g)\ + proof(unfold pcs_def, rule CollectI) + show \t \ Pow(nat \ A)\ + proof(rule tgb) + show \t \ succ(succ(n)) \ A\ by (rule Y1) + next + from snnat + show \succ(succ(n)) \ nat\ by auto + qed + next + show \\m\nat. partcomp(A, t, m, a, g)\ + by(rule bexI, rule Y, rule snnat) + qed + next + show \t \ succ(succ(n)) \ A\ by (rule Y1) + next + show \succ(n) \ succ(succ(n))\ by auto + next + show \t ` succ(n) = t ` succ(n)\ by (rule refl) + qed + have e2:\(\pcs(A, a, g)) ` n = t ` n\ + proof(rule valofunion, rule pcs_lem, rule hyp1) + show \t \ pcs(A, a, g)\ + proof(unfold pcs_def, rule CollectI) + show \t \ Pow(nat \ A)\ + proof(rule tgb) + show \t \ succ(succ(n)) \ A\ by (rule Y1) + next + from snnat + show \succ(succ(n)) \ nat\ by auto + qed + next + show \\m\nat. partcomp(A, t, m, a, g)\ + by(rule bexI, rule Y, rule snnat) + qed + next + show \t \ succ(succ(n)) \ A\ by (rule Y1) + next + show \n \ succ(succ(n))\ by auto + next + show \t ` n = t ` n\ by (rule refl) + qed + have e3:\g ` \(\pcs(A, a, g)) ` n, n\ = g ` \t ` n, n\\ + by (rule subst[OF e2], rule refl) + show \(\pcs(A, a, g)) ` succ(n) = g ` \(\pcs(A, a, g)) ` n, n\\ + by (rule trans, rule e1,rule trans, rule Y3, rule sym, rule e3) + qed + qed +qed + +section "Recursion theorem" + +theorem recursionthm: + shows \\!f. ((f \ (nat\A)) \ ((f`0) = a) \ satpc(f,nat,g))\ +(* where \satpc(t,\,g) == \n \ \ . t`succ(n) = g ` \ *) +proof + show \\f. f \ nat -> A \ f ` 0 = a \ satpc(f, nat, g)\ + proof + show \(\pcs(A,a,g)) \ nat -> A \ (\pcs(A,a,g)) ` 0 = a \ satpc(\pcs(A,a,g), nat, g)\ + proof + show \\pcs(A, a, g) \ nat -> A\ + by (rule l4) + next + show \(\pcs(A, a, g)) ` 0 = a \ satpc(\pcs(A, a, g), nat, g)\ + proof + show \(\pcs(A, a, g)) ` 0 = a\ + by (rule l5) + next + show \satpc(\pcs(A, a, g), nat, g)\ + by (rule l6new) + qed + qed + qed +next + show \\f y. f \ nat -> A \ + f ` 0 = a \ + satpc(f, nat, g) \ + y \ nat -> A \ + y ` 0 = a \ + satpc(y, nat, g) \ + f = y\ + by (rule recuniq) +qed + +end + +section "Lemmas for addition" + +text \ +Let's define function t(x) = (a+x). +Firstly we need to define a function \g:nat \ nat \ nat\, such that +\g`\t`n, n\ = t`succ(n) = a + (n + 1) = (a + n) + 1 = (t`n) + 1\ +So \g`\a, b\ = a + 1\ and \g(p) = succ(pr1(p))\ +and \satpc(t,\,g) \ \n \ \ . t`succ(n) = succ(t`n)\. +\ + +definition addg :: \i\ + where addg_def : \addg == \x\(nat*nat). succ(fst(x))\ + +lemma addgfun: \function(addg)\ + by (unfold addg_def, rule func.function_lam) + +lemma addgsubpow : \addg \ Pow((nat \ nat) \ nat)\ +proof (unfold addg_def, rule subsetD) + show \(\x\nat \ nat. succ(fst(x))) \ nat \ nat \ nat\ + proof(rule func.lam_type) + fix x + assume \x\nat \ nat\ + hence \fst(x)\nat\ by auto + thus \succ(fst(x)) \ nat\ by auto + qed +next + show \nat \ nat \ nat \ Pow((nat \ nat) \ nat)\ + by (rule pisubsig) +qed + +lemma addgdom : \nat \ nat \ domain(addg)\ +proof(unfold addg_def) + have e:\domain(\x\nat \ nat. succ(fst(x))) = nat \ nat\ + by (rule domain_lam) (* "domain(Lambda(A,b)) = A" *) + show \nat \ nat \ + domain(\x\nat \ nat. succ(fst(x)))\ + by (rule subst, rule sym, rule e, auto) +qed + +lemma plussucc: + assumes F:\f \ (nat\nat)\ + assumes H:\satpc(f,nat,addg)\ + shows \\n \ nat . f`succ(n) = succ(f`n)\ +proof + fix n + assume J:\n\nat\ + from H + have H:\\n \ nat . f`succ(n) = (\x\(nat*nat). succ(fst(x)))` \ + by (unfold satpc_def, unfold addg_def) + have H:\f`succ(n) = (\x\(nat*nat). succ(fst(x)))` \ + by (rule bspec[OF H J]) + have Q:\(\x\(nat*nat). succ(fst(x)))` = succ(fst())\ + proof(rule func.beta) + show \\f ` n, n\ \ nat \ nat\ + proof + show \f ` n \ nat\ + by (rule func.apply_funtype[OF F J]) + show \n \ nat\ + by (rule J) + qed + qed + have HQ:\f`succ(n) = succ(fst())\ + by (rule trans[OF H Q]) + have K:\fst() = f`n\ + by auto + hence K:\succ(fst()) = succ(f`n)\ + by (rule subst_context) + show \f`succ(n) = succ(f`n)\ + by (rule trans[OF HQ K]) +qed + +section "Definition of addition" + +text \Theorem that addition of natural numbers exists +and unique in some sense. Due to theorem 'plussucc' the term + \satpc(f,nat,addg)\ + can be replaced here with + \\n \ nat . f`succ(n) = succ(f`n)\.\ +theorem addition: + assumes \a\nat\ + shows + \\!f. ((f \ (nat\nat)) \ ((f`0) = a) \ satpc(f,nat,addg))\ +proof(rule recthm.recursionthm, unfold recthm_def) + show \a \ nat \ addg \ nat \ nat \ nat\ + proof + show \a\nat\ by (rule assms(1)) + next + show \addg \ nat \ nat \ nat\ + proof(unfold Pi_def, rule CollectI) + show \addg \ Pow((nat \ nat) \ nat)\ + by (rule addgsubpow) + next + have A2: \nat \ nat \ domain(addg)\ + by(rule addgdom) + have A3: \function(addg)\ + by (rule addgfun) + show \nat \ nat \ domain(addg) \ function(addg)\ + by(rule conjI[OF A2 A3]) + qed + qed +qed + +end