# HG changeset patch # User wenzelm # Date 1527625772 -7200 # Node ID 88c07fabd5b411422dba001138ac334c426a1fdd # Parent 56c57e91edf915fd7d4c5929521b79348ee5a840# Parent bf733673198191d80f1018272557b3ab93121075 merged diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/Data_Structures/AVL_Set.thy --- a/src/HOL/Data_Structures/AVL_Set.thy Tue May 29 22:25:59 2018 +0200 +++ b/src/HOL/Data_Structures/AVL_Set.thy Tue May 29 22:29:32 2018 +0200 @@ -455,80 +455,93 @@ subsection \Height-Size Relation\ -text \By Daniel St\"uwe\ - -fun fib_tree :: "nat \ unit avl_tree" where -"fib_tree 0 = Leaf" | -"fib_tree (Suc 0) = Node 1 Leaf () Leaf" | -"fib_tree (Suc(Suc n)) = Node (Suc(Suc(n))) (fib_tree (Suc n)) () (fib_tree n)" - -lemma [simp]: "ht (fib_tree h) = h" -by (induction h rule: "fib_tree.induct") auto +text \By Daniel St\"uwe, Manuel Eberl and Peter Lammich.\ -lemma [simp]: "height (fib_tree h) = h" -by (induction h rule: "fib_tree.induct") auto - -lemma "avl(fib_tree h)" -by (induction h rule: "fib_tree.induct") auto - -lemma fib_tree_size1: "size1 (fib_tree h) = fib (h+2)" -by (induction h rule: fib_tree.induct) auto - -lemma height_invers[simp]: +lemma height_invers: "(height t = 0) = (t = Leaf)" "avl t \ (height t = Suc h) = (\ l a r . t = Node (Suc h) l a r)" by (induction t) auto -lemma fib_Suc_lt: "fib n \ fib (Suc n)" -by (induction n rule: fib.induct) auto +text \Any AVL tree of height \h\ has at least \fib (h+2)\ leaves:\ -lemma fib_mono: "n \ m \ fib n \ fib m" -proof (induction n arbitrary: m rule: fib.induct ) - case (2 m) - thus ?case using fib_neq_0_nat[of m] by auto +lemma avl_fib_bound: "avl t \ height t = h \ fib (h+2) \ size1 t" +proof (induction h arbitrary: t rule: fib.induct) + case 1 thus ?case by (simp add: height_invers) next - case (3 n m) - from 3 obtain m' where "m = Suc (Suc m')" - by (metis le_Suc_ex plus_nat.simps(2)) - thus ?case using 3(1)[of "Suc m'"] 3(2)[of m'] 3(3) by auto -qed simp - -lemma size1_fib_tree_mono: - assumes "n \ m" - shows "size1 (fib_tree n) \ size1 (fib_tree m)" -using fib_tree_size1 fib_mono[OF assms] fib_mono[of "Suc n"] add_le_mono assms by fastforce - -lemma fib_tree_minimal: "avl t \ size1 (fib_tree (ht t)) \ size1 t" -proof (induction "ht t" arbitrary: t rule: fib_tree.induct) - case (2 t) - from 2 obtain l a r where "t = Node (Suc 0) l a r" by (cases t) auto - with 2 show ?case by auto + case 2 thus ?case by (cases t) (auto simp: height_invers) next - case (3 h t) - note [simp] = 3(3)[symmetric] - from 3 obtain l a r where [simp]: "t = Node (Suc (Suc h)) l a r" by (cases t) auto - show ?case proof (cases rule: linorder_cases[of "ht l" "ht r"]) - case equal - with 3(3,4) have ht: "ht l = Suc h" "ht r = Suc h" by auto - with 3 have "size1 (fib_tree (ht l)) \ size1 l" by auto moreover - from 3(1)[of r] 3(3,4) ht(2) have "size1 (fib_tree (ht r)) \ size1 r" by auto ultimately - show ?thesis using ht size1_fib_tree_mono[of h "Suc h"] by auto - next - case greater - with 3(3,4) have ht: "ht l = Suc h" "ht r = h" by auto - from ht 3(1,2,4) have "size1 (fib_tree (Suc h)) \ size1 l" by auto moreover - from ht 3(1,2,4) have "size1 (fib_tree h) \ size1 r" by auto ultimately - show ?thesis by auto - next - case less (* analogously *) - with 3 have ht: "ht l = h" "Suc h = ht r" by auto - from ht 3 have "size1 (fib_tree h) \ size1 l" by auto moreover - from ht 3 have "size1 (fib_tree (Suc h)) \ size1 r" by auto ultimately - show ?thesis by auto - qed -qed auto + case (3 h) + from "3.prems" obtain l a r where + [simp]: "t = Node (Suc(Suc h)) l a r" "avl l" "avl r" + and C: " + height r = Suc h \ height l = Suc h + \ height r = Suc h \ height l = h + \ height r = h \ height l = Suc h" (is "?C1 \ ?C2 \ ?C3") + by (cases t) (simp, fastforce) + { + assume ?C1 + with "3.IH"(1) + have "fib (h + 3) \ size1 l" "fib (h + 3) \ size1 r" + by (simp_all add: eval_nat_numeral) + hence ?case by (auto simp: eval_nat_numeral) + } moreover { + assume ?C2 + hence ?case using "3.IH"(1)[of r] "3.IH"(2)[of l] by auto + } moreover { + assume ?C3 + hence ?case using "3.IH"(1)[of l] "3.IH"(2)[of r] by auto + } 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}:\ -theorem avl_size_bound: "avl t \ fib(height t + 2) \ size1 t" -using fib_tree_minimal fib_tree_size1 by fastforce +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" + 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) +qed + +text \The size of an AVL tree is (at least) exponential in its height:\ + +lemma avl_lowerbound: + defines "\ \ (1 + sqrt 5) / 2" + assumes "avl t" + shows "real (size1 t) \ \ ^ (height 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 "\ \ real (fib (height t + 2))" + using fib_lowerbound[of "height t + 2"] by(simp add: \_def) + also have "\ \ real (size1 t)" + using avl_fib_bound[of t "height t"] assms by simp + finally show ?thesis . +qed end diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/Library/Fun_Lexorder.thy --- a/src/HOL/Library/Fun_Lexorder.thy Tue May 29 22:25:59 2018 +0200 +++ b/src/HOL/Library/Fun_Lexorder.thy Tue May 29 22:29:32 2018 +0200 @@ -1,6 +1,6 @@ (* Author: Florian Haftmann, TU Muenchen *) -section \Lexical order on functions\ +section \Lexicographic order on functions\ theory Fun_Lexorder imports Main diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/Library/List_Lexorder.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/List_Lexorder.thy Tue May 29 22:29:32 2018 +0200 @@ -0,0 +1,121 @@ +(* Title: HOL/Library/List_Lexorder.thy + Author: Norbert Voelker +*) + +section \Lexicographic order on lists\ + +theory List_Lexorder +imports Main +begin + +instantiation list :: (ord) ord +begin + +definition + list_less_def: "xs < ys \ (xs, ys) \ lexord {(u, v). u < v}" + +definition + list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" + +instance .. + +end + +instance list :: (order) order +proof + fix xs :: "'a list" + show "xs \ xs" by (simp add: list_le_def) +next + fix xs ys zs :: "'a list" + assume "xs \ ys" and "ys \ zs" + then show "xs \ zs" + apply (auto simp add: list_le_def list_less_def) + apply (rule lexord_trans) + apply (auto intro: transI) + done +next + fix xs ys :: "'a list" + assume "xs \ ys" and "ys \ xs" + then show "xs = ys" + apply (auto simp add: list_le_def list_less_def) + apply (rule lexord_irreflexive [THEN notE]) + defer + apply (rule lexord_trans) + apply (auto intro: transI) + done +next + fix xs ys :: "'a list" + show "xs < ys \ xs \ ys \ \ ys \ xs" + apply (auto simp add: list_less_def list_le_def) + defer + apply (rule lexord_irreflexive [THEN notE]) + apply auto + apply (rule lexord_irreflexive [THEN notE]) + defer + apply (rule lexord_trans) + apply (auto intro: transI) + done +qed + +instance list :: (linorder) linorder +proof + fix xs ys :: "'a list" + have "(xs, ys) \ lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ lexord {(u, v). u < v}" + by (rule lexord_linear) auto + then show "xs \ ys \ ys \ xs" + by (auto simp add: list_le_def list_less_def) +qed + +instantiation list :: (linorder) distrib_lattice +begin + +definition "(inf :: 'a list \ _) = min" + +definition "(sup :: 'a list \ _) = max" + +instance + by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2) + +end + +lemma not_less_Nil [simp]: "\ x < []" + by (simp add: list_less_def) + +lemma Nil_less_Cons [simp]: "[] < a # x" + by (simp add: list_less_def) + +lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" + by (simp add: list_less_def) + +lemma le_Nil [simp]: "x \ [] \ x = []" + unfolding list_le_def by (cases x) auto + +lemma Nil_le_Cons [simp]: "[] \ x" + unfolding list_le_def by (cases x) auto + +lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" + unfolding list_le_def by auto + +instantiation list :: (order) order_bot +begin + +definition "bot = []" + +instance + by standard (simp add: bot_list_def) + +end + +lemma less_list_code [code]: + "xs < ([]::'a::{equal, order} list) \ False" + "[] < (x::'a::{equal, order}) # xs \ True" + "(x::'a::{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" + by simp_all + +lemma less_eq_list_code [code]: + "x # xs \ ([]::'a::{equal, order} list) \ False" + "[] \ (xs::'a::{equal, order} list) \ True" + "(x::'a::{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" + by simp_all + +end diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/Library/List_lexord.thy --- a/src/HOL/Library/List_lexord.thy Tue May 29 22:25:59 2018 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,121 +0,0 @@ -(* Title: HOL/Library/List_lexord.thy - Author: Norbert Voelker -*) - -section \Lexicographic order on lists\ - -theory List_lexord -imports Main -begin - -instantiation list :: (ord) ord -begin - -definition - list_less_def: "xs < ys \ (xs, ys) \ lexord {(u, v). u < v}" - -definition - list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" - -instance .. - -end - -instance list :: (order) order -proof - fix xs :: "'a list" - show "xs \ xs" by (simp add: list_le_def) -next - fix xs ys zs :: "'a list" - assume "xs \ ys" and "ys \ zs" - then show "xs \ zs" - apply (auto simp add: list_le_def list_less_def) - apply (rule lexord_trans) - apply (auto intro: transI) - done -next - fix xs ys :: "'a list" - assume "xs \ ys" and "ys \ xs" - then show "xs = ys" - apply (auto simp add: list_le_def list_less_def) - apply (rule lexord_irreflexive [THEN notE]) - defer - apply (rule lexord_trans) - apply (auto intro: transI) - done -next - fix xs ys :: "'a list" - show "xs < ys \ xs \ ys \ \ ys \ xs" - apply (auto simp add: list_less_def list_le_def) - defer - apply (rule lexord_irreflexive [THEN notE]) - apply auto - apply (rule lexord_irreflexive [THEN notE]) - defer - apply (rule lexord_trans) - apply (auto intro: transI) - done -qed - -instance list :: (linorder) linorder -proof - fix xs ys :: "'a list" - have "(xs, ys) \ lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ lexord {(u, v). u < v}" - by (rule lexord_linear) auto - then show "xs \ ys \ ys \ xs" - by (auto simp add: list_le_def list_less_def) -qed - -instantiation list :: (linorder) distrib_lattice -begin - -definition "(inf :: 'a list \ _) = min" - -definition "(sup :: 'a list \ _) = max" - -instance - by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2) - -end - -lemma not_less_Nil [simp]: "\ x < []" - by (simp add: list_less_def) - -lemma Nil_less_Cons [simp]: "[] < a # x" - by (simp add: list_less_def) - -lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" - by (simp add: list_less_def) - -lemma le_Nil [simp]: "x \ [] \ x = []" - unfolding list_le_def by (cases x) auto - -lemma Nil_le_Cons [simp]: "[] \ x" - unfolding list_le_def by (cases x) auto - -lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" - unfolding list_le_def by auto - -instantiation list :: (order) order_bot -begin - -definition "bot = []" - -instance - by standard (simp add: bot_list_def) - -end - -lemma less_list_code [code]: - "xs < ([]::'a::{equal, order} list) \ False" - "[] < (x::'a::{equal, order}) # xs \ True" - "(x::'a::{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" - by simp_all - -lemma less_eq_list_code [code]: - "x # xs \ ([]::'a::{equal, order} list) \ False" - "[] \ (xs::'a::{equal, order} list) \ True" - "(x::'a::{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" - by simp_all - -end diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/ROOT --- a/src/HOL/ROOT Tue May 29 22:25:59 2018 +0200 +++ b/src/HOL/ROOT Tue May 29 22:29:32 2018 +0200 @@ -29,7 +29,7 @@ Library (*conflicting type class instantiations and dependent applications*) Finite_Lattice - List_lexord + List_Lexorder Prefix_Order Product_Lexorder Product_Order diff -r bf7336731981 -r 88c07fabd5b4 src/HOL/ex/Radix_Sort.thy --- a/src/HOL/ex/Radix_Sort.thy Tue May 29 22:25:59 2018 +0200 +++ b/src/HOL/ex/Radix_Sort.thy Tue May 29 22:29:32 2018 +0200 @@ -2,7 +2,7 @@ theory Radix_Sort imports - "HOL-Library.List_lexord" + "HOL-Library.List_Lexorder" "HOL-Library.Sublist" "HOL-Library.Multiset" begin