# HG changeset patch # User nipkow # Date 1582757356 -3600 # Node ID 0e1b9b308d8f3f25c88fd5d3cbe8738f4258334d # Parent 29e297fd547304463cd901bff9562e79c3f8f776 simplified proofs diff -r 29e297fd5473 -r 0e1b9b308d8f src/HOL/Data_Structures/AVL_Set.thy --- a/src/HOL/Data_Structures/AVL_Set.thy Wed Feb 26 19:50:04 2020 +0100 +++ b/src/HOL/Data_Structures/AVL_Set.thy Wed Feb 26 23:49:16 2020 +0100 @@ -483,38 +483,28 @@ } ultimately show ?case using C by blast qed -lemma fib_alt_induct [consumes 1, case_names 1 2 rec]: - assumes "n > 0" "P 1" "P 2" "\n. n > 0 \ P n \ P (Suc n) \ P (Suc (Suc n))" - shows "P n" - using assms(1) -proof (induction n rule: fib.induct) - case (3 n) - thus ?case using assms by (cases n) (auto simp: eval_nat_numeral) -qed (insert assms, auto) - text \An exponential lower bound for \<^const>\fib\:\ lemma fib_lowerbound: defines "\ \ (1 + sqrt 5) / 2" - defines "c \ 1 / \ ^ 2" - assumes "n > 0" - shows "real (fib n) \ c * \ ^ n" -proof - - have "\ > 1" by (simp add: \_def) - hence "c > 0" by (simp add: c_def) - from \n > 0\ show ?thesis - proof (induction n rule: fib_alt_induct) - case (rec n) - have "c * \ ^ Suc (Suc n) = \ ^ 2 * (c * \ ^ n)" - by (simp add: field_simps power2_eq_square) - also have "\ \ (\ + 1) * (c * \ ^ n)" - by (rule mult_right_mono) (insert \c > 0\, simp_all add: \_def power2_eq_square field_simps) - also have "\ = c * \ ^ Suc n + c * \ ^ n" + shows "real (fib(n+2)) \ \ ^ n" +proof (induction n rule: fib.induct) + case 1 + then show ?case by simp +next + case 2 + then show ?case by (simp add: \_def real_le_lsqrt) +next + case (3 n) term ?case + have "\ ^ Suc (Suc n) = \ ^ 2 * \ ^ n" + by (simp add: field_simps power2_eq_square) + also have "\ = (\ + 1) * \ ^ n" + by (simp_all add: \_def power2_eq_square field_simps) + also have "\ = \ ^ Suc n + \ ^ n" by (simp add: field_simps) - also have "\ \ real (fib (Suc n)) + real (fib n)" - by (intro add_mono rec.IH) - finally show ?case by simp - qed (insert \\ > 1\, simp_all add: c_def power2_eq_square eval_nat_numeral) + also have "\ \ real (fib (Suc n + 2)) + real (fib (n + 2))" + by (intro add_mono "3.IH") + finally show ?case by simp qed text \The size of an AVL tree is (at least) exponential in its height:\ @@ -524,11 +514,8 @@ assumes "avl t" shows "\ ^ (height t) \ size1 t" proof - - have "\ > 0" by(simp add: \_def add_pos_nonneg) - hence "\ ^ height t = (1 / \ ^ 2) * \ ^ (height t + 2)" - by(simp add: field_simps power2_eq_square) - also have "\ \ fib (height t + 2)" - using fib_lowerbound[of "height t + 2"] by(simp add: \_def) + have "\ ^ height t \ fib (height t + 2)" + unfolding \_def by(rule fib_lowerbound) also have "\ \ size1 t" using avl_fib_bound[of t "height t"] assms by simp finally show ?thesis .