--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/ex13/hw13tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,43 @@
+(*<*)
+theory hw13tmpl
+imports Main
+begin
+(*>*)
+
+text \<open>\Homework{Amortized Complexity}{13.07.2018}
+
+ This is an old exam question, which we have converted to a homework to
+ be done with Isabelle!
+\<close>
+
+text \<open>A ``stack with multipop'' is a list with the following two interface functions:\<close>
+
+fun push :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+"push x xs = x # xs"
+
+fun pop :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+"pop n xs = drop n xs"
+
+text \<open>You may assume\<close>
+
+definition t_push :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
+"t_push x xs = 1"
+
+definition t_pop :: "nat \<Rightarrow> 'a list \<Rightarrow> nat" where
+"t_pop n xs = min n (length xs)"
+
+text \<open>Use the potential method to show that
+the amortized complexity of \<open>push\<close> and \<open>pop\<close> is constant.
+
+Provide complete proofs in Isabelle!
+
+Original text here was:
+\emph{
+If you need any properties of the auxiliary functions \<open>length\<close>, \<open>drop\<close> and \<open>min\<close>,
+you should state them but you do not need to prove them.}\<close>
+
+(*<*)
+end
+(*>*)
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Gottfriedsen_Jakob_j.gottfriedsen@tum.de_557/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+j.gottfriedsen@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Gottfriedsen_Jakob_j.gottfriedsen@tum.de_557/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,150 @@
+(** Score: 4/5
+*)
+(*<*)
+theory hw11
+ imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+ (*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+thm sift_up.induct
+
+(* Auxiliary lemma for a "single step" of sift_up which could be useful in the induction step of
+ a proof of sift_up_restore_heap *)
+lemma l1:
+ assumes "is_heap_except j h"
+ and "j<length h"
+ and "h!j < h!parent j"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ have 1: "\<forall>i<length h. i \<noteq> j \<and> i \<noteq> parent j \<longrightarrow> swap j (parent j) h ! parent i \<le> swap j (parent j) h ! i"
+ using assms is_heap_except_def[of j h] swap_def[of j "parent j" h] parent_less
+ by (smt length_list_update less_trans not_le nth_list_update_eq nth_list_update_neq)
+ have 2: "\<forall>i<length h. parent i = parent j \<longrightarrow> swap j (parent j) h ! parent (parent j) \<le> swap j (parent j) h ! i"
+ using assms is_heap_except_def[of j h] swap_def parent_def parent_less
+ by (smt One_nat_def Suc_eq_plus1 diff_is_0_eq' leD leI le_numeral_extra(1) length_list_update less_imp_not_less nth_list_update one_div_two_eq_zero order_trans)
+ have 3: "swap j (parent j) h ! parent j \<le> swap j (parent j) h ! j"
+ using assms is_heap_except_def swap_def parent_def parent_less
+ by (smt One_nat_def Suc_eq_plus1 diff_is_0_eq' le_numeral_extra(1) length_list_update less_imp_le less_trans nth_list_update one_div_two_eq_zero)
+ show ?thesis using assms 1 2 3 swap_def is_heap_except_def by (metis length_list_update)
+qed
+
+(* This is the wrong induction scheme. I think an induction using sift_up.induct would be needed to
+ get the right IH, but "induction rule: sift_up.induct" resulted in an unprovable goal and adding
+ "arbitrary: h" didn't work either. *)
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using is_heap_def[of h] is_heap_except_def[of 0 h] parent_def[of 0] by auto
+next
+ case (2 h v)
+ then show ?case
+
+ (** Your intuition was right! See above how to set up the computation induction correctly. *)
+ sorry
+qed
+
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction j arbitrary: h)
+ case 0
+ hence "is_heap h" using is_heap_def[of h] is_heap_except_def[of 0 h] parent_def[of 0] by force
+ then show ?case by simp
+next
+ case (Suc j)
+ show ?case
+ proof (cases "h ! parent (Suc j) \<le> h ! Suc j")
+ case True
+ have "sift_up h (Suc j) = h" using True sift_up.simps(2)[of h j] by simp
+ moreover have "is_heap h"
+ using Suc.prems(1) True is_heap_def[of h] is_heap_except_def[of "Suc j" h] by auto
+ ultimately show ?thesis by simp
+ next
+ case False
+ from False Suc.IH Suc.prems show ?thesis using sift_up.simps(2)[of h j] oops
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ then show ?case
+ by (metis swap_def less_trans mset_swap nat.simps(3) parent_less sift_up.simps(2) size_mset)
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+
+lemma invar_empty: "is_heap emp"
+ by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (simp add: emp_def)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+ (*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Gottfriedsen_Jakob_j.gottfriedsen@tum.de_557/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,277 @@
+Using temporary directory '/tmp/tmp.lYmBg2eXqN'
+Files in /tmp/eval-557-UAOHs0: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<defs-imports>
+Complex_Main
+</defs-imports>
+<image>
+HOL-Library
+</image>
+<check>
+
+</check>
+<version>
+2016-1
+</version>
+<submission-imports>
+
+</submission-imports>
+<submission>
+
+ (*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+thm sift_up.induct
+
+(* Auxiliary lemma for a "single step" of sift_up which could be useful in the induction step of
+ a proof of sift_up_restore_heap *)
+lemma l1:
+ assumes "is_heap_except j h"
+ and "j<length h"
+ and "h!j < h!parent j"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ have 1: "\<forall>i<length h. i \<noteq> j \<and> i \<noteq> parent j \<longrightarrow> swap j (parent j) h ! parent i \<le> swap j (parent j) h ! i"
+ using assms is_heap_except_def[of j h] swap_def[of j "parent j" h] parent_less
+ by (smt length_list_update less_trans not_le nth_list_update_eq nth_list_update_neq)
+ have 2: "\<forall>i<length h. parent i = parent j \<longrightarrow> swap j (parent j) h ! parent (parent j) \<le> swap j (parent j) h ! i"
+ using assms is_heap_except_def[of j h] swap_def parent_def parent_less
+ by (smt One_nat_def Suc_eq_plus1 diff_is_0_eq' leD leI le_numeral_extra(1) length_list_update less_imp_not_less nth_list_update one_div_two_eq_zero order_trans)
+ have 3: "swap j (parent j) h ! parent j \<le> swap j (parent j) h ! j"
+ using assms is_heap_except_def swap_def parent_def parent_less
+ by (smt One_nat_def Suc_eq_plus1 diff_is_0_eq' le_numeral_extra(1) length_list_update less_imp_le less_trans nth_list_update one_div_two_eq_zero)
+ show ?thesis using assms 1 2 3 swap_def is_heap_except_def by (metis length_list_update)
+qed
+
+(* This is the wrong induction scheme. I think an induction using sift_up.induct would be needed to
+ get the right IH, but "induction rule: sift_up.induct" resulted in an unprovable goal and adding
+ "arbitrary: h" didn't work either. *)
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction j arbitrary: h)
+ case 0
+ hence "is_heap h" using is_heap_def[of h] is_heap_except_def[of 0 h] parent_def[of 0] by force
+ then show ?case by simp
+next
+ case (Suc j)
+ show ?case
+ proof (cases "h ! parent (Suc j) \<le> h ! Suc j")
+ case True
+ have "sift_up h (Suc j) = h" using True sift_up.simps(2)[of h j] by simp
+ moreover have "is_heap h"
+ using Suc.prems(1) True is_heap_def[of h] is_heap_except_def[of "Suc j" h] by auto
+ ultimately show ?thesis by simp
+ next
+ case False
+ from False Suc.IH Suc.prems show ?thesis using sift_up.simps(2)[of h j] oops
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ then show ?case
+ by (metis swap_def less_trans mset_swap nat.simps(3) parent_less sift_up.simps(2) size_mset)
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+
+lemma invar_empty: "is_heap emp"
+ by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (simp add: emp_def)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" good_sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+<check-imports>
+Submission
+</check-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Griebel_Simon_s.griebel@tum.de_562/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+s.griebel@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Griebel_Simon_s.griebel@tum.de_562/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,167 @@
+(** Score: 8/5
+
+ nice job!
+*)
+(*<*)
+theory hw11
+ imports Main
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+
+lemma is_heap_except_is_heap: "h \<noteq> [] \<Longrightarrow> is_heap_except 0 h = is_heap h"
+ unfolding is_heap_def is_heap_except_def parent_def by force
+
+(*This slightly reduces the lemmas needed for the smt solver in the next proof*)
+lemma parent_less2: "a < b \<Longrightarrow> parent a < b"
+ using parent_def by auto
+
+lemma is_heap_except_swap_parent:
+ assumes "i < length h" "is_heap_except i h" "h!i < h!parent i"
+ shows "is_heap_except (parent i) (swap i (parent i) h)"
+proof -
+ have aux1: "\<forall>ia<length (swap i (parent i) h).ia \<noteq> parent i
+ \<longrightarrow> swap i (parent i) h ! parent ia \<le> swap i (parent i) h ! ia"
+ using assms
+ unfolding swap_def is_heap_except_def
+ by (smt leD le_cases le_less_trans length_list_update nth_list_update_eq nth_list_update_neq parent_less2)
+ have aux2: "\<forall>ia<length (swap i (parent i) h). parent ia = parent i
+ \<longrightarrow> swap i (parent i) h ! parent (parent i) \<le> swap i (parent i) h ! ia"
+ using assms
+ unfolding swap_def is_heap_except_def
+ by (smt dual_order.asym dual_order.trans leI length_list_update nth_list_update_eq nth_list_update_neq parent_less2)
+ show ?thesis using aux1 aux2 unfolding is_heap_except_def by blast
+qed
+
+lemma swap_presv_length: "length (swap i j h) = length h"
+ unfolding swap_def by simp
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ using is_heap_except_is_heap by auto
+next
+ case (2 h v)
+ then show ?case
+ unfolding is_heap_def is_heap_except_def
+ using swap_presv_length parent_less2 is_heap_except_swap_parent 2
+ by (metis is_heap_def leI sift_up.simps(2))
+ qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ by (auto simp add: swap_def parent_def mset_swap)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h!0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h =
+ (let h = h@[x] in sift_up h (length h - 1))"
+
+lemma invar_empty: "is_heap emp"
+ by(simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by(simp add: emp_def is_heap_def)
+
+(*helper function, so that a recursive/ inductive proof over parent is possible*)
+fun rec_parent :: "nat \<Rightarrow> nat" where
+ "rec_parent 0 = 0"
+| "rec_parent (Suc n) = rec_parent (parent (Suc n))"
+
+lemma rec_parent_el_less: "n < length (h::'a::linorder list) \<Longrightarrow> \<forall>n < length h. h!parent n \<le> h!n \<Longrightarrow> h!rec_parent n \<le> h!n"
+proof(induction n rule: rec_parent.induct)
+ case 1
+ then show ?case by simp
+next
+ case (2 n)
+ hence "h!parent (Suc n) \<le> h!(Suc n)" by fast
+ hence "h!rec_parent (parent (Suc n)) \<le> h!(Suc n)"
+ using 2 dual_order.trans parent_less2 by auto
+ then show ?case by auto
+qed
+
+lemma rec_parent_is_zero: "rec_parent n = 0"
+ apply(induction n rule: rec_parent.induct)
+ by(auto)
+
+lemma first_el_smallest: "is_heap h \<Longrightarrow> length h \<noteq> 0 \<Longrightarrow> n < length h \<Longrightarrow> h!0 \<le> h!n"
+ unfolding is_heap_def
+ using rec_parent_el_less rec_parent_is_zero by auto
+
+lemma aux: "length h \<noteq> 0 \<Longrightarrow> (\<forall>n < length h. x \<le> h!n) \<Longrightarrow> x \<le> Min_mset (mset h)"
+ by (auto simp add: in_set_conv_nth)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h \<le> Min_mset (mset h)"
+ using first_el_smallest get_min_def aux
+ by (metis length_0_conv mset_zero_iff)
+
+lemma append_el: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[a])"
+ unfolding is_heap_def is_heap_except_def
+ using parent_less2 less_antisym nth_append less_Suc_eq_le
+ by (metis length_append_singleton not_less_eq order_refl)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ by (simp add: append_el sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (auto simp add: ins_def sift_up_mset)
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Griebel_Simon_s.griebel@tum.de_562/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,306 @@
+Using temporary directory '/tmp/tmp.zEO93VFYCv'
+Files in /tmp/eval-562-ZSqBfY: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<defs-imports>
+Complex_Main
+</defs-imports>
+<check-imports>
+Submission
+</check-imports>
+<submission>
+
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+
+lemma is_heap_except_is_heap: "h \<noteq> [] \<Longrightarrow> is_heap_except 0 h = is_heap h"
+ unfolding is_heap_def is_heap_except_def parent_def by force
+
+(*This slightly reduces the lemmas needed for the smt solver in the next proof*)
+lemma parent_less2: "a < b \<Longrightarrow> parent a < b"
+ using parent_def by auto
+
+lemma is_heap_except_swap_parent:
+ assumes "i < length h" "is_heap_except i h" "h!i < h!parent i"
+ shows "is_heap_except (parent i) (swap i (parent i) h)"
+proof -
+ have aux1: "\<forall>ia<length (swap i (parent i) h).ia \<noteq> parent i
+ \<longrightarrow> swap i (parent i) h ! parent ia \<le> swap i (parent i) h ! ia"
+ using assms
+ unfolding swap_def is_heap_except_def
+ by (smt leD le_cases le_less_trans length_list_update nth_list_update_eq nth_list_update_neq parent_less2)
+ have aux2: "\<forall>ia<length (swap i (parent i) h). parent ia = parent i
+ \<longrightarrow> swap i (parent i) h ! parent (parent i) \<le> swap i (parent i) h ! ia"
+ using assms
+ unfolding swap_def is_heap_except_def
+ by (smt dual_order.asym dual_order.trans leI length_list_update nth_list_update_eq nth_list_update_neq parent_less2)
+ show ?thesis using aux1 aux2 unfolding is_heap_except_def by blast
+qed
+
+lemma swap_presv_length: "length (swap i j h) = length h"
+ unfolding swap_def by simp
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ using is_heap_except_is_heap by auto
+next
+ case (2 h v)
+ then show ?case
+ unfolding is_heap_def is_heap_except_def
+ using swap_presv_length parent_less2 is_heap_except_swap_parent 2
+ by (metis is_heap_def leI sift_up.simps(2))
+ qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ by (auto simp add: swap_def parent_def mset_swap)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h!0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h =
+ (let h = h@[x] in sift_up h (length h - 1))"
+
+lemma invar_empty: "is_heap emp"
+ by(simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by(simp add: emp_def is_heap_def)
+
+(*helper function, so that a recursive/ inductive proof over parent is possible*)
+fun rec_parent :: "nat \<Rightarrow> nat" where
+ "rec_parent 0 = 0"
+| "rec_parent (Suc n) = rec_parent (parent (Suc n))"
+
+lemma rec_parent_el_less: "n < length (h::'a::linorder list) \<Longrightarrow> \<forall>n < length h. h!parent n \<le> h!n \<Longrightarrow> h!rec_parent n \<le> h!n"
+proof(induction n rule: rec_parent.induct)
+ case 1
+ then show ?case by simp
+next
+ case (2 n)
+ hence "h!parent (Suc n) \<le> h!(Suc n)" by fast
+ hence "h!rec_parent (parent (Suc n)) \<le> h!(Suc n)"
+ using 2 dual_order.trans parent_less2 by auto
+ then show ?case by auto
+qed
+
+lemma rec_parent_is_zero: "rec_parent n = 0"
+ apply(induction n rule: rec_parent.induct)
+ by(auto)
+
+lemma first_el_smallest: "is_heap h \<Longrightarrow> length h \<noteq> 0 \<Longrightarrow> n < length h \<Longrightarrow> h!0 \<le> h!n"
+ unfolding is_heap_def
+ using rec_parent_el_less rec_parent_is_zero by auto
+
+lemma aux: "length h \<noteq> 0 \<Longrightarrow> (\<forall>n < length h. x \<le> h!n) \<Longrightarrow> x \<le> Min_mset (mset h)"
+ by (auto simp add: in_set_conv_nth)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h \<le> Min_mset (mset h)"
+ using first_el_smallest get_min_def aux
+ by (metis length_0_conv mset_zero_iff)
+
+lemma append_el: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[a])"
+ unfolding is_heap_def is_heap_except_def
+ using parent_less2 less_antisym nth_append less_Suc_eq_le
+ by (metis length_append_singleton not_less_eq order_refl)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ by (simp add: append_el sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (auto simp add: ins_def sift_up_mset)
+
+(*<*)
+</submission>
+<submission-imports>
+
+</submission-imports>
+<version>
+2016-1
+</version>
+<image>
+HOL-Library
+</image>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<check>
+
+</check>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Jonischkeit_Clemens_clemens.jonischkeit@tum.de_563/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+clemens.jonischkeit@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Jonischkeit_Clemens_clemens.jonischkeit@tum.de_563/jonischkeit11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,274 @@
+(** Score: 0/5
+
+ submission equals template, nothing done
+*)
+(*<*)
+theory jonischkeit11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert a Leaf = (Node 1 Leaf a Leaf)" |
+ "lh_insert a (Node n l v r) = (if a < v
+ then node l a (lh_insert v r)
+ else node l v (lh_insert a r))"
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x, auto simp add:node_def)
+
+value "1 :: int"
+
+lemma node_heap:
+ "\<lbrakk>heap l; heap r; \<forall>x \<in> set_mset (mset_tree l + mset_tree r) . v \<le> x\<rbrakk> \<Longrightarrow> heap (node l v r)"
+ unfolding node_def by (auto)
+
+lemma node_ltree:
+ "\<lbrakk>ltree l; ltree r\<rbrakk> \<Longrightarrow> ltree (node l v r)"
+ unfolding node_def by auto
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+proof (induction t arbitrary: x)
+ case (Node n l v r)
+ have xs: "\<forall>x \<in> set_mset (mset_tree l + mset_tree r). v \<le> x" using Node.prems by auto
+ show ?case
+ proof (cases "x < v")
+ case True
+ have xsr: "\<forall>x \<in> set_mset (mset_tree r). v \<le> x" using xs by simp
+ have hr: "heap (lh_insert v r)" using Node by auto
+ have "\<forall>y \<in> set_mset (mset_tree (lh_insert v r)). v \<le> y" using xsr mset_lh_insert[of v r] by simp
+ hence gr: "\<forall>y \<in> set_mset (mset_tree (lh_insert v r)). x \<le> y" using True by auto
+ have "\<forall>y \<in> set_mset (mset_tree l). v \<le> y" using xs by simp
+ hence gl: "\<forall>y \<in> set_mset (mset_tree l). x \<le> y" using True by auto
+ hence "\<forall>y \<in> set_mset (mset_tree l + mset_tree (lh_insert v r)). x \<le> y"
+ using gr by auto
+ thus ?thesis using hr Node.prems node_heap True by auto
+ next
+ case False
+ have xsr: "\<forall>x \<in> set_mset (mset_tree r). v \<le> x" using xs by simp
+ have hr: "heap (lh_insert x r)" using Node by auto
+ have gr: "\<forall>y \<in> set_mset (mset_tree (lh_insert x r)). v \<le> y" using xsr mset_lh_insert[of x r] False by simp
+ have "\<forall>y \<in> set_mset (mset_tree l). v \<le> y" using xs by simp
+ hence "\<forall>y \<in> set_mset (mset_tree l + mset_tree (lh_insert x r)). v \<le> y"
+ using gr by auto
+ then show ?thesis using hr Node.prems node_heap False by auto
+ qed
+qed (simp)
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ by (induction t arbitrary: x, auto simp add: node_ltree)
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+ where
+ "t_lh_insert _ Leaf = 1" |
+ "t_lh_insert a (Node n l v r) = (if a < v
+ then 1 + t_lh_insert v r
+ else 1 + t_lh_insert a r)"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x, auto)
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | BSPQ 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty a \<longleftrightarrow> a = Empty"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert a Empty = BSPQ a orig_empty" |
+ "insert a (BSPQ b q) = (if a < b
+ then BSPQ a (orig_insert b q)
+ else BSPQ b (orig_insert a q))"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min (BSPQ a _) = a" |
+ "get_min Empty = undefined"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min (BSPQ a q) = (if orig_is_empty q
+ then Empty
+ else BSPQ (orig_get_min q) (orig_del_min q))" |
+ "del_min Empty = Empty"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty = True" |
+ "invar (BSPQ a q) = (orig_invar q \<and> (\<forall>x \<in> set_mset (orig_mset q). a \<le> x))"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (BSPQ a q) = orig_mset q + {# a #}"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ thus ?case using local.empty_def by simp
+ next case (2 q)
+ thus ?case using mset.elims by force
+ next case (3 q)
+ thus ?case by (cases q, auto)
+ next case (4 q)
+ thus ?case by (cases q, auto)
+ next case (5 q)
+ hence ?case proof (cases q)
+ case (BSPQ a q')
+ hence "\<forall>x \<in> set_mset (orig_mset q'). a \<le> x" using "5" by simp
+ thus ?thesis using BSPQ antisym by (simp add: antisym)
+ qed (simp)
+ next case 6
+ thus ?case using local.empty_def by simp
+ next case (7 q)
+ thus ?case by (cases q, auto)
+ next case (8 q)
+ thus ?case proof (cases q)
+ case (BSPQ a q')
+ thus ?thesis using 8 finite_set_mset in_diffD by fastforce
+ qed (simp)
+ let ?case2 = "get_min q = Min_mset (mset q)"
+ show ?case2 using 8 proof (cases q)
+ case (BSPQ a q')
+ hence "\<forall>x \<in> set_mset (orig_mset q'). a \<le> x" using "8" by simp
+ thus ?thesis using 8 BSPQ by (simp add: antisym)
+ qed (simp)
+ qed
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ sorry
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ sorry
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" sorry
+
+lemma mset_empty: "mset emp = {#}" sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Jonischkeit_Clemens_clemens.jonischkeit@tum.de_563/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,413 @@
+Using temporary directory '/tmp/tmp.aa95zzpT9j'
+Files in /tmp/eval-563-ZCyGxC: jonischkeit11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<check>
+
+</check>
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert a Leaf = (Node 1 Leaf a Leaf)" |
+ "lh_insert a (Node n l v r) = (if a < v
+ then node l a (lh_insert v r)
+ else node l v (lh_insert a r))"
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x, auto simp add:node_def)
+
+value "1 :: int"
+
+lemma node_heap:
+ "\<lbrakk>heap l; heap r; \<forall>x \<in> set_mset (mset_tree l + mset_tree r) . v \<le> x\<rbrakk> \<Longrightarrow> heap (node l v r)"
+ unfolding node_def by (auto)
+
+lemma node_ltree:
+ "\<lbrakk>ltree l; ltree r\<rbrakk> \<Longrightarrow> ltree (node l v r)"
+ unfolding node_def by auto
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+proof (induction t arbitrary: x)
+ case (Node n l v r)
+ have xs: "\<forall>x \<in> set_mset (mset_tree l + mset_tree r). v \<le> x" using Node.prems by auto
+ show ?case
+ proof (cases "x < v")
+ case True
+ have xsr: "\<forall>x \<in> set_mset (mset_tree r). v \<le> x" using xs by simp
+ have hr: "heap (lh_insert v r)" using Node by auto
+ have "\<forall>y \<in> set_mset (mset_tree (lh_insert v r)). v \<le> y" using xsr mset_lh_insert[of v r] by simp
+ hence gr: "\<forall>y \<in> set_mset (mset_tree (lh_insert v r)). x \<le> y" using True by auto
+ have "\<forall>y \<in> set_mset (mset_tree l). v \<le> y" using xs by simp
+ hence gl: "\<forall>y \<in> set_mset (mset_tree l). x \<le> y" using True by auto
+ hence "\<forall>y \<in> set_mset (mset_tree l + mset_tree (lh_insert v r)). x \<le> y"
+ using gr by auto
+ thus ?thesis using hr Node.prems node_heap True by auto
+ next
+ case False
+ have xsr: "\<forall>x \<in> set_mset (mset_tree r). v \<le> x" using xs by simp
+ have hr: "heap (lh_insert x r)" using Node by auto
+ have gr: "\<forall>y \<in> set_mset (mset_tree (lh_insert x r)). v \<le> y" using xsr mset_lh_insert[of x r] False by simp
+ have "\<forall>y \<in> set_mset (mset_tree l). v \<le> y" using xs by simp
+ hence "\<forall>y \<in> set_mset (mset_tree l + mset_tree (lh_insert x r)). v \<le> y"
+ using gr by auto
+ then show ?thesis using hr Node.prems node_heap False by auto
+ qed
+qed (simp)
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ by (induction t arbitrary: x, auto simp add: node_ltree)
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+ where
+ "t_lh_insert _ Leaf = 1" |
+ "t_lh_insert a (Node n l v r) = (if a < v
+ then 1 + t_lh_insert v r
+ else 1 + t_lh_insert a r)"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x, auto)
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | BSPQ 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty a \<longleftrightarrow> a = Empty"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert a Empty = BSPQ a orig_empty" |
+ "insert a (BSPQ b q) = (if a < b
+ then BSPQ a (orig_insert b q)
+ else BSPQ b (orig_insert a q))"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min (BSPQ a _) = a" |
+ "get_min Empty = undefined"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min (BSPQ a q) = (if orig_is_empty q
+ then Empty
+ else BSPQ (orig_get_min q) (orig_del_min q))" |
+ "del_min Empty = Empty"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty = True" |
+ "invar (BSPQ a q) = (orig_invar q \<and> (\<forall>x \<in> set_mset (orig_mset q). a \<le> x))"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (BSPQ a q) = orig_mset q + {# a #}"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ thus ?case using local.empty_def by simp
+ next case (2 q)
+ thus ?case using mset.elims by force
+ next case (3 q)
+ thus ?case by (cases q, auto)
+ next case (4 q)
+ thus ?case by (cases q, auto)
+ next case (5 q)
+ hence ?case proof (cases q)
+ case (BSPQ a q')
+ hence "\<forall>x \<in> set_mset (orig_mset q'). a \<le> x" using "5" by simp
+ thus ?thesis using BSPQ antisym by (simp add: antisym)
+ qed (simp)
+ next case 6
+ thus ?case using local.empty_def by simp
+ next case (7 q)
+ thus ?case by (cases q, auto)
+ next case (8 q)
+ thus ?case proof (cases q)
+ case (BSPQ a q')
+ thus ?thesis using 8 finite_set_mset in_diffD by fastforce
+ qed (simp)
+ let ?case2 = "get_min q = Min_mset (mset q)"
+ show ?case2 using 8 proof (cases q)
+ case (BSPQ a q')
+ hence "\<forall>x \<in> set_mset (orig_mset q'). a \<le> x" using "8" by simp
+ thus ?thesis using 8 BSPQ by (simp add: antisym)
+ qed (simp)
+ qed
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ good_sorry
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ good_sorry
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" good_sorry
+
+lemma mset_empty: "mset emp = {#}" good_sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" good_sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+<check-imports>
+Submission
+</check-imports>
+<image>
+HOL-Library
+</image>
+<submission-imports>
+
+</submission-imports>
+<version>
+2016-1
+</version>
+<defs-imports>
+Complex_Main
+</defs-imports>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kirchmeier_Maximilian_max.kirchmeier@tum.de_558/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+max.kirchmeier@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kirchmeier_Maximilian_max.kirchmeier@tum.de_558/ex11_tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,344 @@
+(** Score: 8/5
+
+ nice job
+*)
+(*<*)
+theory ex11_tmpl
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+lemma not_first:
+ assumes "is_heap_except 0 h"
+ shows "is_heap h"
+proof (cases h)
+ case Nil
+ then show ?thesis
+ by (simp add: is_heap_def)
+next
+ case Cons
+
+ then have "h!0 \<ge> h!(parent 0)" unfolding parent_def by auto
+ moreover have "\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)"
+ using assms by (simp add: is_heap_except_def parent_def)
+ ultimately show ?thesis
+ using is_heap_def by auto
+qed
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+fun parents :: "nat \<Rightarrow> nat set" where
+ "parents 0 = {0}"
+| "parents n = {parent n} \<union> parents (parent n)"
+
+thm "parents.induct" (* now that's interesting *)
+
+lemma parent_induct: "P (0::nat) \<Longrightarrow> (\<And>i. P (parent i) \<Longrightarrow> P i) \<Longrightarrow> P j"
+proof -
+ obtain j' where "j' = j * 2 + 1"
+ by simp
+ then have "parent j' = j" unfolding parent_def
+ by linarith
+ then show "P (0::nat) \<Longrightarrow> (\<And>i. P (parent i) \<Longrightarrow> P i) \<Longrightarrow> P j"
+ using infinite_descent0 parent_less by blast
+qed
+
+lemma is_heap_min: "h \<noteq> [] \<Longrightarrow> is_heap h \<Longrightarrow> i < length h \<Longrightarrow> h!i \<ge> h!0" unfolding is_heap_def parent_def
+proof (induction i arbitrary: h rule: parent_induct)
+ case 1
+ then show ?case by auto
+next
+ case (2 i)
+ then show ?case
+ using eq_iff leD leI parent_def by fastforce
+qed
+
+lemma swap_len: "length (swap i j h) = length h" unfolding swap_def by auto
+
+lemma swap_pres: "i < length h \<and> j < length h \<Longrightarrow> swap i j (h @ h') = swap i j h @ h'"
+ unfolding swap_def
+ by (simp add: list_update_append1 nth_append)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+lemma sift_len: "i < length h \<Longrightarrow> length (sift_up h i) = length h"
+proof (induction i arbitrary: h rule: parent_induct)
+ case 1
+ then show ?case by auto
+next
+ case (2 i')
+ then show ?case
+ proof (cases i')
+ case 0
+ then show ?thesis by auto
+ next
+ case (Suc nat)
+ then show ?thesis
+ by (metis (no_types, lifting) "2.IH" "2.prems" Suc_lessD less_numeral_extra(3) less_trans_Suc parent_less sift_up.simps(2) swap_len zero_less_Suc)
+ qed
+qed
+
+lemma sift_pres: "i < length h \<Longrightarrow> sift_up (h @ h') i = sift_up h i @ h'"
+proof (induction i arbitrary: h rule: parent_induct)
+case 1
+ then show ?case by auto
+next
+ case (2 i')
+ show ?case
+ proof (cases i')
+ case 0
+ then show ?thesis by auto
+ next
+ case (Suc nat)
+ then show ?thesis
+ proof (cases "h!i' \<ge> h!parent i'")
+ case True
+ then show ?thesis
+ by (smt "2.prems" Suc dual_order.strict_implies_order less_le_trans nth_append old.nat.distinct(2) parent_less sift_up.simps(2))
+ next
+ case False
+ then show ?thesis using Suc swap_pres
+ by (smt "2.IH" "2.prems" less_imp_le_nat less_le_trans nat.simps(3) nth_append parent_less sift_up.simps(2) swap_len)
+ qed
+ qed
+qed
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+thm sift_up.induct
+
+value "is_heap [finite_2.a\<^sub>1, finite_2.a\<^sub>2]"
+
+lemma swap_maybe_restore:
+ assumes "\<not> is_heap h" "is_heap_except j h" "j<length h"
+ shows "is_heap (swap j (parent j) h) \<or> (\<not> is_heap (swap j (parent j) h) \<and> is_heap_except (parent j) (swap j (parent j) h))"
+proof (cases h)
+ case Nil
+ then show ?thesis
+ by (simp add: ex11_tmpl.swap_def is_heap_def)
+next
+ case (Cons a h')
+
+ have almost_heap:
+ "\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!parent i"
+ "\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j)"
+ using assms is_heap_except_def by auto
+ have j_less: "h!j < h!parent j" using assms is_heap_def is_heap_except_def by (metis leI)
+
+ let ?h_s = "swap j (parent j) h"
+
+ have almost_heap2: "\<forall>i<length ?h_s. i\<noteq>parent j \<longrightarrow> ?h_s!i \<ge> ?h_s!parent i" using almost_heap j_less
+ by (smt assms(2) assms(3) dual_order.strict_trans dual_order.trans ex11_tmpl.swap_def is_heap_def leD length_list_update less_imp_le not_first nth_list_update_eq nth_list_update_neq parent_less)
+
+ show ?thesis
+ proof (cases "is_heap ?h_s")
+ case True
+ then show ?thesis
+ by simp
+ next
+ case False
+
+ then have "is_heap_except (parent j) (swap j (parent j) h)" using almost_heap2 almost_heap
+ by (smt assms(3) dual_order.trans eq_iff ex11_tmpl.swap_def is_heap_def is_heap_except_def nth_list_update_eq nth_list_update_neq swap_len)
+
+ then show ?thesis
+ by simp
+ qed
+qed
+
+lemma sift_nop: "j < length h \<Longrightarrow> is_heap h \<Longrightarrow> sift_up h j = h"
+ by (metis is_heap_def sift_up.elims)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using not_first by auto
+next
+ case (2 h v)
+ let ?v = "Suc v"
+ let ?h_s = "swap ?v (parent ?v) h"
+
+ show ?case
+ proof (cases "h!?v \<ge> h!parent ?v")
+ case True
+ then show ?thesis
+ by (metis "2.prems"(1) is_heap_def is_heap_except_def sift_up.elims)
+ next
+ case False
+ then have noheap: "\<not> is_heap h"
+ using "2.prems"(2) is_heap_def by blast
+ from False have "sift_up h ?v = sift_up ?h_s (parent ?v)"
+ by simp
+
+ consider "is_heap ?h_s" | "(\<not> is_heap ?h_s \<and> is_heap_except (parent ?v) ?h_s)"
+ using swap_maybe_restore[of h ?v] using noheap 2 by auto
+
+ then show ?thesis
+ proof cases
+ case 1
+ then have "sift_up h (Suc v) = ?h_s" using sift_nop
+ by (metis "2.prems"(2) \<open>sift_up h (Suc v) = sift_up (swap (Suc v) (parent (Suc v)) h) (parent (Suc v))\<close> dual_order.strict_trans nat.simps(3) parent_less swap_len)
+ then show ?thesis
+ using "1" by auto
+ next
+ case 2
+ then show ?thesis
+ by (metis "2.IH" "2.prems"(2) False \<open>sift_up h (Suc v) = sift_up (swap (Suc v) (parent (Suc v)) h) (parent (Suc v))\<close> dual_order.strict_trans length_code nat.simps(3) parent_less swap_len)
+ qed
+ qed
+qed
+
+lemma swap_mset:
+ assumes "i < length h" "j < length h"
+ shows "mset (swap i j h) = mset h"
+ by (simp add: assms(1) assms(2) ex11_tmpl.swap_def mset_swap)
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case using swap_mset
+ by (metis dual_order.strict_trans nat.simps(3) parent_less sift_up.simps(2) swap_len)
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> hd h"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding emp_def is_heap_def by auto
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by auto
+
+lemma mset_get_min:
+ assumes "is_heap h" "mset h \<noteq> {#}"
+ shows "get_min h = Min_mset (mset h)"
+proof (cases h)
+ case Nil
+ then show ?thesis using assms
+ by simp
+next
+ case (Cons a h')
+
+ then have "\<forall>i < length h. h ! 0 \<le> h ! i" using is_heap_min
+ using assms(1) by blast
+
+ then have "\<forall>e \<in> set h. a \<le> e"
+ by (metis in_set_conv_nth local.Cons nth_Cons_0)
+ moreover
+ have "finite (set h)"
+ by simp
+ ultimately have "a = Min (set h)"
+ by (simp add: Min_insert2 local.Cons)
+
+ moreover
+
+ have "get_min h = a" using Cons
+ by (simp add: get_min_def)
+ moreover
+ have "set_mset (mset h) = set h"
+ by simp
+
+ ultimately show ?thesis
+ by auto
+qed
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+proof (cases h)
+ case Nil
+ then show ?thesis unfolding ins_def
+ by (auto simp add: is_heap_def parent_def)
+next
+ case (Cons a list)
+ assume heap: "is_heap h"
+
+ let ?h' = "h @ [x]"
+
+ let ?j = "length h"
+
+ have "i < length ?h' \<Longrightarrow> (i\<noteq>?j \<longrightarrow> ?h'!i \<ge> ?h'!(parent i))
+ \<and> (parent i = ?j \<longrightarrow> ?h'!i \<ge> ?h'!(parent ?j))" for i
+ proof (cases "i = ?j")
+ case True
+ then show ?thesis
+ by simp
+ next
+ case False
+
+ assume "i < length (h @ [x])"
+ then have "i < length h"
+ by (simp add: False less_Suc_eq)
+ then have "parent i < ?j" unfolding parent_def using parent_less
+ by auto
+ then show ?thesis using heap
+ by (metis \<open>i < length h\<close> is_heap_def neq_iff nth_append)
+ qed
+
+ then have "is_heap_except ?j ?h'" unfolding is_heap_except_def
+ by blast
+
+ then show ?thesis
+ by (simp add: ins_def sift_up_restore_heap)
+qed
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+proof -
+ have "mset (ins x h) = mset (h @ [x])" unfolding ins_def
+ using sift_up_mset[of "length h" "h @ [x]"] by auto
+ also have "\<dots> = mset h + {#x#}"
+ by simp
+ finally show "mset (ins x h) = mset h + {#x#}"
+ by auto
+qed
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kirchmeier_Maximilian_max.kirchmeier@tum.de_558/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,596 @@
+Using temporary directory '/tmp/tmp.qsDMk4pX5V'
+Files in /tmp/eval-558-IhnEbA: ex11_tmpl.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<check>
+
+</check>
+<check-imports>
+Submission
+</check-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<version>
+2016-1
+</version>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<image>
+HOL-Library
+</image>
+<submission-imports>
+
+</submission-imports>
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert _ _ = undefined"
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ good_sorry
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ good_sorry
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ good_sorry
+
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+ where
+ "t_lh_insert _ _ = undefined"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ good_sorry
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = PUT_CONSTRUCTORS_HERE
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = undefined"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty _ \<longleftrightarrow> undefined"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert _ _ = undefined"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min _ = undefined"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min _ = undefined"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar _ = undefined"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset _ = undefined"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ then show ?case good_sorry
+ next
+ case (2 q) -- \<open>and so on\<close>
+ oops
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+lemma not_first:
+ assumes "is_heap_except 0 h"
+ shows "is_heap h"
+proof (cases h)
+ case Nil
+ then show ?thesis
+ by (simp add: is_heap_def)
+next
+ case Cons
+
+ then have "h!0 \<ge> h!(parent 0)" unfolding parent_def by auto
+ moreover have "\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)"
+ using assms by (simp add: is_heap_except_def parent_def)
+ ultimately show ?thesis
+ using is_heap_def by auto
+qed
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+fun parents :: "nat \<Rightarrow> nat set" where
+ "parents 0 = {0}"
+| "parents n = {parent n} \<union> parents (parent n)"
+
+thm "parents.induct" (* now that's interesting *)
+
+lemma parent_induct: "P (0::nat) \<Longrightarrow> (\<And>i. P (parent i) \<Longrightarrow> P i) \<Longrightarrow> P j"
+proof -
+ obtain j' where "j' = j * 2 + 1"
+ by simp
+ then have "parent j' = j" unfolding parent_def
+ by linarith
+ then show "P (0::nat) \<Longrightarrow> (\<And>i. P (parent i) \<Longrightarrow> P i) \<Longrightarrow> P j"
+ using infinite_descent0 parent_less by blast
+qed
+
+lemma is_heap_min: "h \<noteq> [] \<Longrightarrow> is_heap h \<Longrightarrow> i < length h \<Longrightarrow> h!i \<ge> h!0" unfolding is_heap_def parent_def
+proof (induction i arbitrary: h rule: parent_induct)
+ case 1
+ then show ?case by auto
+next
+ case (2 i)
+ then show ?case
+ using eq_iff leD leI parent_def by fastforce
+qed
+
+lemma swap_len: "length (swap i j h) = length h" unfolding swap_def by auto
+
+lemma swap_pres: "i < length h \<and> j < length h \<Longrightarrow> swap i j (h @ h') = swap i j h @ h'"
+ unfolding swap_def
+ by (simp add: list_update_append1 nth_append)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+lemma sift_len: "i < length h \<Longrightarrow> length (sift_up h i) = length h"
+proof (induction i arbitrary: h rule: parent_induct)
+ case 1
+ then show ?case by auto
+next
+ case (2 i')
+ then show ?case
+ proof (cases i')
+ case 0
+ then show ?thesis by auto
+ next
+ case (Suc nat)
+ then show ?thesis
+ by (metis (no_types, lifting) "2.IH" "2.prems" Suc_lessD less_numeral_extra(3) less_trans_Suc parent_less sift_up.simps(2) swap_len zero_less_Suc)
+ qed
+qed
+
+lemma sift_pres: "i < length h \<Longrightarrow> sift_up (h @ h') i = sift_up h i @ h'"
+proof (induction i arbitrary: h rule: parent_induct)
+case 1
+ then show ?case by auto
+next
+ case (2 i')
+ show ?case
+ proof (cases i')
+ case 0
+ then show ?thesis by auto
+ next
+ case (Suc nat)
+ then show ?thesis
+ proof (cases "h!i' \<ge> h!parent i'")
+ case True
+ then show ?thesis
+ by (smt "2.prems" Suc dual_order.strict_implies_order less_le_trans nth_append old.nat.distinct(2) parent_less sift_up.simps(2))
+ next
+ case False
+ then show ?thesis using Suc swap_pres
+ by (smt "2.IH" "2.prems" less_imp_le_nat less_le_trans nat.simps(3) nth_append parent_less sift_up.simps(2) swap_len)
+ qed
+ qed
+qed
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+thm sift_up.induct
+
+value "is_heap [finite_2.a\<^sub>1, finite_2.a\<^sub>2]"
+
+lemma swap_maybe_restore:
+ assumes "\<not> is_heap h" "is_heap_except j h" "j<length h"
+ shows "is_heap (swap j (parent j) h) \<or> (\<not> is_heap (swap j (parent j) h) \<and> is_heap_except (parent j) (swap j (parent j) h))"
+proof (cases h)
+ case Nil
+ then show ?thesis
+ by (simp add: ex11_tmpl.swap_def is_heap_def)
+next
+ case (Cons a h')
+
+ have almost_heap:
+ "\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!parent i"
+ "\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j)"
+ using assms is_heap_except_def by auto
+ have j_less: "h!j < h!parent j" using assms is_heap_def is_heap_except_def by (metis leI)
+
+ let ?h_s = "swap j (parent j) h"
+
+ have almost_heap2: "\<forall>i<length ?h_s. i\<noteq>parent j \<longrightarrow> ?h_s!i \<ge> ?h_s!parent i" using almost_heap j_less
+ by (smt assms(2) assms(3) dual_order.strict_trans dual_order.trans ex11_tmpl.swap_def is_heap_def leD length_list_update less_imp_le not_first nth_list_update_eq nth_list_update_neq parent_less)
+
+ show ?thesis
+ proof (cases "is_heap ?h_s")
+ case True
+ then show ?thesis
+ by simp
+ next
+ case False
+
+ then have "is_heap_except (parent j) (swap j (parent j) h)" using almost_heap2 almost_heap
+ by (smt assms(3) dual_order.trans eq_iff ex11_tmpl.swap_def is_heap_def is_heap_except_def nth_list_update_eq nth_list_update_neq swap_len)
+
+ then show ?thesis
+ by simp
+ qed
+qed
+
+lemma sift_nop: "j < length h \<Longrightarrow> is_heap h \<Longrightarrow> sift_up h j = h"
+ by (metis is_heap_def sift_up.elims)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using not_first by auto
+next
+ case (2 h v)
+ let ?v = "Suc v"
+ let ?h_s = "swap ?v (parent ?v) h"
+
+ show ?case
+ proof (cases "h!?v \<ge> h!parent ?v")
+ case True
+ then show ?thesis
+ by (metis "2.prems"(1) is_heap_def is_heap_except_def sift_up.elims)
+ next
+ case False
+ then have noheap: "\<not> is_heap h"
+ using "2.prems"(2) is_heap_def by blast
+ from False have "sift_up h ?v = sift_up ?h_s (parent ?v)"
+ by simp
+
+ consider "is_heap ?h_s" | "(\<not> is_heap ?h_s \<and> is_heap_except (parent ?v) ?h_s)"
+ using swap_maybe_restore[of h ?v] using noheap 2 by auto
+
+ then show ?thesis
+ proof cases
+ case 1
+ then have "sift_up h (Suc v) = ?h_s" using sift_nop
+ by (metis "2.prems"(2) \<open>sift_up h (Suc v) = sift_up (swap (Suc v) (parent (Suc v)) h) (parent (Suc v))\<close> dual_order.strict_trans nat.simps(3) parent_less swap_len)
+ then show ?thesis
+ using "1" by auto
+ next
+ case 2
+ then show ?thesis
+ by (metis "2.IH" "2.prems"(2) False \<open>sift_up h (Suc v) = sift_up (swap (Suc v) (parent (Suc v)) h) (parent (Suc v))\<close> dual_order.strict_trans length_code nat.simps(3) parent_less swap_len)
+ qed
+ qed
+qed
+
+lemma swap_mset:
+ assumes "i < length h" "j < length h"
+ shows "mset (swap i j h) = mset h"
+ by (simp add: assms(1) assms(2) ex11_tmpl.swap_def mset_swap)
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case using swap_mset
+ by (metis dual_order.strict_trans nat.simps(3) parent_less sift_up.simps(2) swap_len)
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> hd h"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding emp_def is_heap_def by auto
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by auto
+
+lemma mset_get_min:
+ assumes "is_heap h" "mset h \<noteq> {#}"
+ shows "get_min h = Min_mset (mset h)"
+proof (cases h)
+ case Nil
+ then show ?thesis using assms
+ by simp
+next
+ case (Cons a h')
+
+ then have "\<forall>i < length h. h ! 0 \<le> h ! i" using is_heap_min
+ using assms(1) by blast
+
+ then have "\<forall>e \<in> set h. a \<le> e"
+ by (metis in_set_conv_nth local.Cons nth_Cons_0)
+ moreover
+ have "finite (set h)"
+ by simp
+ ultimately have "a = Min (set h)"
+ by (simp add: Min_insert2 local.Cons)
+
+ moreover
+
+ have "get_min h = a" using Cons
+ by (simp add: get_min_def)
+ moreover
+ have "set_mset (mset h) = set h"
+ by simp
+
+ ultimately show ?thesis
+ by auto
+qed
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+proof (cases h)
+ case Nil
+ then show ?thesis unfolding ins_def
+ by (auto simp add: is_heap_def parent_def)
+next
+ case (Cons a list)
+ assume heap: "is_heap h"
+
+ let ?h' = "h @ [x]"
+
+ let ?j = "length h"
+
+ have "i < length ?h' \<Longrightarrow> (i\<noteq>?j \<longrightarrow> ?h'!i \<ge> ?h'!(parent i))
+ \<and> (parent i = ?j \<longrightarrow> ?h'!i \<ge> ?h'!(parent ?j))" for i
+ proof (cases "i = ?j")
+ case True
+ then show ?thesis
+ by simp
+ next
+ case False
+
+ assume "i < length (h @ [x])"
+ then have "i < length h"
+ by (simp add: False less_Suc_eq)
+ then have "parent i < ?j" unfolding parent_def using parent_less
+ by auto
+ then show ?thesis using heap
+ by (metis \<open>i < length h\<close> is_heap_def neq_iff nth_append)
+ qed
+
+ then have "is_heap_except ?j ?h'" unfolding is_heap_except_def
+ by blast
+
+ then show ?thesis
+ by (simp add: ins_def sift_up_restore_heap)
+qed
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+proof -
+ have "mset (ins x h) = mset (h @ [x])" unfolding ins_def
+ using sift_up_mset[of "length h" "h @ [x]"] by auto
+ also have "\<dots> = mset h + {#x#}"
+ by simp
+ finally show "mset (ins x h) = mset h + {#x#}"
+ by auto
+qed
+
+
+(*<*)
+</submission>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Koepke_Eric_eric.koepke@tum.de_554/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+eric.koepke@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Koepke_Eric_eric.koepke@tum.de_554/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,178 @@
+(** Score: 7/5
+
+ I deducted one point b/c you did not prove the required spec, but only a "roughly equivalent" one.
+*)
+(*<*)
+theory hw11
+ imports
+ Main
+ "~~/src/HOL/Data_Structures/Leftist_Heap"
+begin
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+value "parent 0"
+
+lemma grandparents: "is_heap_except j h \<Longrightarrow> i<length h \<Longrightarrow> i\<noteq>j \<Longrightarrow> h!i \<ge> h!(parent (parent i))"
+ apply (cases "parent i = j")
+ apply (simp add: is_heap_except_def)
+proof -
+ assume a1: "i < length h"
+ assume a2: "is_heap_except j h"
+ assume a3: "parent i \<noteq> j"
+ assume a4: "i \<noteq> j"
+ have f5: "\<not> Suc (Suc i) \<le> Suc (i + 0)"
+ by auto
+ have f6: "\<forall>n na. \<not> n < na \<or> Suc n \<le> na"
+ using Suc_leI by satx
+ have f7: "\<forall>n na. \<not> (n::nat) \<le> na \<or> (\<not> n < na) = (n = na)"
+ by simp
+ have f8: "(i + 1) div 2 - 1 \<le> (i + 1) div 2"
+ by (meson diff_le_self)
+ have f9: "Suc i \<le> length h"
+ using f6 a1 by metis
+ have "(\<forall>n. \<not> n < length h \<or> n = j \<or> h ! parent n \<le> h ! n) \<and> (\<forall>n. \<not> n < length h \<or> parent n \<noteq> j \<or> h ! parent j \<le> h ! n)"
+ using a2 by (meson is_heap_except_def)
+ then show ?thesis
+ using f9 f8 f7 f6 f5 a4 a3 a1 by (metis (no_types) One_nat_def add.right_neutral add_Suc_right add_diff_cancel_right' div_le_dividend dual_order.trans parent_def)
+qed
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+lemma swap_length: "length (swap i j h) = length h"
+ by (simp add: hw11.swap_def)
+
+lemma swap_comp: "j < length h \<Longrightarrow> h ! j \<le> h ! parent j \<Longrightarrow> (swap j (parent j) h) ! j \<ge> (swap j (parent j) h) ! parent j"
+ by (smt One_nat_def Suc_eq_plus1 Suc_leI Suc_pred antisym_conv2 diff_le_self div_le_dividend dual_order.strict_trans hw11.swap_def length_greater_0_conv length_list_update lessI nth_list_update parent_def)
+
+lemma swap_const: "i \<noteq> j \<Longrightarrow> i \<noteq> k \<Longrightarrow> (swap j k h) ! i = h ! i"
+ by (simp add: hw11.swap_def)
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ proof(induction h j rule: sift_up.induct)
+ case c1: (1 h)
+ have "parent 0 = 0" using diff_is_0_eq' div_le_dividend parent_def by presburger
+ then have "h!0 \<ge> h!parent 0" by auto
+ from this c1 show ?case using is_heap_def is_heap_except_def by fastforce
+ next
+ case c1: (2 h v)
+ then have s0: "parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ by (metis dual_order.strict_trans nat.simps(3) parent_less swap_length)
+ from c1(2,3) have s1: "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> (swap (Suc v) (parent (Suc v)) h) ! (Suc v) \<ge> (swap (Suc v) (parent (Suc v)) h) ! (parent(Suc v))"
+ by (simp add: swap_comp)
+ from c1(2,3) have "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> (\<forall>i<length h. parent i = (parent (Suc v)) \<longrightarrow> (swap (Suc v) (parent (Suc v)) h)!i \<ge> (swap (Suc v) (parent (Suc v)) h) ! (parent(Suc v)))"
+ by (metis (no_types, lifting) eq_iff hw11.swap_def is_heap_except_def length_list_update nth_list_update_eq order_trans s0 s1 swap_const)
+ from this s0 c1 have "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ apply (simp add: is_heap_except_def swap_length)
+ by (smt hw11.swap_def le_less_trans not_less nth_list_update_eq nth_list_update_neq)
+ from c1 s0 this show ?case
+ apply auto
+ apply (metis is_heap_def is_heap_except_def)
+ done
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case
+ by (metis dual_order.strict_trans hw11.swap_def mset_swap nat.simps(3) parent_less sift_up.simps(2) swap_length)
+qed
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+
+lemma invar_empty: "is_heap emp"
+ by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (simp add: emp_def)
+
+text \<open> I didn't manage to import the Min_mset function so I proved this instead which should be roughly equivalent.\<close>
+
+lemma get_min_alt: "is_heap h \<Longrightarrow> h \<noteq> [] \<Longrightarrow> i < length h \<Longrightarrow> j\<le>i \<Longrightarrow> get_min h \<le> h!j"
+ proof(induction i arbitrary: j)
+ case 0
+ then show ?case by (simp add: get_min_def)
+ next
+ case (Suc i)
+ then show ?case
+ by (metis Suc_lessD dual_order.trans is_heap_def le_SucE less_Suc_eq_le nat.simps(3) parent_less)
+qed
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma ins_exc: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ apply (auto simp: is_heap_except_def)
+ apply (metis One_nat_def diff_is_0_eq' div_le_dividend dual_order.strict_trans is_heap_def length_Cons less_SucE list.size(3) list.size(4) nth_append parent_def parent_less)
+ by (metis One_nat_def Suc_eq_plus1 diff_is_0_eq div_le_dividend eq_iff not_less_eq parent_def parent_less)
+
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: ins_def ins_exc sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: ins_def sift_up_mset)
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Koepke_Eric_eric.koepke@tum.de_554/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,316 @@
+Using temporary directory '/tmp/tmp.yigNtT6UDd'
+Files in /tmp/eval-554-PFf3iI: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<image>
+HOL-Library
+</image>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<check>
+
+</check>
+<submission-imports>
+
+</submission-imports>
+<check-imports>
+Submission
+</check-imports>
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+value "parent 0"
+
+lemma grandparents: "is_heap_except j h \<Longrightarrow> i<length h \<Longrightarrow> i\<noteq>j \<Longrightarrow> h!i \<ge> h!(parent (parent i))"
+ apply (cases "parent i = j")
+ apply (simp add: is_heap_except_def)
+proof -
+ assume a1: "i < length h"
+ assume a2: "is_heap_except j h"
+ assume a3: "parent i \<noteq> j"
+ assume a4: "i \<noteq> j"
+ have f5: "\<not> Suc (Suc i) \<le> Suc (i + 0)"
+ by auto
+ have f6: "\<forall>n na. \<not> n < na \<or> Suc n \<le> na"
+ using Suc_leI by satx
+ have f7: "\<forall>n na. \<not> (n::nat) \<le> na \<or> (\<not> n < na) = (n = na)"
+ by simp
+ have f8: "(i + 1) div 2 - 1 \<le> (i + 1) div 2"
+ by (meson diff_le_self)
+ have f9: "Suc i \<le> length h"
+ using f6 a1 by metis
+ have "(\<forall>n. \<not> n < length h \<or> n = j \<or> h ! parent n \<le> h ! n) \<and> (\<forall>n. \<not> n < length h \<or> parent n \<noteq> j \<or> h ! parent j \<le> h ! n)"
+ using a2 by (meson is_heap_except_def)
+ then show ?thesis
+ using f9 f8 f7 f6 f5 a4 a3 a1 by (metis (no_types) One_nat_def add.right_neutral add_Suc_right add_diff_cancel_right' div_le_dividend dual_order.trans parent_def)
+qed
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+lemma swap_length: "length (swap i j h) = length h"
+ by (simp add: hw11.swap_def)
+
+lemma swap_comp: "j < length h \<Longrightarrow> h ! j \<le> h ! parent j \<Longrightarrow> (swap j (parent j) h) ! j \<ge> (swap j (parent j) h) ! parent j"
+ by (smt One_nat_def Suc_eq_plus1 Suc_leI Suc_pred antisym_conv2 diff_le_self div_le_dividend dual_order.strict_trans hw11.swap_def length_greater_0_conv length_list_update lessI nth_list_update parent_def)
+
+lemma swap_const: "i \<noteq> j \<Longrightarrow> i \<noteq> k \<Longrightarrow> (swap j k h) ! i = h ! i"
+ by (simp add: hw11.swap_def)
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ proof(induction h j rule: sift_up.induct)
+ case c1: (1 h)
+ have "parent 0 = 0" using diff_is_0_eq' div_le_dividend parent_def by presburger
+ then have "h!0 \<ge> h!parent 0" by auto
+ from this c1 show ?case using is_heap_def is_heap_except_def by fastforce
+ next
+ case c1: (2 h v)
+ then have s0: "parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ by (metis dual_order.strict_trans nat.simps(3) parent_less swap_length)
+ from c1(2,3) have s1: "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> (swap (Suc v) (parent (Suc v)) h) ! (Suc v) \<ge> (swap (Suc v) (parent (Suc v)) h) ! (parent(Suc v))"
+ by (simp add: swap_comp)
+ from c1(2,3) have "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> (\<forall>i<length h. parent i = (parent (Suc v)) \<longrightarrow> (swap (Suc v) (parent (Suc v)) h)!i \<ge> (swap (Suc v) (parent (Suc v)) h) ! (parent(Suc v)))"
+ by (metis (no_types, lifting) eq_iff hw11.swap_def is_heap_except_def length_list_update nth_list_update_eq order_trans s0 s1 swap_const)
+ from this s0 c1 have "h ! parent (Suc v) \<ge> h ! Suc v \<Longrightarrow> is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ apply (simp add: is_heap_except_def swap_length)
+ by (smt hw11.swap_def le_less_trans not_less nth_list_update_eq nth_list_update_neq)
+ from c1 s0 this show ?case
+ apply auto
+ apply (metis is_heap_def is_heap_except_def)
+ done
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case
+ by (metis dual_order.strict_trans hw11.swap_def mset_swap nat.simps(3) parent_less sift_up.simps(2) swap_length)
+qed
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+
+lemma invar_empty: "is_heap emp"
+ by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (simp add: emp_def)
+
+text \<open> I didn't manage to import the Min_mset function so I proved this instead which should be roughly equivalent.\<close>
+
+lemma get_min_alt: "is_heap h \<Longrightarrow> h \<noteq> [] \<Longrightarrow> i < length h \<Longrightarrow> j\<le>i \<Longrightarrow> get_min h \<le> h!j"
+ proof(induction i arbitrary: j)
+ case 0
+ then show ?case by (simp add: get_min_def)
+ next
+ case (Suc i)
+ then show ?case
+ by (metis Suc_lessD dual_order.trans is_heap_def le_SucE less_Suc_eq_le nat.simps(3) parent_less)
+qed
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" good_sorry
+
+lemma ins_exc: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ apply (auto simp: is_heap_except_def)
+ apply (metis One_nat_def diff_is_0_eq' div_le_dividend dual_order.strict_trans is_heap_def length_Cons less_SucE list.size(3) list.size(4) nth_append parent_def parent_less)
+ by (metis One_nat_def Suc_eq_plus1 diff_is_0_eq div_le_dividend eq_iff not_less_eq parent_def parent_less)
+
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: ins_def ins_exc sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: ins_def sift_up_mset)
+
+
+(*<*)
+</submission>
+<version>
+2016-1
+</version>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Krebs_Mitja_mitja.krebs@tum.de_553/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+mitja.krebs@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Krebs_Mitja_mitja.krebs@tum.de_553/ex11_tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,428 @@
+(** Score: 8/5
+ nice job!
+*)
+(*<*)
+theory ex11_tmpl
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
+ "lh_insert x Leaf = Node 1 Leaf x Leaf"
+| "lh_insert x (Node _ l a r) = (
+ if x \<le> a then node l x (lh_insert a r)
+ else node l a (lh_insert x r)
+ )"
+
+lemma mset_tree_node[simp]: "mset_tree (node l a r) = mset_tree l + {#a#} + mset_tree r"
+ by (auto simp: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x) (auto)
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ apply (induction t arbitrary: x)
+ apply (force simp: heap_node mset_lh_insert)+
+ done
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ apply (induction t arbitrary: x)
+ apply (auto simp: ltree_node)
+ done
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat" where
+ "t_lh_insert x Leaf = 1"
+| "t_lh_insert x (Node _ l a r) = (
+ if x \<le> a then t_lh_insert a r
+ else t_lh_insert x r) + 1"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x) (auto)
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | Heap 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+definition empty :: "('a,'s) bs_pq" where
+ "empty = Empty"
+
+fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool" where
+ "is_empty Empty \<longleftrightarrow> True"
+| "is_empty _ \<longleftrightarrow> False"
+
+fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq" where
+ "insert x Empty = Heap x (orig_insert x orig_empty)"
+| "insert x (Heap a s) = Heap (min x a) (orig_insert x s)"
+
+fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a" where
+ "get_min Empty = undefined"
+| "get_min (Heap a s) = a"
+
+fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq" where
+ "del_min Empty = undefined"
+| "del_min (Heap a s) = (let q = orig_del_min s in
+ if orig_is_empty q then Empty
+ else Heap (orig_get_min q) q
+ )"
+
+fun invar :: "('a,'s) bs_pq \<Rightarrow> bool" where
+ "invar Empty \<longleftrightarrow> True"
+| "invar (Heap a s) \<longleftrightarrow> a\<in>#orig_mset s \<and> a = Min_mset (orig_mset s) \<and> orig_invar s"
+
+fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset" where
+ "mset Empty = {#}"
+| "mset (Heap _ s) = orig_mset s"
+
+lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+lemma [simp]: "Min_mset m\<in>#m \<longleftrightarrow> m\<noteq>{#}" by auto
+
+text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+proof goal_cases
+ case 1
+ then show ?case unfolding empty_def by simp
+next
+ case (2 q)
+ then show ?case by (cases q) auto
+next
+case (3 q x)
+ then show ?case by (cases q) (auto)
+next
+ case (4 q)
+ then show ?case by (cases q) (auto)
+next
+ case (5 q)
+ then show ?case by (cases q) (auto)
+next
+ case 6
+ then show ?case unfolding empty_def by simp
+next
+ case (7 q x)
+ then show ?case
+ apply (cases q)
+ apply (auto simp: min_def)
+ done
+next
+case (8 q)
+then show ?case by (cases q) (auto)
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i)
+ )"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma aux1_1:
+ assumes "is_heap_except 0 h"
+ shows "is_heap h"
+proof -
+ have *: "\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)"
+ using assms is_heap_except_def by blast
+ have "parent 0 = 0" by (simp add: parent_def)
+ hence "h!0 \<ge> h!(parent 0)" by simp
+ thus ?thesis using * is_heap_def by auto
+qed
+
+lemma aux1_2: "i < length h \<Longrightarrow> j < length h \<Longrightarrow> mset (swap i j h) = mset h"
+ by (simp add: swap_def mset_swap)
+
+lemma aux1_3: "i < length h \<Longrightarrow> j < length h \<Longrightarrow> length (swap i j h) = length h"
+ by (metis aux1_2 size_mset)
+
+lemma aux1_4:
+ assumes "i < length h" and "j < length h"
+ shows "h!i = (swap i j h)!j" and "h!j = (swap i j h)!i" "\<forall>k<length h. k\<noteq>i \<and> k\<noteq>j \<longrightarrow> h!k = (swap i j h)!k"
+ by (auto simp: assms swap_def nth_list_update)
+
+lemma aux1_5:
+ fixes h :: "'a::linorder list"
+ assumes "j\<noteq>0" "j<length h" "\<not>h!j \<ge> h!parent j"
+ shows "(swap j (parent j) h)!j \<ge> (swap j (parent j) h)!parent j"
+proof -
+ have *: "parent j < length h" using assms parent_less by fastforce
+ have "(\<not>h!j \<ge> h!parent j) = (\<not>(swap j (parent j) h)!parent j \<ge> (swap j (parent j) h)!j)"
+ using aux1_4[OF assms(2) *] by auto
+ thus ?thesis using assms not_le_imp_less dual_order.strict_implies_order by auto
+qed
+
+lemma aux1_6:
+ assumes "j\<noteq>0" "j<length h" "\<not>h!j \<ge> h!parent j" "is_heap_except j h"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ let ?h = "(swap j (parent j) h)"
+ have a: "parent j < length h" using parent_less assms by fastforce
+ note aux1_4 = aux1_4[OF assms(2) a]
+ note aux1_3 = aux1_3[OF assms(2) a]
+ have b: "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<longrightarrow> ?h!i = h!i"
+ using aux1_4 aux1_3 by simp
+ have c: "\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!parent j"
+ using assms is_heap_except_def by blast
+ have d: "\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!parent i"
+ using assms is_heap_except_def by blast
+ (* show "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i\<noteq>j \<and> parent i\<noteq>parent j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ hence *: "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i\<noteq>j \<and> parent i\<noteq>parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using aux1_4 by (metis swap_def length_list_update nth_list_update_neq)
+ (* show "\<forall>i<length ?h. i\<noteq>parent j \<and> i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "?h!j \<ge> ?h!parent j"
+ using aux1_5 assms by blast
+ hence **: "\<forall>i<length ?h. i\<noteq>parent j \<and> i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" by blast
+ (* show: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> i\<noteq>j" by auto
+ hence "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> i\<noteq>j \<and> i\<noteq>parent j" by force
+ hence "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i = h!i"
+ using b by simp
+ hence ***: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using aux1_4 aux1_3 c by auto
+ (* show: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i=parent j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i = h!i"
+ using b aux1_3 by simp
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using d aux1_3 by metis
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!j"
+ using assms(3) by auto
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent j"
+ using aux1_4 by simp
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ by presburger
+ hence ****: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using ** by blast
+ (* show: "\<forall>i<length ?h. parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)" *)
+ have "parent (parent j)\<noteq>parent j \<longrightarrow> parent j\<noteq>0"
+ using parent_def by auto
+ hence e: "parent (parent j)\<noteq>parent j \<longrightarrow> parent (parent j) < parent j"
+ using parent_less by blast
+ hence f: "parent (parent j) < length h" using a by fastforce
+ have "parent (parent j)\<noteq>parent j \<longrightarrow> parent (parent j)\<noteq>j"
+ using e parent_less[OF assms(1)] by linarith
+ hence g: "parent (parent j)\<noteq>parent j \<longrightarrow> ?h!parent (parent j) = h!parent (parent j)"
+ using f aux1_3 b by presburger
+ have "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using aux1_4 by simp
+ hence "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent (parent j)"
+ using a d by fastforce
+ hence h: "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using g by presburger
+ have i: "parent j\<noteq>j" using parent_less[OF assms(1)] by fastforce
+ have "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> i\<noteq>parent j"
+ by blast
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i = h!i"
+ using b by blast
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using aux1_3 d by fastforce
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent (parent j)"
+ using aux1_3 d a i by force
+ hence j: "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using g by presburger
+ have "\<forall>i<length ?h. parent (parent j) = parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using **** by auto
+ hence *****: "\<forall>i<length ?h. parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using h j by blast
+ show ?thesis using is_heap_except_def * ** *** **** ***** by blast
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (simp add: aux1_1)
+next
+ case (2 h v)
+ let ?j = "Suc v"
+ show ?case proof (cases "h!?j \<ge> h!parent ?j")
+ case True
+ have "\<forall>i<length h. i\<noteq>?j \<longrightarrow> h!i \<ge> h!(parent i)"
+ using "2.prems" is_heap_except_def by blast
+ hence "is_heap h" using True is_heap_def by blast
+ thus ?thesis using True by simp
+ next
+ case False
+ have *: "is_heap_except (parent ?j) (swap ?j (parent ?j) h)"
+ using aux1_6 "2.prems" False by blast
+ have **: "parent ?j < length (swap ?j (parent ?j) h)"
+ using "2.prems" parent_less aux1_3 dual_order.strict_trans by (metis nat.simps(3))
+ show ?thesis using "2.IH"[OF False * **] by (simp add: False)
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ let ?j = "Suc v"
+ show ?case proof (cases "h!?j \<ge> h!parent ?j")
+ case True
+ then show ?thesis by simp
+ next
+ case False
+ hence "mset (sift_up h ?j) = mset (sift_up (swap ?j (parent ?j) h) (parent ?j))" by simp
+ also have "... = mset (swap ?j (parent ?j) h)"
+ using "2.IH" "2.prems" aux1_3 False dual_order.strict_trans by fastforce
+ also have "... = mset h"
+ using aux1_2 "2.prems" dual_order.strict_trans parent_less by blast
+ finally show ?thesis .
+ qed
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where
+ "get_min h \<equiv> (case h of
+ [] \<Rightarrow> undefined
+ | (x#xs) \<Rightarrow> fold min xs x
+ )"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x h \<equiv> sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp" by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}" by (simp add: emp_def)
+
+lemma aux1_7:
+ fixes x :: "'a::linorder"
+ shows "fold min xs x = Min_mset (mset (x#xs))"
+ by (metis Min.set_eq_fold set_mset_mset)
+
+lemma aux1_8:
+ fixes x :: "'a::linorder"
+ shows "get_min (x#xs) = fold min xs x"
+ by (simp add: get_min_def)
+
+lemma aux1_9:
+ fixes h :: "'a::linorder list"
+ assumes "h\<noteq>[]"
+ shows "get_min h = Min_mset (mset h)"
+ using assms aux1_7 aux1_8 by (metis list.exhaust)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ by (simp add: aux1_9)
+
+lemma aux1_10:
+ fixes x :: "'a::linorder"
+ and h :: "'a::linorder list"
+ assumes "is_heap h"
+ shows "is_heap_except (length h) (h@[x])"
+proof -
+ have a: "\<forall>i<length h. (h@[x])!i = h!i" by (simp add: nth_append)
+ have "\<forall>i<length h. parent i<length h" using parent_def by fastforce
+ hence b: "\<forall>i<length h. (h@[x])!parent i = h!parent i" using a by blast
+ have "\<forall>i<length h. h!i \<ge> h!parent i" using assms is_heap_def by blast
+ hence "\<forall>i<length h. (h@[x])!i \<ge> (h@[x])!parent i" using a b by presburger
+ hence *: "\<forall>i<length h + 1. i\<noteq>length h \<longrightarrow> (h@[x])!i \<ge> (h@[x])!parent i" by simp
+ have "\<forall>i<length h + 1. parent i = length h \<longrightarrow> i = parent i"
+ using parent_def by fastforce
+ hence **: "\<forall>i<length h + 1. parent i = length h \<longrightarrow> (h@[x])!i \<ge> (h@[x])!parent (length h)"
+ by fastforce
+ have "length (h@[x]) = length h + 1" by force
+ thus ?thesis using * ** by (simp add: is_heap_except_def)
+qed
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: aux1_10 ins_def sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: aux1_10 ins_def sift_up_mset)
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Krebs_Mitja_mitja.krebs@tum.de_553/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,568 @@
+Using temporary directory '/tmp/tmp.lWDMziTAGL'
+Files in /tmp/eval-553-SVlVcH: ex11_tmpl.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<check-imports>
+Submission
+</check-imports>
+<submission-imports>
+
+</submission-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<check>
+
+</check>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
+ "lh_insert x Leaf = Node 1 Leaf x Leaf"
+| "lh_insert x (Node _ l a r) = (
+ if x \<le> a then node l x (lh_insert a r)
+ else node l a (lh_insert x r)
+ )"
+
+lemma mset_tree_node[simp]: "mset_tree (node l a r) = mset_tree l + {#a#} + mset_tree r"
+ by (auto simp: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x) (auto)
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ apply (induction t arbitrary: x)
+ apply (force simp: heap_node mset_lh_insert)+
+ done
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ apply (induction t arbitrary: x)
+ apply (auto simp: ltree_node)
+ done
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat" where
+ "t_lh_insert x Leaf = 1"
+| "t_lh_insert x (Node _ l a r) = (
+ if x \<le> a then t_lh_insert a r
+ else t_lh_insert x r) + 1"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x) (auto)
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | Heap 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+definition empty :: "('a,'s) bs_pq" where
+ "empty = Empty"
+
+fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool" where
+ "is_empty Empty \<longleftrightarrow> True"
+| "is_empty _ \<longleftrightarrow> False"
+
+fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq" where
+ "insert x Empty = Heap x (orig_insert x orig_empty)"
+| "insert x (Heap a s) = Heap (min x a) (orig_insert x s)"
+
+fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a" where
+ "get_min Empty = undefined"
+| "get_min (Heap a s) = a"
+
+fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq" where
+ "del_min Empty = undefined"
+| "del_min (Heap a s) = (let q = orig_del_min s in
+ if orig_is_empty q then Empty
+ else Heap (orig_get_min q) q
+ )"
+
+fun invar :: "('a,'s) bs_pq \<Rightarrow> bool" where
+ "invar Empty \<longleftrightarrow> True"
+| "invar (Heap a s) \<longleftrightarrow> a\<in>#orig_mset s \<and> a = Min_mset (orig_mset s) \<and> orig_invar s"
+
+fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset" where
+ "mset Empty = {#}"
+| "mset (Heap _ s) = orig_mset s"
+
+lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+lemma [simp]: "Min_mset m\<in>#m \<longleftrightarrow> m\<noteq>{#}" by auto
+
+text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+proof goal_cases
+ case 1
+ then show ?case unfolding empty_def by simp
+next
+ case (2 q)
+ then show ?case by (cases q) auto
+next
+case (3 q x)
+ then show ?case by (cases q) (auto)
+next
+ case (4 q)
+ then show ?case by (cases q) (auto)
+next
+ case (5 q)
+ then show ?case by (cases q) (auto)
+next
+ case 6
+ then show ?case unfolding empty_def by simp
+next
+ case (7 q x)
+ then show ?case
+ apply (cases q)
+ apply (auto simp: min_def)
+ done
+next
+case (8 q)
+then show ?case by (cases q) (auto)
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i)
+ )"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma aux1_1:
+ assumes "is_heap_except 0 h"
+ shows "is_heap h"
+proof -
+ have *: "\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)"
+ using assms is_heap_except_def by blast
+ have "parent 0 = 0" by (simp add: parent_def)
+ hence "h!0 \<ge> h!(parent 0)" by simp
+ thus ?thesis using * is_heap_def by auto
+qed
+
+lemma aux1_2: "i < length h \<Longrightarrow> j < length h \<Longrightarrow> mset (swap i j h) = mset h"
+ by (simp add: swap_def mset_swap)
+
+lemma aux1_3: "i < length h \<Longrightarrow> j < length h \<Longrightarrow> length (swap i j h) = length h"
+ by (metis aux1_2 size_mset)
+
+lemma aux1_4:
+ assumes "i < length h" and "j < length h"
+ shows "h!i = (swap i j h)!j" and "h!j = (swap i j h)!i" "\<forall>k<length h. k\<noteq>i \<and> k\<noteq>j \<longrightarrow> h!k = (swap i j h)!k"
+ by (auto simp: assms swap_def nth_list_update)
+
+lemma aux1_5:
+ fixes h :: "'a::linorder list"
+ assumes "j\<noteq>0" "j<length h" "\<not>h!j \<ge> h!parent j"
+ shows "(swap j (parent j) h)!j \<ge> (swap j (parent j) h)!parent j"
+proof -
+ have *: "parent j < length h" using assms parent_less by fastforce
+ have "(\<not>h!j \<ge> h!parent j) = (\<not>(swap j (parent j) h)!parent j \<ge> (swap j (parent j) h)!j)"
+ using aux1_4[OF assms(2) *] by auto
+ thus ?thesis using assms not_le_imp_less dual_order.strict_implies_order by auto
+qed
+
+lemma aux1_6:
+ assumes "j\<noteq>0" "j<length h" "\<not>h!j \<ge> h!parent j" "is_heap_except j h"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ let ?h = "(swap j (parent j) h)"
+ have a: "parent j < length h" using parent_less assms by fastforce
+ note aux1_4 = aux1_4[OF assms(2) a]
+ note aux1_3 = aux1_3[OF assms(2) a]
+ have b: "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<longrightarrow> ?h!i = h!i"
+ using aux1_4 aux1_3 by simp
+ have c: "\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!parent j"
+ using assms is_heap_except_def by blast
+ have d: "\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!parent i"
+ using assms is_heap_except_def by blast
+ (* show "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i\<noteq>j \<and> parent i\<noteq>parent j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ hence *: "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i\<noteq>j \<and> parent i\<noteq>parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using aux1_4 by (metis swap_def length_list_update nth_list_update_neq)
+ (* show "\<forall>i<length ?h. i\<noteq>parent j \<and> i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "?h!j \<ge> ?h!parent j"
+ using aux1_5 assms by blast
+ hence **: "\<forall>i<length ?h. i\<noteq>parent j \<and> i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" by blast
+ (* show: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> i\<noteq>j" by auto
+ hence "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> i\<noteq>j \<and> i\<noteq>parent j" by force
+ hence "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i = h!i"
+ using b by simp
+ hence ***: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using aux1_4 aux1_3 c by auto
+ (* show: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i=parent j \<longrightarrow> ?h!i \<ge> ?h!parent i" *)
+ have "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i = h!i"
+ using b aux1_3 by simp
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using d aux1_3 by metis
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!j"
+ using assms(3) by auto
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent j"
+ using aux1_4 by simp
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ by presburger
+ hence ****: "\<forall>i<length ?h. i\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent i"
+ using ** by blast
+ (* show: "\<forall>i<length ?h. parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)" *)
+ have "parent (parent j)\<noteq>parent j \<longrightarrow> parent j\<noteq>0"
+ using parent_def by auto
+ hence e: "parent (parent j)\<noteq>parent j \<longrightarrow> parent (parent j) < parent j"
+ using parent_less by blast
+ hence f: "parent (parent j) < length h" using a by fastforce
+ have "parent (parent j)\<noteq>parent j \<longrightarrow> parent (parent j)\<noteq>j"
+ using e parent_less[OF assms(1)] by linarith
+ hence g: "parent (parent j)\<noteq>parent j \<longrightarrow> ?h!parent (parent j) = h!parent (parent j)"
+ using f aux1_3 b by presburger
+ have "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using aux1_4 by simp
+ hence "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent (parent j)"
+ using a d by fastforce
+ hence h: "\<forall>i<length ?h. i = j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using g by presburger
+ have i: "parent j\<noteq>j" using parent_less[OF assms(1)] by fastforce
+ have "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> i\<noteq>parent j"
+ by blast
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i = h!i"
+ using b by blast
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent j"
+ using aux1_3 d by fastforce
+ hence "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> h!parent (parent j)"
+ using aux1_3 d a i by force
+ hence j: "\<forall>i<length ?h. i\<noteq>j \<and> parent (parent j)\<noteq>parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using g by presburger
+ have "\<forall>i<length ?h. parent (parent j) = parent j \<and> parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using **** by auto
+ hence *****: "\<forall>i<length ?h. parent i = parent j \<longrightarrow> ?h!i \<ge> ?h!parent (parent j)"
+ using h j by blast
+ show ?thesis using is_heap_except_def * ** *** **** ***** by blast
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (simp add: aux1_1)
+next
+ case (2 h v)
+ let ?j = "Suc v"
+ show ?case proof (cases "h!?j \<ge> h!parent ?j")
+ case True
+ have "\<forall>i<length h. i\<noteq>?j \<longrightarrow> h!i \<ge> h!(parent i)"
+ using "2.prems" is_heap_except_def by blast
+ hence "is_heap h" using True is_heap_def by blast
+ thus ?thesis using True by simp
+ next
+ case False
+ have *: "is_heap_except (parent ?j) (swap ?j (parent ?j) h)"
+ using aux1_6 "2.prems" False by blast
+ have **: "parent ?j < length (swap ?j (parent ?j) h)"
+ using "2.prems" parent_less aux1_3 dual_order.strict_trans by (metis nat.simps(3))
+ show ?thesis using "2.IH"[OF False * **] by (simp add: False)
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ let ?j = "Suc v"
+ show ?case proof (cases "h!?j \<ge> h!parent ?j")
+ case True
+ then show ?thesis by simp
+ next
+ case False
+ hence "mset (sift_up h ?j) = mset (sift_up (swap ?j (parent ?j) h) (parent ?j))" by simp
+ also have "... = mset (swap ?j (parent ?j) h)"
+ using "2.IH" "2.prems" aux1_3 False dual_order.strict_trans by fastforce
+ also have "... = mset h"
+ using aux1_2 "2.prems" dual_order.strict_trans parent_less by blast
+ finally show ?thesis .
+ qed
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where
+ "get_min h \<equiv> (case h of
+ [] \<Rightarrow> undefined
+ | (x#xs) \<Rightarrow> fold min xs x
+ )"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x h \<equiv> sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp" by (simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}" by (simp add: emp_def)
+
+lemma aux1_7:
+ fixes x :: "'a::linorder"
+ shows "fold min xs x = Min_mset (mset (x#xs))"
+ by (metis Min.set_eq_fold set_mset_mset)
+
+lemma aux1_8:
+ fixes x :: "'a::linorder"
+ shows "get_min (x#xs) = fold min xs x"
+ by (simp add: get_min_def)
+
+lemma aux1_9:
+ fixes h :: "'a::linorder list"
+ assumes "h\<noteq>[]"
+ shows "get_min h = Min_mset (mset h)"
+ using assms aux1_7 aux1_8 by (metis list.exhaust)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ by (simp add: aux1_9)
+
+lemma aux1_10:
+ fixes x :: "'a::linorder"
+ and h :: "'a::linorder list"
+ assumes "is_heap h"
+ shows "is_heap_except (length h) (h@[x])"
+proof -
+ have a: "\<forall>i<length h. (h@[x])!i = h!i" by (simp add: nth_append)
+ have "\<forall>i<length h. parent i<length h" using parent_def by fastforce
+ hence b: "\<forall>i<length h. (h@[x])!parent i = h!parent i" using a by blast
+ have "\<forall>i<length h. h!i \<ge> h!parent i" using assms is_heap_def by blast
+ hence "\<forall>i<length h. (h@[x])!i \<ge> (h@[x])!parent i" using a b by presburger
+ hence *: "\<forall>i<length h + 1. i\<noteq>length h \<longrightarrow> (h@[x])!i \<ge> (h@[x])!parent i" by simp
+ have "\<forall>i<length h + 1. parent i = length h \<longrightarrow> i = parent i"
+ using parent_def by fastforce
+ hence **: "\<forall>i<length h + 1. parent i = length h \<longrightarrow> (h@[x])!i \<ge> (h@[x])!parent (length h)"
+ by fastforce
+ have "length (h@[x]) = length h + 1" by force
+ thus ?thesis using * ** by (simp add: is_heap_except_def)
+qed
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: aux1_10 ins_def sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: aux1_10 ins_def sift_up_mset)
+
+(*<*)
+</submission>
+<image>
+HOL-Library
+</image>
+<version>
+2016-1
+</version>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kutasi_Daniel_daniel.kutasi@mytum.de_561/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+daniel.kutasi@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kutasi_Daniel_daniel.kutasi@mytum.de_561/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,123 @@
+(** Score: 2/5
+*)
+(*<*)
+theory hw11
+imports "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma [simp]: "\<lbrakk>is_heap_except j h; j < length h; h ! parent j \<le> h ! j\<rbrakk>
+ \<Longrightarrow> is_heap h"
+ sorry
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (auto simp: is_heap_except_def parent_def is_heap_def)
+next
+ case (2 h j)
+ from "2.prems" show ?case apply auto apply (rule "2.IH") apply (auto simp: is_heap_except_def nth_append nth_list_update) sorry
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto simp: swap_def)
+ by (smt le_less_trans mset_swap nat.simps(3) not_less order.asym parent_less)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h!0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h @ [x]) (length h)"
+
+lemma invar_empty: "is_heap emp" by (simp add: is_heap_def emp_def)
+
+lemma mset_empty: "mset emp = {#}" by (simp add: emp_def)
+
+lemma [simp]: "is_heap h \<Longrightarrow> \<forall>x . x < length h \<Longrightarrow> h!0 \<le> h!x"
+ by auto
+
+lemma [simp]: "is_heap (x # y # h) \<Longrightarrow> x \<le> y"
+ by (induction h) (auto simp: is_heap_def parent_def)
+
+lemma [simp]: "is_heap (x # y # h) \<Longrightarrow> \<forall>a \<in> set h . x \<le> a"
+ apply (induction h)
+ apply (auto) (* loops if I add is_heap_def to the simp set, because of the unfolding of \<le>
+ over all following elements. I don't know how else
+ to describe the transitivity of \<ge> such that the following lemmas
+ can use it. Therefore, I'm sorry *)
+
+(** If apply auto with some simpset does not work, do your argument step-by-step in Isar ... *)
+
+ sorry
+
+lemma [simp]: "is_heap (i#h) \<Longrightarrow> \<forall>x \<in> set h. i \<le> x"
+ by (induction h) (auto)
+
+lemma [simp]: "is_heap h \<Longrightarrow> i=h!0 \<Longrightarrow> \<forall>x \<in> set h. i \<le> x"
+ by (induction h) (auto)
+
+lemma [simp]: "is_heap (a # h) \<Longrightarrow> a = Min (Set.insert a (set h))"
+ by (induction h) (auto simp: min_def)
+
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ by (induction h) (auto simp: get_min_def)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Kutasi_Daniel_daniel.kutasi@mytum.de_561/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,262 @@
+Using temporary directory '/tmp/tmp.QumbFpoNBg'
+Files in /tmp/eval-561-g63XuU: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<defs-imports>
+Complex_Main
+</defs-imports>
+<submission-imports>
+
+</submission-imports>
+<check-imports>
+Submission
+</check-imports>
+<check>
+
+</check>
+<submission>
+
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma [simp]: "\<lbrakk>is_heap_except j h; j < length h; h ! parent j \<le> h ! j\<rbrakk>
+ \<Longrightarrow> is_heap h"
+ good_sorry
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (auto simp: is_heap_except_def parent_def is_heap_def)
+next
+ case (2 h j)
+ from "2.prems" show ?case apply auto apply (rule "2.IH") apply (auto simp: is_heap_except_def nth_append nth_list_update) good_sorry
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto simp: swap_def)
+ by (smt le_less_trans mset_swap nat.simps(3) not_less order.asym parent_less)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h!0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h @ [x]) (length h)"
+
+lemma invar_empty: "is_heap emp" by (simp add: is_heap_def emp_def)
+
+lemma mset_empty: "mset emp = {#}" by (simp add: emp_def)
+
+lemma [simp]: "is_heap h \<Longrightarrow> \<forall>x . x < length h \<Longrightarrow> h!0 \<le> h!x"
+ by auto
+
+lemma [simp]: "is_heap (x # y # h) \<Longrightarrow> x \<le> y"
+ by (induction h) (auto simp: is_heap_def parent_def)
+
+lemma [simp]: "is_heap (x # y # h) \<Longrightarrow> \<forall>a \<in> set h . x \<le> a"
+ apply (induction h)
+ apply (auto) (* loops if I add is_heap_def to the simp set, because of the unfolding of \<le>
+ over all following elements. I don't know how else
+ to describe the transitivity of \<ge> such that the following lemmas
+ can use it. Therefore, I'm good_sorry *)
+ good_sorry
+
+lemma [simp]: "is_heap (i#h) \<Longrightarrow> \<forall>x \<in> set h. i \<le> x"
+ by (induction h) (auto)
+
+lemma [simp]: "is_heap h \<Longrightarrow> i=h!0 \<Longrightarrow> \<forall>x \<in> set h. i \<le> x"
+ by (induction h) (auto)
+
+lemma [simp]: "is_heap (a # h) \<Longrightarrow> a = Min (Set.insert a (set h))"
+ by (induction h) (auto simp: min_def)
+
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ by (induction h) (auto simp: get_min_def)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<image>
+HOL-Library
+</image>
+<version>
+2016-1
+</version>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Mutius_Joshua_ga96koz@mytum.de_560/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+ga96koz@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Mutius_Joshua_ga96koz@mytum.de_560/ex11_tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,241 @@
+(** Score: 5/5
+*)
+(*<*)
+theory ex11_tmpl
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+
+find_theorems "_[_:=_, _:=_] ! _"
+thm mset_swap length_list_update
+
+
+lemma len_swap: "length h = length (swap i j h)"
+ unfolding swap_def by simp
+
+lemma except_start: "is_heap_except 0 h \<Longrightarrow> is_heap h" unfolding is_heap_def is_heap_except_def parent_def
+ by force
+
+ lemma aux_except: "\<not> is_heap h \<Longrightarrow> is_heap_except j h \<Longrightarrow> h!j < h ! (parent j )"
+ unfolding is_heap_def is_heap_except_def parent_def by auto
+
+
+lemma except2: assumes "j < length h" shows "\<not> is_heap h \<and> is_heap_except j h \<longleftrightarrow> h!j < h ! (parent j ) \<and> is_heap_except j h"
+proof
+ assume P: "h!j < h ! (parent j ) \<and> is_heap_except j h"
+ then have L: "\<not> is_heap h" unfolding is_heap_def using assms by force
+ have R: "is_heap_except j h"
+ using P unfolding is_heap_except_def by simp
+ show "\<not> is_heap h \<and> is_heap_except j h" using R L by simp
+
+ next
+ assume P: "\<not> is_heap h \<and> is_heap_except j h"
+ show "h!j < h ! (parent j ) \<and> is_heap_except j h" using P aux_except by auto
+qed
+
+
+lemma swap_aux: assumes "j < length h" and "h' = swap j (parent j) h" shows "h' ! j = h ! parent j \<and> h'! parent j = h ! j"
+ unfolding swap_def
+proof -
+ let ?h_fst = "h[j:= h ! (parent j)]"
+ show ?thesis
+ proof (cases "j = parent j")
+ case True
+ then have "j = 0" unfolding parent_def by simp
+ then have "h' = h [j := h ! j]"
+ using list_update_overwrite[of h j "h ! parent j" "h ! j"] assms True unfolding swap_def by simp
+ then have "h' ! j = h ! j" by simp
+ then show ?thesis using True by simp
+ next
+ case False
+ then have fst: "h' = ?h_fst[parent j := h ! j]" using assms swap_def by auto
+ then have "h' ! j = ?h_fst ! parent j " using parent_def
+ using False assms(1) by auto
+ also have "... = h ! (parent j)"
+ by (simp add: False)
+ finally have L: "h' ! j = h ! parent j" by simp
+ have snd: "h' ! parent j = h ! j"
+ using assms fst False nth_list_update_eq[of "parent j" ?h_fst "h ! j"] unfolding parent_def by auto
+ then show ?thesis using L by blast
+ qed
+qed
+lemma rewrite_except:
+ assumes "\<not> is_heap h" and "is_heap_except j h" and "j < length h \<and> j \<noteq> 0" and "sw = swap j (parent j) h"
+ shows "is_heap_except (parent j) sw"
+proof -
+ have "is_heap_except j h" using assms by simp
+ then have except_std: "(\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+ unfolding is_heap_except_def by simp
+ have eq: "\<forall>x. x < length h \<and> x \<noteq>j \<and> x \<noteq> (parent j) \<longrightarrow> sw ! x = h ! x"
+ using len_swap assms
+ by (simp add: swap_def)
+ have bad_elem: "h ! j < h ! (parent j)"
+ using assms aux_except by blast
+ have diff: "sw ! j = h ! parent j \<and> sw ! (parent j) = h ! j" using assms by (simp add: swap_aux assms(3))
+ have diff1: "sw ! j \<ge> sw ! (parent j)" using diff bad_elem by simp
+ have "is_heap_except (parent j) sw" unfolding is_heap_except_def
+ by (smt One_nat_def Suc_eq_plus1 assms(4) bad_elem diff diff_is_0_eq'
+ dual_order.strict_trans eq except_std leD le_numeral_extra(1) len_swap
+ linear one_div_two_eq_zero order_trans parent_def parent_less)
+ then show ?thesis by simp
+qed
+
+lemma is_heap_sift: "j < length h \<Longrightarrow> is_heap h \<Longrightarrow> sift_up h j = h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ have "h ! (Suc v) \<ge> h ! parent (Suc v)" using "2.prems" unfolding is_heap_def by simp
+ then have "sift_up h (Suc v) = h" by simp
+ then show ?case.
+qed
+
+lemma not_heap_sift: "\<lbrakk>is_heap_except j h; h' = sift_up h j; j < length h\<rbrakk> \<Longrightarrow> \<forall>i < length h'. h' ! i \<ge> h' ! parent i"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using except_start is_heap_def by auto
+next
+ case (2 h v)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then have "h = h'" using 2 is_heap_sift[of "Suc v" h] by simp
+ then show ?thesis using True is_heap_def by auto
+ next
+ case False
+ then have A:"\<not> h ! parent (Suc v) \<le> h ! Suc v" using aux_except[of h "Suc v"] "2.prems" by simp
+ have B:"is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ using "2.prems" False rewrite_except[of h "Suc v"] by simp
+ have C:"parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ using parent_less[of "Suc v"] "2.prems"(3) len_swap[of h "Suc v" "parent (Suc v)"]
+ by linarith
+ then show ?thesis using A B "2.prems" "2.IH" by auto
+ qed
+qed
+
+lemma sift_up_one: assumes "is_heap_except j h" and "h' = sift_up h j" and " j < length h"
+ shows "\<forall>i < length h'. h' ! i \<ge> h' ! parent i"
+proof (cases "is_heap h")
+ case True
+ then have "h = h'" using assms is_heap_sift[of j h] by simp
+ then show ?thesis using True is_heap_def by auto
+next
+ case False
+ then show ?thesis using not_heap_sift assms by blast
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using except_start by simp
+next
+ case (2 h v)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then show ?thesis using is_heap_sift[of "(Suc v)" h] "2.prems" by simp
+ next
+ case False
+ let ?sift_h = "sift_up h (Suc v)"
+ have "\<forall>i < length ?sift_h. ?sift_h ! i \<ge> ?sift_h ! parent i"
+ using sift_up_one[of "Suc v" h "?sift_h"] "2.prems" by auto
+ then have "is_heap ?sift_h" unfolding is_heap_def by simp
+ then show ?thesis.
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto simp: swap_def mset_swap)
+ using dual_order.strict_trans parent_less by blast
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+fun ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x [] = [x]" |
+ "ins x (y#ys) = (if x \<ge> y then y#ins x ys else x # ins y ys)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding is_heap_def emp_def by simp
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by simp
+
+lemma "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> \<forall>i < length h. h ! 0 \<le> h ! i"
+proof (rule ccontr)
+ assume "\<not>(\<forall>i < length h. h ! 0 \<le> h ! i)"
+ then have "is_heap h = False" sorry
+ then show "False" sorry
+qed
+
+
+
+find_theorems " Min_mset _"
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ unfolding is_heap_def get_min_def parent_def
+ sorry
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Mutius_Joshua_ga96koz@mytum.de_560/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,382 @@
+Using temporary directory '/tmp/tmp.NdGodFOpww'
+Files in /tmp/eval-560-uZPrKR: ex11_tmpl.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<version>
+2016-1
+</version>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<check>
+
+</check>
+<check-imports>
+Submission
+</check-imports>
+<image>
+HOL-Library
+</image>
+<submission-imports>
+
+</submission-imports>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<submission>
+
+(*>*)
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+
+find_theorems "_[_:=_, _:=_] ! _"
+thm mset_swap length_list_update
+
+
+lemma len_swap: "length h = length (swap i j h)"
+ unfolding swap_def by simp
+
+lemma except_start: "is_heap_except 0 h \<Longrightarrow> is_heap h" unfolding is_heap_def is_heap_except_def parent_def
+ by force
+
+ lemma aux_except: "\<not> is_heap h \<Longrightarrow> is_heap_except j h \<Longrightarrow> h!j < h ! (parent j )"
+ unfolding is_heap_def is_heap_except_def parent_def by auto
+
+
+lemma except2: assumes "j < length h" shows "\<not> is_heap h \<and> is_heap_except j h \<longleftrightarrow> h!j < h ! (parent j ) \<and> is_heap_except j h"
+proof
+ assume P: "h!j < h ! (parent j ) \<and> is_heap_except j h"
+ then have L: "\<not> is_heap h" unfolding is_heap_def using assms by force
+ have R: "is_heap_except j h"
+ using P unfolding is_heap_except_def by simp
+ show "\<not> is_heap h \<and> is_heap_except j h" using R L by simp
+
+ next
+ assume P: "\<not> is_heap h \<and> is_heap_except j h"
+ show "h!j < h ! (parent j ) \<and> is_heap_except j h" using P aux_except by auto
+qed
+
+
+lemma swap_aux: assumes "j < length h" and "h' = swap j (parent j) h" shows "h' ! j = h ! parent j \<and> h'! parent j = h ! j"
+ unfolding swap_def
+proof -
+ let ?h_fst = "h[j:= h ! (parent j)]"
+ show ?thesis
+ proof (cases "j = parent j")
+ case True
+ then have "j = 0" unfolding parent_def by simp
+ then have "h' = h [j := h ! j]"
+ using list_update_overwrite[of h j "h ! parent j" "h ! j"] assms True unfolding swap_def by simp
+ then have "h' ! j = h ! j" by simp
+ then show ?thesis using True by simp
+ next
+ case False
+ then have fst: "h' = ?h_fst[parent j := h ! j]" using assms swap_def by auto
+ then have "h' ! j = ?h_fst ! parent j " using parent_def
+ using False assms(1) by auto
+ also have "... = h ! (parent j)"
+ by (simp add: False)
+ finally have L: "h' ! j = h ! parent j" by simp
+ have snd: "h' ! parent j = h ! j"
+ using assms fst False nth_list_update_eq[of "parent j" ?h_fst "h ! j"] unfolding parent_def by auto
+ then show ?thesis using L by blast
+ qed
+qed
+lemma rewrite_except:
+ assumes "\<not> is_heap h" and "is_heap_except j h" and "j < length h \<and> j \<noteq> 0" and "sw = swap j (parent j) h"
+ shows "is_heap_except (parent j) sw"
+proof -
+ have "is_heap_except j h" using assms by simp
+ then have except_std: "(\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+ unfolding is_heap_except_def by simp
+ have eq: "\<forall>x. x < length h \<and> x \<noteq>j \<and> x \<noteq> (parent j) \<longrightarrow> sw ! x = h ! x"
+ using len_swap assms
+ by (simp add: swap_def)
+ have bad_elem: "h ! j < h ! (parent j)"
+ using assms aux_except by blast
+ have diff: "sw ! j = h ! parent j \<and> sw ! (parent j) = h ! j" using assms by (simp add: swap_aux assms(3))
+ have diff1: "sw ! j \<ge> sw ! (parent j)" using diff bad_elem by simp
+ have "is_heap_except (parent j) sw" unfolding is_heap_except_def
+ by (smt One_nat_def Suc_eq_plus1 assms(4) bad_elem diff diff_is_0_eq'
+ dual_order.strict_trans eq except_std leD le_numeral_extra(1) len_swap
+ linear one_div_two_eq_zero order_trans parent_def parent_less)
+ then show ?thesis by simp
+qed
+
+lemma is_heap_sift: "j < length h \<Longrightarrow> is_heap h \<Longrightarrow> sift_up h j = h"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by simp
+next
+ case (2 h v)
+ have "h ! (Suc v) \<ge> h ! parent (Suc v)" using "2.prems" unfolding is_heap_def by simp
+ then have "sift_up h (Suc v) = h" by simp
+ then show ?case.
+qed
+
+lemma not_heap_sift: "\<lbrakk>is_heap_except j h; h' = sift_up h j; j < length h\<rbrakk> \<Longrightarrow> \<forall>i < length h'. h' ! i \<ge> h' ! parent i"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using except_start is_heap_def by auto
+next
+ case (2 h v)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then have "h = h'" using 2 is_heap_sift[of "Suc v" h] by simp
+ then show ?thesis using True is_heap_def by auto
+ next
+ case False
+ then have A:"\<not> h ! parent (Suc v) \<le> h ! Suc v" using aux_except[of h "Suc v"] "2.prems" by simp
+ have B:"is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ using "2.prems" False rewrite_except[of h "Suc v"] by simp
+ have C:"parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ using parent_less[of "Suc v"] "2.prems"(3) len_swap[of h "Suc v" "parent (Suc v)"]
+ by linarith
+ then show ?thesis using A B "2.prems" "2.IH" by auto
+ qed
+qed
+
+lemma sift_up_one: assumes "is_heap_except j h" and "h' = sift_up h j" and " j < length h"
+ shows "\<forall>i < length h'. h' ! i \<ge> h' ! parent i"
+proof (cases "is_heap h")
+ case True
+ then have "h = h'" using assms is_heap_sift[of j h] by simp
+ then show ?thesis using True is_heap_def by auto
+next
+ case False
+ then show ?thesis using not_heap_sift assms by blast
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case using except_start by simp
+next
+ case (2 h v)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then show ?thesis using is_heap_sift[of "(Suc v)" h] "2.prems" by simp
+ next
+ case False
+ let ?sift_h = "sift_up h (Suc v)"
+ have "\<forall>i < length ?sift_h. ?sift_h ! i \<ge> ?sift_h ! parent i"
+ using sift_up_one[of "Suc v" h "?sift_h"] "2.prems" by auto
+ then have "is_heap ?sift_h" unfolding is_heap_def by simp
+ then show ?thesis.
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto simp: swap_def mset_swap)
+ using dual_order.strict_trans parent_less by blast
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+fun ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x [] = [x]" |
+ "ins x (y#ys) = (if x \<ge> y then y#ins x ys else x # ins y ys)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding is_heap_def emp_def by simp
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by simp
+
+lemma "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> \<forall>i < length h. h ! 0 \<le> h ! i"
+proof (rule ccontr)
+ assume "\<not>(\<forall>i < length h. h ! 0 \<le> h ! i)"
+ then have "is_heap h = False" good_sorry
+ then show "False" good_sorry
+qed
+
+
+
+find_theorems " Min_mset _"
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ unfolding is_heap_def get_min_def parent_def
+ good_sorry
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Pfeiffer_Marcus_ge72lic@mytum.de_555/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+ge72lic@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Pfeiffer_Marcus_ge72lic@mytum.de_555/ex11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,130 @@
+(** Score: 1/5
+*)
+(*<*)
+theory ex11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(cases j)
+ assume A1:"is_heap_except j h"
+ case 0
+ have "parent j = 0" using "0" unfolding parent_def by auto
+ hence T1: "h!j \<ge> h!parent j"
+ by (simp add: "0")
+ have T2: "sift_up h j = h"
+ by (simp add: "0")
+ also have "(\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))" using A1 by (simp add: is_heap_except_def)
+ hence "\<forall>i<length h. h!i \<ge> h!parent i"
+ using T1 by blast
+ then show ?thesis
+ by (simp add: T2 is_heap_def)
+next
+ assume A1:"is_heap_except j h"
+ case (Suc nat)
+ then show ?thesis
+ proof (cases "h!j \<ge> h!parent j")
+ case True
+ then show ?thesis using A1
+ using Suc is_heap_def is_heap_except_def by force
+ next
+ case F: False
+ have "\<not>(h!j \<ge> h!parent j)"
+ by (simp add: F)
+ hence "sift_up h j = sift_up (swap j (parent j) h) (parent j)"
+ by (simp add: Suc)
+ then show ?thesis sorry
+ qed
+qed
+
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof(cases j)
+ case 0
+ then show ?thesis
+ by simp
+next
+ assume A1: "j < length h"
+ case (Suc nat)
+ then show ?thesis
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+then show ?case by auto
+next
+ case (2 h v)
+ then show ?case sorry
+qed
+qed
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" sorry
+
+lemma mset_empty: "mset emp = {#}" sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Pfeiffer_Marcus_ge72lic@mytum.de_555/testoutput.html Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>No test cases passed! -- Is your submission compiling?</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Failed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Putwattana_Attakorn_a.putwattana@tum.de_552/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+a.putwattana@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Putwattana_Attakorn_a.putwattana@tum.de_552/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,131 @@
+(** Score: 6/5
+*)
+(*<*)
+theory hw11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+(* auxiliary lemma *)
+lemma aux_basis:
+ "is_heap_except 0 h \<Longrightarrow> 0 < length h \<Longrightarrow> is_heap h"
+ unfolding is_heap_except_def is_heap_def parent_def
+ by auto
+
+lemma aux_is:
+ "i < length h \<Longrightarrow> is_heap_except i h \<Longrightarrow> \<not> h ! i \<ge> h ! parent i \<Longrightarrow>
+ is_heap_except (parent i) (swap i (parent i) h)"
+ unfolding is_heap_except_def
+ by (smt add_cancel_right_left diff_is_0_eq div_le_dividend hw11.swap_def le_less_trans length_list_update linorder_not_less nth_list_update order.strict_iff_order parent_def parent_less)
+(* *** *)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (simp add: aux_basis)
+next
+ case (2 h v)
+ then show ?case
+ apply (auto simp: is_heap_def)
+ apply (metis is_heap_except_def)
+ by (smt aux_is hw11.swap_def le_less_trans less_imp_le_nat mset_swap nat.simps(3) parent_less size_mset)
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ by (smt hw11.swap_def linorder_not_less min_def min_less_iff_conj mset_swap nat.simps(3) parent_less size_mset)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x h = sift_up (h@[x]) (length h)"
+
+(* auxiliary lemma *)
+lemma invar_ins_aux:
+ "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ unfolding is_heap_def is_heap_except_def parent_def
+ apply (induction h arbitrary: x)
+ apply (auto)
+ by (smt Suc_mono Suc_pred append_is_Nil_conv butlast.simps(2) butlast_snoc diff_is_0_eq' div_le_dividend le_less_trans length_Cons less_Suc_eq not_less nth_butlast zero_less_Suc)
+(* *** *)
+
+lemma invar_empty: "is_heap emp"
+ unfolding is_heap_def emp_def
+ apply (auto)
+ done
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def
+ apply auto
+ done
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ oops
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ apply (induction h arbitrary: x)
+ apply (simp add: is_heap_def parent_def)
+ apply (metis invar_ins_aux length_append_singleton less_Suc_eq sift_up_restore_heap)
+ done
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ oops
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Putwattana_Attakorn_a.putwattana@tum.de_552/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,273 @@
+Using temporary directory '/tmp/tmp.jJk2hs870q'
+Files in /tmp/eval-552-1wCVE8: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<submission>
+
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+(* auxiliary lemma *)
+lemma aux_basis:
+ "is_heap_except 0 h \<Longrightarrow> 0 < length h \<Longrightarrow> is_heap h"
+ unfolding is_heap_except_def is_heap_def parent_def
+ by auto
+
+lemma aux_is:
+ "i < length h \<Longrightarrow> is_heap_except i h \<Longrightarrow> \<not> h ! i \<ge> h ! parent i \<Longrightarrow>
+ is_heap_except (parent i) (swap i (parent i) h)"
+ unfolding is_heap_except_def
+ by (smt add_cancel_right_left diff_is_0_eq div_le_dividend hw11.swap_def le_less_trans length_list_update linorder_not_less nth_list_update order.strict_iff_order parent_def parent_less)
+(* *** *)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by (simp add: aux_basis)
+next
+ case (2 h v)
+ then show ?case
+ apply (auto simp: is_heap_def)
+ apply (metis is_heap_except_def)
+ by (smt aux_is hw11.swap_def le_less_trans less_imp_le_nat mset_swap nat.simps(3) parent_less size_mset)
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ by (smt hw11.swap_def linorder_not_less min_def min_less_iff_conj mset_swap nat.simps(3) parent_less size_mset)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x h = sift_up (h@[x]) (length h)"
+
+(* auxiliary lemma *)
+lemma invar_ins_aux:
+ "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ unfolding is_heap_def is_heap_except_def parent_def
+ apply (induction h arbitrary: x)
+ apply (auto)
+ by (smt Suc_mono Suc_pred append_is_Nil_conv butlast.simps(2) butlast_snoc diff_is_0_eq' div_le_dividend le_less_trans length_Cons less_Suc_eq not_less nth_butlast zero_less_Suc)
+(* *** *)
+
+lemma invar_empty: "is_heap emp"
+ unfolding is_heap_def emp_def
+ apply (auto)
+ done
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def
+ apply auto
+ done
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ oops
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ apply (induction h arbitrary: x)
+ apply (simp add: is_heap_def parent_def)
+ apply (metis invar_ins_aux length_append_singleton less_Suc_eq sift_up_restore_heap)
+ done
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ oops
+
+(*<*)
+</submission>
+<submission-imports>
+
+</submission-imports>
+<check-imports>
+Submission
+</check-imports>
+<check>
+
+</check>
+<version>
+2016-1
+</version>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<image>
+HOL-Library
+</image>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rau_Martin_martin.rau@tum.de_559/Home11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,112 @@
+(** Score: 6/5
+*)
+theory Home11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+
+(* Homework 11 Heap *)
+
+definition parent :: "nat \<Rightarrow> nat" where
+ "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool" where
+ "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+definition swap :: "nat \<Rightarrow> nat \<Rightarrow> 'a::linorder list \<Rightarrow> 'a list" where
+ "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+lemma is_heap_except_cases_n: "n = length h \<Longrightarrow> is_heap_except j h \<Longrightarrow> is_heap h \<or> h!j < h!parent j"
+ by(induction n) (auto simp add: is_heap_except_def is_heap_def)
+
+lemma is_heap_except_cases: "is_heap_except j h \<Longrightarrow> is_heap h \<or> h!j < h!parent j"
+ using is_heap_except_cases_n by blast
+
+lemma except_j: "\<not>is_heap h \<Longrightarrow> is_heap_except j h \<Longrightarrow> j < length h \<Longrightarrow> h!j < h!parent j"
+ by (induction h) (auto simp add: is_heap_except_def parent_def is_heap_def)
+
+lemma swap_length: "length h = length (swap i j h)"
+ by (induction h) (auto simp add: swap_def split: nat.splits)
+
+lemma swap_0: "length h > 0 \<Longrightarrow> swap 0 (parent 0) h = h"
+ by (induction h) (auto simp add: parent_def swap_def)
+
+lemma swap_parent_j: "j < length h \<Longrightarrow> (swap j (parent j) h)!parent j = h!j"
+ by (induction j) (auto simp add: swap_def parent_def)
+
+lemma swap_is_heap_except:
+ assumes "n = length h" "j < n" "is_heap_except j h" "h!j < h!parent j"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+ using assms
+ apply (induction n)
+ apply (auto simp add: is_heap_except_def swap_def)
+ apply (smt dual_order.strict_trans2 le_less_linear less_asym' list_update_beyond nth_list_update_eq nth_list_update_neq)
+ by (smt dual_order.strict_trans2 le_less_linear length_list_update less_asym' nth_list_update_eq nth_list_update_neq parent_less swap_0 swap_parent_j)
+
+lemma sift_up_restore_heap: "is_heap_except j h \<Longrightarrow> j < length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by(auto simp add: is_heap_def is_heap_except_def parent_def)
+next
+ case (2 h j)
+ then have "is_heap h \<or> h!Suc j < h!parent (Suc j)" using is_heap_except_cases by blast
+ then show ?case by (metis (no_types, lifting) "2.IH" "2.prems"(1) "2.prems"(2) dual_order.strict_trans not_le old.nat.distinct(2) parent_less sift_up.simps(2) swap_is_heap_except swap_length)
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ by (induction h j rule: sift_up.induct) (auto simp add: swap_def parent_def mset_swap)
+
+(* Bonus *)
+
+definition emp :: "'a::linorder list" where
+ "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where
+ "get_min h \<equiv> h!0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "ins x h = sift_up (h @ [x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ by (simp add: is_heap_def emp_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (simp add: emp_def)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ sorry
+
+lemma invar_ins:
+ assumes "is_heap h"
+ shows "is_heap (ins x h)"
+ using assms
+proof -
+ let ?k = "h @ [x]"
+ show ?thesis
+ proof cases
+ assume "?k!parent (length h) \<le> ?k!length h"
+ then have "is_heap ?k" using assms by (smt butlast_snoc dual_order.strict_trans is_heap_def length_append_singleton less_Suc_eq nth_butlast parent_less swap_0 swap_parent_j)
+ then show ?thesis unfolding ins_def by (metis \<open>(h @ [x]) ! parent (length h) \<le> (h @ [x]) ! length h\<close> sift_up.elims)
+ next
+ assume "\<not>(?k!parent (length h) \<le> ?k!length h)"
+ then have "is_heap_except (length h) ?k" using assms by (smt One_nat_def Suc_eq_plus1 diff_is_0_eq div_le_dividend dual_order.strict_trans eq_iff is_heap_def is_heap_except_def length_append_singleton less_Suc_eq nth_append parent_def parent_less)
+ then show ?thesis by (simp add: ins_def sift_up_restore_heap)
+ qed
+qed
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ unfolding ins_def by (simp add: sift_up_mset)
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rau_Martin_martin.rau@tum.de_559/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+martin.rau@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rau_Martin_martin.rau@tum.de_559/testoutput.html Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>No test cases passed! -- Is your submission compiling?</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Failed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rieder_Sabine_sabine.rieder@tum.de_551/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+sabine.rieder@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rieder_Sabine_sabine.rieder@tum.de_551/ex11_tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,364 @@
+(** Score: 8/5
+ nice job
+*)
+(*<*)
+theory ex11_tmpl
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert x Leaf = Node 1 Leaf x Leaf" |
+ "lh_insert x (Node _ l a r) = (if a < x then node l a (lh_insert x r) else node l x (lh_insert a r))"
+
+lemma mset_tree_node[simp]: "mset_tree (node l a r) = mset_tree l + mset_tree r + {#a#}"
+ by(auto simp add: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ apply(induction t arbitrary: x)
+ apply(auto)
+ done
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ apply(induction t arbitrary: x)
+ apply(force simp add: heap_node mset_lh_insert)+
+ done
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ apply(induction t arbitrary: x)
+ apply (auto simp: ltree_node)
+ done
+
+(*count node or Node calls*)
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+ where
+ "t_lh_insert x Leaf = 1" |
+ "t_lh_insert x (Node _ l a r) = (if a < x then 1 + (t_lh_insert x r) else 1 +(t_lh_insert a r))"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ apply(induction t arbitrary: x)
+ apply(auto)
+ done
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'q) bs_pq = Empty | Heap 'a 'q
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty Empty \<longleftrightarrow> True" |
+ "is_empty _ \<longleftrightarrow> False"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert x Empty = Heap x (orig_insert x orig_empty)" |
+ "insert x (Heap y q) = Heap (min x y) (orig_insert x q)"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min Empty = undefined" |
+ "get_min (Heap y q) = y"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min Empty = undefined" |
+ "del_min (Heap x q) = (let q = orig_del_min q in (if orig_is_empty q then
+ Empty else Heap (orig_get_min q) q))"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty \<longleftrightarrow> True" |
+ "invar (Heap a q) \<longleftrightarrow> (orig_invar q) \<and> a \<in># orig_mset q \<and> a = Min_mset (orig_mset q)"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (Heap m q) = orig_mset q"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+lemma aux1: "\<lbrakk>m\<noteq>{#}; x\<le> Min_mset m\<rbrakk> \<Longrightarrow> x = Min (Set.insert x (set_mset m))"
+ by (simp add: min_def)
+
+lemma[simp]: "Min_mset m \<in># m \<longleftrightarrow> m\<noteq>{#} " by auto
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ then show ?case unfolding empty_def by (auto)
+ next
+ case (2 q)
+ then show ?case apply(cases q) by auto
+next
+case (3 q x)
+ then show ?case apply(cases q) by auto
+next
+ case (4 q)
+ then show ?case apply(cases q) by auto
+next
+ case (5 q)
+ then show ?case apply(cases q) by auto
+next
+ case 6
+ then show ?case unfolding empty_def by auto
+next
+ case (7 q x)
+ then show ?case apply(cases q) by (auto simp: min_def)
+next
+case (8 q)
+then show ?case apply(cases q) by auto
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma[simp]:
+ assumes "is_heap_except 0 h" "h \<noteq> []"
+ shows "is_heap h"
+proof -
+ have h1: " (\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i))" using assms is_heap_except_def by auto
+ have h2: "parent 0 = 0" using parent_def by auto
+ have h3: "h!0 \<ge> h!(parent 0)" using assms h2 by auto
+ show ?thesis using h1 h2 h3 is_heap_def by metis
+qed
+
+lemma aux1_1: "\<lbrakk>x \<noteq> a; x \<noteq> b\<rbrakk> \<Longrightarrow> (swap a b h)!x = h!x"
+ by(auto simp add: swap_def)
+
+lemma aux1_2: "a<length h \<Longrightarrow> parent a < length (swap a (parent a) h)"
+ by(auto simp add: swap_def parent_def)
+
+lemma aux1_3: "\<lbrakk>a< length h; b< length h\<rbrakk> \<Longrightarrow> (swap a b h)!a = h!b"
+ apply(auto simp add: swap_def)
+ by (simp add: nth_list_update)
+
+lemma aux1_4: "\<lbrakk>a< length h; b< length h\<rbrakk> \<Longrightarrow> (swap a b h)!b = h!a"
+ by(auto simp add: swap_def)
+
+lemma aux1:
+ assumes "is_heap_except j h" "j < length h" "h ! parent j > h !j"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ from assms have h1: "(\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))" using is_heap_except_def by auto
+ from assms have h2: "(\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))" using is_heap_except_def by auto
+ obtain h' where "h' = (swap j (parent j) h)" by auto
+ from assms h1 h2 have h3: "(\<forall>i<length h'. i\<noteq>j \<and> i\<noteq>parent j \<longrightarrow> h'!i \<ge> h'!(parent i))" by (smt \<open>h' = swap j (parent j) h\<close> aux1_2 dual_order.strict_trans2 ex11_tmpl.swap_def length_list_update less_imp_not_less not_le_imp_less nth_list_update_eq nth_list_update_neq)
+ have "h'! j = h!parent j \<and> h'!parent j = h!j" by (metis \<open>h' = swap j (parent j) h\<close> assms(2) aux1_2 ex11_tmpl.swap_def length_list_update nth_list_update_eq nth_list_update_neq)
+ from this assms h1 h2 have h4: " h'!(parent j) \<le> h'!j" by simp
+ from assms h1 h2 h3 h4 have h6: "(\<forall>i<length h'. i\<noteq>parent j \<longrightarrow> h'!i \<ge> h'!(parent i))" by auto
+ show ?thesis apply(auto simp add: h6 is_heap_except_def)
+ using \<open>h' = swap j (parent j) h\<close> h6 apply blast
+ by (smt assms(3) aux1_1 aux1_2 dual_order.trans ex11_tmpl.swap_def h1 leD le_cases length_list_update nth_list_update)
+qed
+
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction h j rule:sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case apply auto
+ apply (auto simp add: is_heap_def is_heap_except_def)[1] by(auto simp add: aux1_2 aux1)
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case apply(auto)
+ by (metis dual_order.strict_trans ex11_tmpl.swap_def mset_swap nat.simps(3) parent_less size_mset)
+qed
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> (case h of
+ [] \<Rightarrow> undefined |
+ a#h \<Rightarrow> a)"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma aux2_2: "\<forall>i<length h. P(h!i) \<Longrightarrow> \<forall>j \<in> set h. P j"
+proof(induction h)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons a h)
+ obtain i' where "i' = length h" by auto
+ from this Cons have "P((a#h)!i')" by (meson impossible_Cons linorder_not_less)
+ from this have "P a" using Cons.prems by auto
+ then show ?case by (metis Cons.prems in_set_conv_nth)
+qed
+
+fun is_grand_parent where
+ "is_grand_parent a b = (if a<b then is_grand_parent a (parent b) else a=b)"
+
+lemma aux_gp3: "\<lbrakk>a<b; is_grand_parent a b\<rbrakk> \<Longrightarrow> is_grand_parent a (parent b)" by simp
+
+lemma aux_gp:
+ assumes "is_grand_parent a b" "is_heap h" "b<length h"
+ shows " h!a \<le> h!b"
+ using assms proof (induction b rule: less_induct)
+ case (less x)
+ then show ?case using assms proof (cases x)
+ case 0
+ then show ?thesis using less.prems(1) by auto
+ next
+ case (Suc nat)
+ then show ?thesis using assms less by (metis dual_order.strict_trans dual_order.trans is_grand_parent.elims(2) is_heap_def le_cases nat.simps(3) parent_less)
+ qed
+qed
+
+
+lemma aux_gp2: "is_grand_parent 0 b"
+ apply(induction b rule: less_induct)
+ by (metis is_grand_parent.elims(3) neq0_conv parent_less)
+
+lemma aux2_3: "is_heap h \<Longrightarrow> \<forall>i<length h. h!i \<ge> h!0"
+ apply(auto)
+ using aux_gp aux_gp2 by auto
+
+lemma aux2_1:
+ assumes "is_heap (a#h)"
+ shows " \<forall>i \<in> set (a#h). i \<ge> a"
+proof -
+ from assms have "\<forall>i<length (a#h). (a#h)!i \<ge> (a#h)!parent i" by (auto simp add: is_heap_def)
+ from this have "\<forall>i<length (a#h). (a#h)!i \<ge>(a#h)!0" using assms aux2_3 by blast
+ from this have "\<forall>i<length (a#h). (a#h)!i \<ge>a" by auto
+ from this have "\<forall>i\<in> set (a#h). i \<ge> a" using aux2_2 by blast
+ then show ?thesis by auto
+qed
+
+lemma invar_empty: "is_heap emp" by(auto simp add: is_heap_def emp_def)
+
+lemma mset_empty: "mset emp = {#}" by (auto simp add: emp_def)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ apply(auto simp add: get_min_def)
+ apply(cases h)
+ apply(auto)
+ by (simp add: Min_insert2 aux2_1)
+
+lemma aux3_1: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ apply(auto simp add: is_heap_def is_heap_except_def)
+ apply (metis One_nat_def diff_is_0_eq' div_le_dividend dual_order.strict_trans length_Cons less_SucE list.size(3) list.size(4) nth_append parent_def parent_less)
+ by (metis One_nat_def diff_is_0_eq' div_le_dividend eq_iff length_Cons list.size(3) list.size(4) not_less_eq parent_def parent_less)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: aux3_1 ins_def sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: ins_def sift_up_mset)
+
+
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rieder_Sabine_sabine.rieder@tum.de_551/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,505 @@
+Using temporary directory '/tmp/tmp.FasUEq1JI3'
+Files in /tmp/eval-551-s5diBB: ex11_tmpl.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert x Leaf = Node 1 Leaf x Leaf" |
+ "lh_insert x (Node _ l a r) = (if a < x then node l a (lh_insert x r) else node l x (lh_insert a r))"
+
+lemma mset_tree_node[simp]: "mset_tree (node l a r) = mset_tree l + mset_tree r + {#a#}"
+ by(auto simp add: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ apply(induction t arbitrary: x)
+ apply(auto)
+ done
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ apply(induction t arbitrary: x)
+ apply(force simp add: heap_node mset_lh_insert)+
+ done
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ apply(induction t arbitrary: x)
+ apply (auto simp: ltree_node)
+ done
+
+(*count node or Node calls*)
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+ where
+ "t_lh_insert x Leaf = 1" |
+ "t_lh_insert x (Node _ l a r) = (if a < x then 1 + (t_lh_insert x r) else 1 +(t_lh_insert a r))"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ apply(induction t arbitrary: x)
+ apply(auto)
+ done
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'q) bs_pq = Empty | Heap 'a 'q
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty Empty \<longleftrightarrow> True" |
+ "is_empty _ \<longleftrightarrow> False"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert x Empty = Heap x (orig_insert x orig_empty)" |
+ "insert x (Heap y q) = Heap (min x y) (orig_insert x q)"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min Empty = undefined" |
+ "get_min (Heap y q) = y"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min Empty = undefined" |
+ "del_min (Heap x q) = (let q = orig_del_min q in (if orig_is_empty q then
+ Empty else Heap (orig_get_min q) q))"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty \<longleftrightarrow> True" |
+ "invar (Heap a q) \<longleftrightarrow> (orig_invar q) \<and> a \<in># orig_mset q \<and> a = Min_mset (orig_mset q)"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (Heap m q) = orig_mset q"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+lemma aux1: "\<lbrakk>m\<noteq>{#}; x\<le> Min_mset m\<rbrakk> \<Longrightarrow> x = Min (Set.insert x (set_mset m))"
+ by (simp add: min_def)
+
+lemma[simp]: "Min_mset m \<in># m \<longleftrightarrow> m\<noteq>{#} " by auto
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ then show ?case unfolding empty_def by (auto)
+ next
+ case (2 q)
+ then show ?case apply(cases q) by auto
+next
+case (3 q x)
+ then show ?case apply(cases q) by auto
+next
+ case (4 q)
+ then show ?case apply(cases q) by auto
+next
+ case (5 q)
+ then show ?case apply(cases q) by auto
+next
+ case 6
+ then show ?case unfolding empty_def by auto
+next
+ case (7 q x)
+ then show ?case apply(cases q) by (auto simp: min_def)
+next
+case (8 q)
+then show ?case apply(cases q) by auto
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma[simp]:
+ assumes "is_heap_except 0 h" "h \<noteq> []"
+ shows "is_heap h"
+proof -
+ have h1: " (\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i))" using assms is_heap_except_def by auto
+ have h2: "parent 0 = 0" using parent_def by auto
+ have h3: "h!0 \<ge> h!(parent 0)" using assms h2 by auto
+ show ?thesis using h1 h2 h3 is_heap_def by metis
+qed
+
+lemma aux1_1: "\<lbrakk>x \<noteq> a; x \<noteq> b\<rbrakk> \<Longrightarrow> (swap a b h)!x = h!x"
+ by(auto simp add: swap_def)
+
+lemma aux1_2: "a<length h \<Longrightarrow> parent a < length (swap a (parent a) h)"
+ by(auto simp add: swap_def parent_def)
+
+lemma aux1_3: "\<lbrakk>a< length h; b< length h\<rbrakk> \<Longrightarrow> (swap a b h)!a = h!b"
+ apply(auto simp add: swap_def)
+ by (simp add: nth_list_update)
+
+lemma aux1_4: "\<lbrakk>a< length h; b< length h\<rbrakk> \<Longrightarrow> (swap a b h)!b = h!a"
+ by(auto simp add: swap_def)
+
+lemma aux1:
+ assumes "is_heap_except j h" "j < length h" "h ! parent j > h !j"
+ shows "is_heap_except (parent j) (swap j (parent j) h)"
+proof -
+ from assms have h1: "(\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))" using is_heap_except_def by auto
+ from assms have h2: "(\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))" using is_heap_except_def by auto
+ obtain h' where "h' = (swap j (parent j) h)" by auto
+ from assms h1 h2 have h3: "(\<forall>i<length h'. i\<noteq>j \<and> i\<noteq>parent j \<longrightarrow> h'!i \<ge> h'!(parent i))" by (smt \<open>h' = swap j (parent j) h\<close> aux1_2 dual_order.strict_trans2 ex11_tmpl.swap_def length_list_update less_imp_not_less not_le_imp_less nth_list_update_eq nth_list_update_neq)
+ have "h'! j = h!parent j \<and> h'!parent j = h!j" by (metis \<open>h' = swap j (parent j) h\<close> assms(2) aux1_2 ex11_tmpl.swap_def length_list_update nth_list_update_eq nth_list_update_neq)
+ from this assms h1 h2 have h4: " h'!(parent j) \<le> h'!j" by simp
+ from assms h1 h2 h3 h4 have h6: "(\<forall>i<length h'. i\<noteq>parent j \<longrightarrow> h'!i \<ge> h'!(parent i))" by auto
+ show ?thesis apply(auto simp add: h6 is_heap_except_def)
+ using \<open>h' = swap j (parent j) h\<close> h6 apply blast
+ by (smt assms(3) aux1_1 aux1_2 dual_order.trans ex11_tmpl.swap_def h1 leD le_cases length_list_update nth_list_update)
+qed
+
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof(induction h j rule:sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case apply auto
+ apply (auto simp add: is_heap_def is_heap_except_def)[1] by(auto simp add: aux1_2 aux1)
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case by auto
+next
+ case (2 h v)
+ then show ?case apply(auto)
+ by (metis dual_order.strict_trans ex11_tmpl.swap_def mset_swap nat.simps(3) parent_less size_mset)
+qed
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> (case h of
+ [] \<Rightarrow> undefined |
+ a#h \<Rightarrow> a)"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma aux2_2: "\<forall>i<length h. P(h!i) \<Longrightarrow> \<forall>j \<in> set h. P j"
+proof(induction h)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons a h)
+ obtain i' where "i' = length h" by auto
+ from this Cons have "P((a#h)!i')" by (meson impossible_Cons linorder_not_less)
+ from this have "P a" using Cons.prems by auto
+ then show ?case by (metis Cons.prems in_set_conv_nth)
+qed
+
+fun is_grand_parent where
+ "is_grand_parent a b = (if a<b then is_grand_parent a (parent b) else a=b)"
+
+lemma aux_gp3: "\<lbrakk>a<b; is_grand_parent a b\<rbrakk> \<Longrightarrow> is_grand_parent a (parent b)" by simp
+
+lemma aux_gp:
+ assumes "is_grand_parent a b" "is_heap h" "b<length h"
+ shows " h!a \<le> h!b"
+ using assms proof (induction b rule: less_induct)
+ case (less x)
+ then show ?case using assms proof (cases x)
+ case 0
+ then show ?thesis using less.prems(1) by auto
+ next
+ case (Suc nat)
+ then show ?thesis using assms less by (metis dual_order.strict_trans dual_order.trans is_grand_parent.elims(2) is_heap_def le_cases nat.simps(3) parent_less)
+ qed
+qed
+
+
+lemma aux_gp2: "is_grand_parent 0 b"
+ apply(induction b rule: less_induct)
+ by (metis is_grand_parent.elims(3) neq0_conv parent_less)
+
+lemma aux2_3: "is_heap h \<Longrightarrow> \<forall>i<length h. h!i \<ge> h!0"
+ apply(auto)
+ using aux_gp aux_gp2 by auto
+
+lemma aux2_1:
+ assumes "is_heap (a#h)"
+ shows " \<forall>i \<in> set (a#h). i \<ge> a"
+proof -
+ from assms have "\<forall>i<length (a#h). (a#h)!i \<ge> (a#h)!parent i" by (auto simp add: is_heap_def)
+ from this have "\<forall>i<length (a#h). (a#h)!i \<ge>(a#h)!0" using assms aux2_3 by blast
+ from this have "\<forall>i<length (a#h). (a#h)!i \<ge>a" by auto
+ from this have "\<forall>i\<in> set (a#h). i \<ge> a" using aux2_2 by blast
+ then show ?thesis by auto
+qed
+
+lemma invar_empty: "is_heap emp" by(auto simp add: is_heap_def emp_def)
+
+lemma mset_empty: "mset emp = {#}" by (auto simp add: emp_def)
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ apply(auto simp add: get_min_def)
+ apply(cases h)
+ apply(auto)
+ by (simp add: Min_insert2 aux2_1)
+
+lemma aux3_1: "is_heap h \<Longrightarrow> is_heap_except (length h) (h@[x])"
+ apply(auto simp add: is_heap_def is_heap_except_def)
+ apply (metis One_nat_def diff_is_0_eq' div_le_dividend dual_order.strict_trans length_Cons less_SucE list.size(3) list.size(4) nth_append parent_def parent_less)
+ by (metis One_nat_def diff_is_0_eq' div_le_dividend eq_iff length_Cons list.size(3) list.size(4) not_less_eq parent_def parent_less)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ by (simp add: aux3_1 ins_def sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (simp add: ins_def sift_up_mset)
+
+
+(*<*)
+</submission>
+<submission-imports>
+
+</submission-imports>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<version>
+2016-1
+</version>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<image>
+HOL-Library
+</image>
+<check-imports>
+Submission
+</check-imports>
+<check>
+
+</check>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rotar_Alexej_ga59zew@mytum.de_564/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+ga59zew@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rotar_Alexej_ga59zew@mytum.de_564/hw11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,132 @@
+(** Score: 5/5
+*)
+(*<*)
+theory hw11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma [simp]: "is_heap h \<Longrightarrow> n<length h \<Longrightarrow> sift_up h n = h"
+proof (induction n rule: nat_less_induct)
+ case (1 n)
+ then show ?case using is_heap_def
+ by (metis sift_up.elims)
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction j arbitrary: h rule: nat_less_induct)
+ case (1 n)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then show ?thesis using 1 by simp
+ next
+ case False
+ then show ?thesis
+ proof (cases "n=0")
+ case True
+ then have "sift_up h n = h" by simp
+ moreover have "is_heap h" using 1 \<open>n=0\<close>
+ by (auto simp: is_heap_except_def is_heap_def parent_def)
+ ultimately show ?thesis by simp
+ next
+ case False
+ let ?h'= "swap n (parent n) h"
+ from 1 \<open>\<not>is_heap h\<close> have "\<not>h!n \<ge> h!parent n" using 1 is_heap_except_def is_heap_def by metis
+ then have [simp]: "sift_up h n = sift_up ?h' (parent n)"
+ using 1 \<open>n\<noteq>0\<close> by (metis (no_types, lifting) sift_up.elims)
+
+ have [simp]: "length ?h' = length h" using 1 \<open>n\<noteq>0\<close>
+ by (simp add: hw11.swap_def)
+
+ have A: "is_heap_except (parent n) ?h'" using 1 unfolding is_heap_except_def
+ by (smt False \<open>\<not> h ! parent n \<le> h ! n\<close> antisym_conv2 dual_order.strict_trans hw11.swap_def length_list_update linear nth_list_update_eq nth_list_update_neq parent_less)
+
+ have "parent n < n" using 1 \<open>n\<noteq>0\<close> by auto
+ also have "\<dots> < length h" using 1 by simp
+ also have "\<dots> = length ?h'" by simp
+ finally have "parent n < length ?h'" .
+
+ then show ?thesis using 1 A
+ by (simp add: "1.IH" \<open>parent n < n\<close>)
+ qed
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ by (metis hw11.swap_def le_less_trans less_imp_le_nat mset_eq_length mset_swap nat.simps(3) parent_less)
+
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" sorry
+
+lemma mset_empty: "mset emp = {#}" sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Rotar_Alexej_ga59zew@mytum.de_564/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,273 @@
+Using temporary directory '/tmp/tmp.mS57MkNlB4'
+Files in /tmp/eval-564-vjGnqC: hw11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<image>
+HOL-Library
+</image>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<submission-imports>
+
+</submission-imports>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<check-imports>
+Submission
+</check-imports>
+<check>
+
+</check>
+<version>
+2016-1
+</version>
+<submission>
+
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma [simp]: "is_heap h \<Longrightarrow> n<length h \<Longrightarrow> sift_up h n = h"
+proof (induction n rule: nat_less_induct)
+ case (1 n)
+ then show ?case using is_heap_def
+ by (metis sift_up.elims)
+qed
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction j arbitrary: h rule: nat_less_induct)
+ case (1 n)
+ then show ?case
+ proof (cases "is_heap h")
+ case True
+ then show ?thesis using 1 by simp
+ next
+ case False
+ then show ?thesis
+ proof (cases "n=0")
+ case True
+ then have "sift_up h n = h" by simp
+ moreover have "is_heap h" using 1 \<open>n=0\<close>
+ by (auto simp: is_heap_except_def is_heap_def parent_def)
+ ultimately show ?thesis by simp
+ next
+ case False
+ let ?h'= "swap n (parent n) h"
+ from 1 \<open>\<not>is_heap h\<close> have "\<not>h!n \<ge> h!parent n" using 1 is_heap_except_def is_heap_def by metis
+ then have [simp]: "sift_up h n = sift_up ?h' (parent n)"
+ using 1 \<open>n\<noteq>0\<close> by (metis (no_types, lifting) sift_up.elims)
+
+ have [simp]: "length ?h' = length h" using 1 \<open>n\<noteq>0\<close>
+ by (simp add: hw11.swap_def)
+
+ have A: "is_heap_except (parent n) ?h'" using 1 unfolding is_heap_except_def
+ by (smt False \<open>\<not> h ! parent n \<le> h ! n\<close> antisym_conv2 dual_order.strict_trans hw11.swap_def length_list_update linear nth_list_update_eq nth_list_update_neq parent_less)
+
+ have "parent n < n" using 1 \<open>n\<noteq>0\<close> by auto
+ also have "\<dots> < length h" using 1 by simp
+ also have "\<dots> = length ?h'" by simp
+ finally have "parent n < length ?h'" .
+
+ then show ?thesis using 1 A
+ by (simp add: "1.IH" \<open>parent n < n\<close>)
+ qed
+ qed
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ by (metis hw11.swap_def le_less_trans less_imp_le_nat mset_eq_length mset_swap nat.simps(3) parent_less)
+
+
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" good_sorry
+
+lemma mset_empty: "mset emp = {#}" good_sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" good_sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Smith_Nicholas_nick.smith@tum.de_567/HW11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,138 @@
+(** Score: 4/5
+*)
+theory HW11
+ imports Main
+ "HOL-Data_Structures.Leftist_Heap"
+
+begin
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma aux_base_heap: "is_heap_except 0 h \<Longrightarrow> is_heap(h)"
+ apply(induction h)
+ apply( auto simp add:is_heap_def parent_def)
+ by (metis Divides.div_less Groups.add_ac(2) One_nat_def Suc_eq_plus1_left Suc_less_eq diff_is_0_eq' is_heap_except_def le_less less_Suc0 list.size(4) numeral_2_eq_2 parent_def)
+
+
+lemma aux_swap_size: "length (swap i j h) = length(h)"
+ apply(induction h arbitrary: i j)
+ apply(auto simp add: swap_def)
+ by (simp add: Nitpick.case_nat_unfold)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule:sift_up.induct )
+ case (1 h)
+ then show ?case by (auto simp add: aux_base_heap)
+next
+ case (2 h v) show ?case
+ proof( cases "h ! parent (Suc v) \<le> h ! Suc v ")
+ case True
+ then show ?thesis using "2.prems" apply(auto) by (metis is_heap_def is_heap_except_def)
+ next
+ case False
+ then show ?thesis using "2.prems" apply(auto) apply (rule "2.IH") apply(simp_all)
+ (*prove the assumptions 2nd statement *) sorry
+ qed
+ qed
+
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply(induction h j rule:sift_up.induct)
+ apply(auto)
+ by (metis HW11.swap_def Zero_not_Suc dual_order.strict_trans mset_eq_length mset_swap parent_less)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where
+"emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a"
+ where "get_min (h) \<equiv> h!0"
+
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ by (auto simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (auto simp add: emp_def is_heap_def)
+
+
+lemma aux_get_min: "is_heap h \<Longrightarrow> \<forall>i\<in> set(h). get_min h \<le> i"
+ apply (induction h)
+ apply (auto simp add: is_heap_def get_min_def aux_base_heap )
+ sorry
+
+(* relies on auxillarly that I can't prove *)
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ apply (induction h)
+ apply (auto simp add: eq_Min_iff get_min_def )
+ by (metis aux_get_min get_min_def nth_Cons_0 set_subset_Cons subsetCE)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ apply (induction h arbitrary:x)
+ apply(auto simp add:is_heap_def ins_def parent_def split:if_split)
+ sorry
+
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ apply (induction h arbitrary:x)
+ apply(auto simp add:ins_def)
+ by (metis (no_types, lifting) HW11.swap_def Zero_not_Suc add_mset_commute aux_swap_size dual_order.strict_trans length_Cons length_append_singleton lessI mset.simps(2) mset_swap parent_less sift_up_mset)
+
+
+(*<*)
+end
+(*>*)
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Smith_Nicholas_nick.smith@tum.de_567/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+nick.smith@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Smith_Nicholas_nick.smith@tum.de_567/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,278 @@
+Using temporary directory '/tmp/tmp.eOQzj51sA9'
+Files in /tmp/eval-567-aDLHm9: HW11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<check>
+
+</check>
+<image>
+HOL-Library
+</image>
+<submission-imports>
+
+</submission-imports>
+<submission>
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma aux_base_heap: "is_heap_except 0 h \<Longrightarrow> is_heap(h)"
+ apply(induction h)
+ apply( auto simp add:is_heap_def parent_def)
+ by (metis Divides.div_less Groups.add_ac(2) One_nat_def Suc_eq_plus1_left Suc_less_eq diff_is_0_eq' is_heap_except_def le_less less_Suc0 list.size(4) numeral_2_eq_2 parent_def)
+
+
+lemma aux_swap_size: "length (swap i j h) = length(h)"
+ apply(induction h arbitrary: i j)
+ apply(auto simp add: swap_def)
+ by (simp add: Nitpick.case_nat_unfold)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule:sift_up.induct )
+ case (1 h)
+ then show ?case by (auto simp add: aux_base_heap)
+next
+ case (2 h v) show ?case
+ proof( cases "h ! parent (Suc v) \<le> h ! Suc v ")
+ case True
+ then show ?thesis using "2.prems" apply(auto) by (metis is_heap_def is_heap_except_def)
+ next
+ case False
+ then show ?thesis using "2.prems" apply(auto) apply (rule "2.IH") apply(simp_all)
+ (*prove the assumptions 2nd statement *) good_sorry
+ qed
+ qed
+
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply(induction h j rule:sift_up.induct)
+ apply(auto)
+ by (metis HW11.swap_def Zero_not_Suc dual_order.strict_trans mset_eq_length mset_swap parent_less)
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where
+"emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a"
+ where "get_min (h) \<equiv> h!0"
+
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ where "ins x h = sift_up (h@[x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ by (auto simp add: emp_def is_heap_def)
+
+lemma mset_empty: "mset emp = {#}"
+ by (auto simp add: emp_def is_heap_def)
+
+
+lemma aux_get_min: "is_heap h \<Longrightarrow> \<forall>i\<in> set(h). get_min h \<le> i"
+ apply (induction h)
+ apply (auto simp add: is_heap_def get_min_def aux_base_heap )
+ good_sorry
+
+(* relies on auxillarly that I can't prove *)
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ apply (induction h)
+ apply (auto simp add: eq_Min_iff get_min_def )
+ by (metis aux_get_min get_min_def nth_Cons_0 set_subset_Cons subsetCE)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ apply (induction h arbitrary:x)
+ apply(auto simp add:is_heap_def ins_def parent_def split:if_split)
+ good_sorry
+
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ apply (induction h arbitrary:x)
+ apply(auto simp add:ins_def)
+ by (metis (no_types, lifting) HW11.swap_def Zero_not_Suc add_mset_commute aux_swap_size dual_order.strict_trans length_Cons length_append_singleton lessI mset.simps(2) mset_swap parent_less sift_up_mset)
+
+
+(*<*)
+</submission>
+<version>
+2016-1
+</version>
+<check-imports>
+Submission
+</check-imports>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Stamer_Florian_florian.stamer@tum.de_565/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+florian.stamer@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Stamer_Florian_florian.stamer@tum.de_565/ex11.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,312 @@
+(** Score: 6/5
+*)
+(*<*)
+theory ex11
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert x \<langle>\<rangle> = \<langle>1, \<langle>\<rangle>, x, \<langle>\<rangle>\<rangle>" |
+ "lh_insert x \<langle>_, l, a, r\<rangle> = (if a<x then node l a (lh_insert x r) else node l x (lh_insert a r))"
+
+lemma mset_tree_node[simp]:
+ "mset_tree (node l a r) = mset_tree l + {#a#} + mset_tree r"
+ by (auto simp: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x) auto
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ by (induction t arbitrary: x) (force simp: heap_node mset_lh_insert)+
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ by (induction t arbitrary: x) (auto simp: ltree_node)
+
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+where
+ "t_lh_insert x \<langle>\<rangle> = 1" |
+ "t_lh_insert x \<langle>_, l, a, r\<rangle> = (if a<x then 1+ (t_lh_insert x r) else 1+ (t_lh_insert a r))"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x) auto
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | Heap 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty Empty \<longleftrightarrow> True" |
+ "is_empty _ \<longleftrightarrow> False"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert x Empty = Heap x (orig_insert x orig_empty)" |
+ "insert x (Heap a s) = Heap (min x a) (orig_insert x s)"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min Empty = undefined" |
+ "get_min (Heap a s) = a"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min Empty = undefined" |
+ "del_min (Heap a s) = (let
+ s = orig_del_min s
+ in
+ if orig_is_empty s then Empty
+ else Heap (orig_get_min s) s)"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty \<longleftrightarrow> True" |
+ "invar (Heap a s) \<longleftrightarrow> orig_invar s \<and> a\<in>#orig_mset s \<and> a=Min_mset (orig_mset s)"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (Heap _ s) = orig_mset s"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ then show ?case
+ unfolding empty_def by simp
+ next
+ case (2 q) -- \<open>and so on\<close>
+ then show ?case
+ by (cases q) auto
+ next
+ case (3 q x)
+ then show ?case
+ by (cases q) auto
+ next
+ case (4 q)
+ then show ?case
+ by (cases q) auto
+ next
+ case (5 q)
+ then show ?case
+ by (cases q) auto
+ next
+ case 6
+ then show ?case
+ unfolding empty_def by auto
+ next
+ case (7 q x)
+ then show ?case
+ apply (cases q)
+ apply (auto simp: min_def)
+ apply (rule Min_insert2[symmetric])
+ apply (simp)
+ using Min.boundedE apply auto[1]
+ by (metis Min_insert ball_empty finite_set_mset min_def)
+ next
+ case (8 q)
+ then show ?case
+ by (cases q) auto
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma parent0: "parent 0 = 0" by (auto simp: parent_def)
+
+lemma is_heap_except0: "is_heap_except 0 h \<Longrightarrow> is_heap h"
+proof (induction h)
+ case Nil
+ then show ?case by (auto simp: is_heap_def)
+next
+ case (Cons a h)
+ then have "\<forall>i<Suc(length h). i\<noteq>0 \<longrightarrow> (a#h)!parent i \<le> (a#h)!i"
+ unfolding is_heap_except_def by auto
+ then show ?case
+ unfolding is_heap_def
+ by (metis Cons.prems is_heap_except_def parent0)
+qed
+
+lemma length_swap: "length (swap v (parent v) h) = length h"
+ unfolding swap_def by auto
+
+lemma aux1: "v < length h \<Longrightarrow> \<not> h ! parent v \<le> h ! v \<Longrightarrow> parent v < length h"
+ by (metis dual_order.strict_trans parent0 parent_less)
+
+lemma aux2: "\<lbrakk>is_heap_except v h; v < length h; \<not> h ! parent v \<le> h ! v\<rbrakk>
+ \<Longrightarrow> is_heap_except (parent v) (swap v (parent v) h)"
+ unfolding swap_def
+ by (smt aux1 dual_order.trans is_heap_except_def le_cases length_list_update nth_list_update)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ by (auto simp: is_heap_except0)
+next
+ case (2 h v)
+ then show ?case
+ apply (auto simp: length_swap)
+ apply (metis is_heap_def is_heap_except_def)
+ by (auto simp: aux1 aux2)
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ unfolding swap_def
+ apply (auto simp: mset_swap)
+ using dual_order.strict_trans parent_less by blast
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ where "ins x h = sift_up (h @ [x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding emp_def is_heap_def by simp
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by simp
+
+lemma transitive: "is_heap h \<Longrightarrow> h ! parent i \<le> h ! i \<Longrightarrow> h ! (parent (parent i)) \<le> h ! parent i
+ \<Longrightarrow> h ! (parent (parent i)) \<le> h ! i" by simp
+
+lemma min0: "is_heap h \<Longrightarrow> \<forall>x\<in>set h. h ! 0 \<le> x"
+ unfolding is_heap_def
+ apply (induction h)
+ apply (auto simp: )
+ sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ unfolding get_min_def
+ apply (auto simp: eq_Min_iff min0)
+ sorry
+
+lemma aux4: "is_heap [] \<Longrightarrow> is_heap [x]"
+ unfolding is_heap_def by (auto simp: parent0)
+
+lemma length_sift_up: "length (sift_up h j) = length h"
+ by (induction h j rule: sift_up.induct) (auto simp: length_swap)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ apply (induction "h@[x]" "length h" rule: sift_up.induct)
+ apply (auto simp: aux4)
+ by (smt dual_order.strict_trans eq_iff is_heap_def is_heap_except_def length_append_singleton lessI less_SucE not_less_eq nth_append parent0 parent_less sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (auto simp: ins_def sift_up_mset)
+
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Stamer_Florian_florian.stamer@tum.de_565/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,454 @@
+Using temporary directory '/tmp/tmp.07VpFUIEVt'
+Files in /tmp/eval-565-nJzuPl: ex11.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<check>
+
+</check>
+<version>
+2016-1
+</version>
+<image>
+HOL-Library
+</image>
+<defs-imports>
+Complex_Main
+</defs-imports>
+<submission>
+
+(*>*)
+text {* \ExerciseSheet{11}{22.~6.~2018} *}
+
+text \<open>
+ \Exercise{Insert for Leftist Heap}
+
+ \<^item> Define a function to directly insert an element into a leftist heap.
+ Do not construct an intermediate heap like insert via merge does!
+
+ \<^item> Show that your function is correct
+
+ \<^item> Define a timing function for your insert function, and show that it is
+ linearly bounded by the rank of the tree.
+\<close>
+
+fun lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"
+ where
+ "lh_insert x \<langle>\<rangle> = \<langle>1, \<langle>\<rangle>, x, \<langle>\<rangle>\<rangle>" |
+ "lh_insert x \<langle>_, l, a, r\<rangle> = (if a<x then node l a (lh_insert x r) else node l x (lh_insert a r))"
+
+lemma mset_tree_node[simp]:
+ "mset_tree (node l a r) = mset_tree l + {#a#} + mset_tree r"
+ by (auto simp: node_def)
+
+lemma mset_lh_insert: "mset_tree (lh_insert x t) = mset_tree t + {# x #}"
+ by (induction t arbitrary: x) auto
+
+lemma "heap t \<Longrightarrow> heap (lh_insert x t)"
+ by (induction t arbitrary: x) (force simp: heap_node mset_lh_insert)+
+
+lemma "ltree t \<Longrightarrow> ltree (lh_insert x t)"
+ by (induction t arbitrary: x) (auto simp: ltree_node)
+
+
+fun t_lh_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat"
+where
+ "t_lh_insert x \<langle>\<rangle> = 1" |
+ "t_lh_insert x \<langle>_, l, a, r\<rangle> = (if a<x then 1+ (t_lh_insert x r) else 1+ (t_lh_insert a r))"
+
+lemma "t_lh_insert x t \<le> rank t + 1"
+ by (induction t arbitrary: x) auto
+
+text \<open>
+ \Exercise{Bootstrapping a Priority Queue}
+
+ Given a generic priority queue implementation with
+ \<open>O(1)\<close> \<open>empty\<close>, \<open>is_empty\<close> operations, \<open>O(f\<^sub>1 n)\<close> insert,
+ and \<open>O(f\<^sub>2 n)\<close> \<open>get_min\<close> and \<open>del_min\<close> operations.
+
+ Derive an implementation with \<open>O(1)\<close> \<open>get_min\<close>, and the
+ asymptotic complexities of the other operations unchanged!
+
+ Hint: Store the current minimal element! As you know nothing
+ about \<open>f\<^sub>1\<close> and \<open>f\<^sub>2\<close>, you must not use \<open>get_min/del_min\<close>
+ in your new \<open>insert\<close> operation, and vice versa!
+\<close>
+
+text \<open>For technical reasons, you have to define the new implementations type
+ outside the locale!\<close>
+datatype ('a,'s) bs_pq = Empty | Heap 'a 's
+
+locale Bs_Priority_Queue =
+ orig: Priority_Queue where
+ empty = orig_empty and
+ is_empty = orig_is_empty and
+ insert = orig_insert and
+ get_min = orig_get_min and
+ del_min = orig_del_min and
+ invar = orig_invar and
+ mset = orig_mset
+ for orig_empty orig_is_empty orig_insert orig_get_min orig_del_min orig_invar
+ and orig_mset :: "'s \<Rightarrow> 'a::linorder multiset"
+begin
+ text \<open>In here, the original implementation is available with the prefix \<open>orig\<close>, e.g. \<close>
+ term orig_empty term orig_invar
+ thm orig.invar_empty
+
+ definition empty :: "('a,'s) bs_pq"
+ where "empty = Empty"
+
+ fun is_empty :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "is_empty Empty \<longleftrightarrow> True" |
+ "is_empty _ \<longleftrightarrow> False"
+
+ fun insert :: "'a \<Rightarrow> ('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "insert x Empty = Heap x (orig_insert x orig_empty)" |
+ "insert x (Heap a s) = Heap (min x a) (orig_insert x s)"
+
+ fun get_min :: "('a,'s) bs_pq \<Rightarrow> 'a"
+ where
+ "get_min Empty = undefined" |
+ "get_min (Heap a s) = a"
+
+ fun del_min :: "('a,'s) bs_pq \<Rightarrow> ('a,'s) bs_pq"
+ where
+ "del_min Empty = undefined" |
+ "del_min (Heap a s) = (let
+ s = orig_del_min s
+ in
+ if orig_is_empty s then Empty
+ else Heap (orig_get_min s) s)"
+
+ fun invar :: "('a,'s) bs_pq \<Rightarrow> bool"
+ where
+ "invar Empty \<longleftrightarrow> True" |
+ "invar (Heap a s) \<longleftrightarrow> orig_invar s \<and> a\<in>#orig_mset s \<and> a=Min_mset (orig_mset s)"
+
+ fun mset :: "('a,'s) bs_pq \<Rightarrow> 'a multiset"
+ where
+ "mset Empty = {#}" |
+ "mset (Heap _ s) = orig_mset s"
+
+ lemmas [simp] = orig.is_empty orig.mset_get_min orig.mset_del_min
+ orig.mset_insert orig.mset_empty
+ orig.invar_empty orig.invar_insert orig.invar_del_min
+
+ text \<open>Show that your new implementation satisfies the priority queue interface!\<close>
+ sublocale Priority_Queue empty is_empty insert get_min del_min invar mset
+ apply unfold_locales
+ proof goal_cases
+ case 1
+ then show ?case
+ unfolding empty_def by simp
+ next
+ case (2 q) -- \<open>and so on\<close>
+ then show ?case
+ by (cases q) auto
+ next
+ case (3 q x)
+ then show ?case
+ by (cases q) auto
+ next
+ case (4 q)
+ then show ?case
+ by (cases q) auto
+ next
+ case (5 q)
+ then show ?case
+ by (cases q) auto
+ next
+ case 6
+ then show ?case
+ unfolding empty_def by auto
+ next
+ case (7 q x)
+ then show ?case
+ apply (cases q)
+ apply (auto simp: min_def)
+ apply (rule Min_insert2[symmetric])
+ apply (simp)
+ using Min.boundedE apply auto[1]
+ by (metis Min_insert ball_empty finite_set_mset min_def)
+ next
+ case (8 q)
+ then show ?case
+ by (cases q) auto
+qed
+
+
+end
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma parent0: "parent 0 = 0" by (auto simp: parent_def)
+
+lemma is_heap_except0: "is_heap_except 0 h \<Longrightarrow> is_heap h"
+proof (induction h)
+ case Nil
+ then show ?case by (auto simp: is_heap_def)
+next
+ case (Cons a h)
+ then have "\<forall>i<Suc(length h). i\<noteq>0 \<longrightarrow> (a#h)!parent i \<le> (a#h)!i"
+ unfolding is_heap_except_def by auto
+ then show ?case
+ unfolding is_heap_def
+ by (metis Cons.prems is_heap_except_def parent0)
+qed
+
+lemma length_swap: "length (swap v (parent v) h) = length h"
+ unfolding swap_def by auto
+
+lemma aux1: "v < length h \<Longrightarrow> \<not> h ! parent v \<le> h ! v \<Longrightarrow> parent v < length h"
+ by (metis dual_order.strict_trans parent0 parent_less)
+
+lemma aux2: "\<lbrakk>is_heap_except v h; v < length h; \<not> h ! parent v \<le> h ! v\<rbrakk>
+ \<Longrightarrow> is_heap_except (parent v) (swap v (parent v) h)"
+ unfolding swap_def
+ by (smt aux1 dual_order.trans is_heap_except_def le_cases length_list_update nth_list_update)
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+proof (induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ by (auto simp: is_heap_except0)
+next
+ case (2 h v)
+ then show ?case
+ apply (auto simp: length_swap)
+ apply (metis is_heap_def is_heap_except_def)
+ by (auto simp: aux1 aux2)
+qed
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply (induction h j rule: sift_up.induct)
+ apply (auto)
+ unfolding swap_def
+ apply (auto simp: mset_swap)
+ using dual_order.strict_trans parent_less by blast
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> []"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> h ! 0"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ where "ins x h = sift_up (h @ [x]) (length h)"
+
+lemma invar_empty: "is_heap emp"
+ unfolding emp_def is_heap_def by simp
+
+lemma mset_empty: "mset emp = {#}"
+ unfolding emp_def by simp
+
+lemma transitive: "is_heap h \<Longrightarrow> h ! parent i \<le> h ! i \<Longrightarrow> h ! (parent (parent i)) \<le> h ! parent i
+ \<Longrightarrow> h ! (parent (parent i)) \<le> h ! i" by simp
+
+lemma min0: "is_heap h \<Longrightarrow> \<forall>x\<in>set h. h ! 0 \<le> x"
+ unfolding is_heap_def
+ apply (induction h)
+ apply (auto simp: )
+ good_sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)"
+ unfolding get_min_def
+ apply (auto simp: eq_Min_iff min0)
+ good_sorry
+
+lemma aux4: "is_heap [] \<Longrightarrow> is_heap [x]"
+ unfolding is_heap_def by (auto simp: parent0)
+
+lemma length_sift_up: "length (sift_up h j) = length h"
+ by (induction h j rule: sift_up.induct) (auto simp: length_swap)
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)"
+ unfolding ins_def
+ apply (induction "h@[x]" "length h" rule: sift_up.induct)
+ apply (auto simp: aux4)
+ by (smt dual_order.strict_trans eq_iff is_heap_def is_heap_except_def length_append_singleton lessI less_SucE not_less_eq nth_append parent0 parent_less sift_up_restore_heap)
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}"
+ by (auto simp: ins_def sift_up_mset)
+
+
+(*<*)
+</submission>
+<submission-imports>
+
+</submission-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<check-imports>
+Submission
+</check-imports>
+=================
+Failed to load submission in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Wielander_Felix_felix.wielander@tum.de_556/SENTMAIL Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,1 @@
+felix.wielander@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Wielander_Felix_felix.wielander@tum.de_556/ex11_tmpl.thy Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,128 @@
+(** Score: 1/5
+*)
+(*<*)
+theory ex11_tmpl
+imports
+ "HOL-Data_Structures.Leftist_Heap"
+begin
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma double_swap: "\<lbrakk>is_heap h; i < length h\<rbrakk> \<Longrightarrow> swap i (parent i) (swap i (parent i) h) = h"
+ apply(induction h)
+ apply auto
+ sorry
+
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ apply auto sorry
+ proof -
+ have "is_heap_except 0 h" by (simp add: "1.prems"(1))
+ then have "(\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)) \<and> (\<forall>i<length h. parent i = 0 \<longrightarrow> h!i \<ge> h!(parent 0))" using is_heap_except_def by blast
+ also have "\<forall>i<length h. parent i = 0 \<longrightarrow> h!i \<ge> h!0" using calculation by auto
+ then have "h \<noteq> []" using "1.prems"(2) by blast
+ (** Das müssen sie beweisen \<dots> zur Not kleinschrittig mit Isar \<dots>
+ Wenn der Defekt an der Wurzel liegt, ist es natürlich auch ein heap
+ *)
+ qed
+ next
+ case (2 h v)
+ show ?case
+ proof (cases "h ! parent (Suc v) \<le> h ! Suc v")
+ case True with \<open>is_heap_except (Suc v) h\<close> show ?thesis
+ using is_heap_def is_heap_except_def by force
+ next
+ case False
+
+ have "is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ (* Das Kernargument: Durch Vertauschen wird der Defekt eins nach oben geschoben *)
+ sorry
+ moreover have "parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ (* Wenn Suc v in range ist, dann natürlich auch parent (Suc v), selbst wenn man im heap was rumtauscht,
+ da swap die Länge natürlich nicht ändert. *)
+ sorry
+ ultimately show ?thesis using False
+ by (simp add: "2.IH")
+
+ qed
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply(induction h j rule: sift_up.induct)
+ apply(auto)
+ sorry
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" sorry
+
+lemma mset_empty: "mset emp = {#}" sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" sorry
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/SS18/Exercises/hwsubm/hw11/Wielander_Felix_felix.wielander@tum.de_556/user_error_log.txt Mon Jul 09 11:57:12 2018 +0200
@@ -0,0 +1,269 @@
+Using temporary directory '/tmp/tmp.0r3UdiZ6oq'
+Files in /tmp/eval-556-YOxNc9: ex11_tmpl.thy
+Command ['java', '-cp', '/jail/tumjudge-assembly.jar', 'TUMJudge', '/home/test-runner/local-storage']
+=================
+<defs-imports>
+Complex_Main
+</defs-imports>
+<defs>
+
+
+lemma power2sum: "(\<Sum>i\<in>{0..<k}. (2::nat)^i) = 2^k-1"
+ by (induction k) auto
+
+subsection \<open>Strictly Ascending and Descending Sequences\<close>
+definition strictly_ascending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_ascending xs \<equiv> sorted xs \<and> distinct xs"
+
+definition strictly_descending :: "'a::linorder list \<Rightarrow> bool"
+ where
+ "strictly_descending xs \<equiv> sorted (rev xs) \<and> distinct xs"
+
+text \<open>The sequence \<open>[r-1,\<dots>,0]\<close>: \<close>
+definition compl_descending :: "nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "compl_descending r ns \<longleftrightarrow> (ns = rev [0..<r])"
+
+lemma strictly_descending_alt:
+ "strictly_descending = strictly_ascending o rev"
+ unfolding strictly_descending_def strictly_ascending_def by auto
+
+lemma strictly_ascending_Nil[simp, intro!]: "strictly_ascending []"
+ unfolding strictly_ascending_def by auto
+
+lemma strictly_ascending_Cons[simp]:
+ "strictly_ascending (x#xs) \<longleftrightarrow> strictly_ascending xs \<and> (\<forall>y\<in>set xs. x<y)"
+ unfolding strictly_ascending_def using le_neq_trans
+ by (auto simp: sorted_Cons)
+
+lemma strictly_ascending_append[simp]:
+ "strictly_ascending (xs@ys)
+ \<longleftrightarrow> strictly_ascending xs \<and> strictly_ascending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x<y)"
+ by (induction xs) auto
+
+lemma strictly_descending_Nil[simp, intro!]: "strictly_descending []"
+ unfolding strictly_descending_def by auto
+
+lemma strictly_descending_Cons[simp]:
+ "strictly_descending (x#xs) \<longleftrightarrow> strictly_descending xs \<and> (\<forall>y\<in>set xs. x>y)"
+ unfolding strictly_descending_alt
+ by (auto)
+
+lemma strictly_descending_append[simp]:
+ "strictly_descending (xs@ys)
+ \<longleftrightarrow> strictly_descending xs \<and> strictly_descending ys
+ \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. x>y)"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_ascending_rev[simp]:
+ "strictly_ascending (rev xs) \<longleftrightarrow> strictly_descending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma strictly_descending_rev[simp]:
+ "strictly_descending (rev xs) \<longleftrightarrow> strictly_ascending xs"
+ unfolding strictly_descending_alt by auto
+
+lemma compl_descending_Nil[simp]:
+ "compl_descending 0 []"
+ "compl_descending r [] \<longleftrightarrow> r = 0"
+ "compl_descending 0 ns \<longleftrightarrow> ns=[]"
+ unfolding compl_descending_def by auto
+
+lemma compl_descending_Cons[simp]:
+ "compl_descending r (n#ns) \<longleftrightarrow> (r = Suc n \<and> compl_descending n ns)"
+ unfolding compl_descending_def
+ by (cases r) (auto)
+
+lemma compl_is_strictly_descending:
+ "compl_descending r ns \<Longrightarrow> strictly_descending ns"
+ unfolding compl_descending_def strictly_descending_def
+ by auto
+
+subsubsection \<open>Strictly Ascending Sequences of Natural Numbers\<close>
+
+text \<open>Each element is greater or equal to its index\<close>
+lemma strictly_ascending_idx:
+ assumes "strictly_ascending ns"
+ shows "\<forall>i<length ns. i\<le>ns!i"
+ using assms
+ apply (induction ns rule: rev_induct)
+ subgoal by simp
+ subgoal
+ apply (clarsimp simp: nth_append)
+ by (metis less_antisym not_less nth_mem)
+ done
+
+text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
+ can be estimated by summation over \<open>{0..<n}\<close>.
+\<close>
+lemma strictly_ascending_sum_mono_lowerbound:
+ fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
+ assumes "strictly_ascending ns"
+ assumes MONO: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
+ shows "(\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
+ using assms
+proof (induction ns rule: rev_induct)
+ case Nil
+ then show ?case by simp
+next
+ case (snoc n ns)
+ hence IH: "sum f {0..<length ns} \<le> sum_list (map f ns)"
+ by auto
+
+ from strictly_ascending_idx[OF snoc.prems(1), THEN spec[of _ "length ns"]]
+ have NGE: "length ns \<le> n" by auto
+
+ have "sum f {0..<length (ns @ [n])}
+ = sum f {0..<length ns} + f (length ns)"
+ by simp
+ also note IH
+ also note NGE
+ finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
+ using MONO add_mono by blast
+ thus ?case by simp
+qed
+
+save_test_theory
+</defs>
+<submission>
+
+(*>*)
+
+text\<open>
+\Homework{Heap}{29.~6.~2018}
+
+A binary tree can be encoded as an array \<open>[a\<^sub>1,...,a\<^sub>n]\<close>, such that
+the parent of node \<open>a\<^sub>i\<close> is node \<open>a\<^sub>(\<^sub>i \<^sub>d\<^sub>i\<^sub>v \<^sub>2\<^sub>)\<close>.
+
+Thus, for a heap, each node is greater than or equal to its parent:
+\<close>
+
+definition parent :: "nat \<Rightarrow> nat" where "parent i \<equiv> (i+1) div 2 - 1"
+definition is_heap :: "'a::linorder list \<Rightarrow> bool"
+ where "is_heap h \<equiv> \<forall>i<length h. h!i \<ge> h!parent i"
+
+
+text \<open>A heap with a single defect at index \<open>j\<close> is characterized as follows:
+ The heap property holds for all elements except \<open>j\<close>,
+ and the children of \<open>j\<close> must also be greater than their grand-parent.
+\<close>
+definition is_heap_except :: "nat \<Rightarrow> 'a::linorder list \<Rightarrow> bool" where
+ "is_heap_except j h \<equiv>
+ (\<forall>i<length h. i\<noteq>j \<longrightarrow> h!i \<ge> h!(parent i))
+ \<and> (\<forall>i<length h. parent i = j \<longrightarrow> h!i \<ge> h!(parent j))"
+
+
+
+text \<open>
+ The function \<open>sift_up\<close> corrects a single defect in a heap by
+ iterated swapping of the defect with its parent, until the heap property
+ is restored.
+\<close>
+
+definition "swap i j h \<equiv> h[i:=h!j, j:=h!i]"
+
+(* Required for termination proof of sift_up *)
+lemma parent_less[simp]: "i\<noteq>0 \<Longrightarrow> parent i < i"
+ by (auto simp: parent_def)
+
+fun sift_up where
+ "sift_up h 0 = h"
+| "sift_up h i = (
+ if h!i \<ge> h!parent i then h
+ else sift_up (swap i (parent i) h) (parent i))"
+
+
+text \<open>Show that @{const sift_up} restores the heap
+ and preserves the multiset of elements in the heap\<close>
+
+lemma double_swap: "\<lbrakk>is_heap h; i < length h\<rbrakk> \<Longrightarrow> swap i (parent i) (swap i (parent i) h) = h"
+ apply(induction h)
+ apply auto
+ good_sorry
+
+
+lemma sift_up_restore_heap:
+ "is_heap_except j h \<Longrightarrow> j<length h \<Longrightarrow> is_heap (sift_up h j)"
+ proof(induction h j rule: sift_up.induct)
+ case (1 h)
+ then show ?case
+ apply auto good_sorry
+ proof -
+ have "is_heap_except 0 h" by (simp add: "1.prems"(1))
+ then have "(\<forall>i<length h. i\<noteq>0 \<longrightarrow> h!i \<ge> h!(parent i)) \<and> (\<forall>i<length h. parent i = 0 \<longrightarrow> h!i \<ge> h!(parent 0))" using is_heap_except_def by blast
+ also have "\<forall>i<length h. parent i = 0 \<longrightarrow> h!i \<ge> h!0" using calculation by auto
+ then have "h \<noteq> []" using "1.prems"(2) by blast
+ (** Das müssen sie beweisen \<dots> zur Not kleinschrittig mit Isar \<dots>
+ Wenn der Defekt an der Wurzel liegt, ist es natürlich auch ein heap
+ *)
+ qed
+ next
+ case (2 h v)
+ show ?case
+ proof (cases "h ! parent (Suc v) \<le> h ! Suc v")
+ case True with \<open>is_heap_except (Suc v) h\<close> show ?thesis
+ using is_heap_def is_heap_except_def by force
+ next
+ case False
+
+ have "is_heap_except (parent (Suc v)) (swap (Suc v) (parent (Suc v)) h)"
+ (* Das Kernargument: Durch Vertauschen wird der Defekt eins nach oben geschoben *)
+ good_sorry
+ moreover have "parent (Suc v) < length (swap (Suc v) (parent (Suc v)) h)"
+ (* Wenn Suc v in range ist, dann natürlich auch parent (Suc v), selbst wenn man im heap was rumtauscht,
+ da swap die Länge natürlich nicht ändert. *)
+ good_sorry
+ ultimately show ?thesis using False
+ by (simp add: "2.IH")
+
+ qed
+qed
+
+
+lemma sift_up_mset: "j<length h \<Longrightarrow> mset (sift_up h j) = mset h"
+ apply(induction h j rule: sift_up.induct)
+ apply(auto)
+ good_sorry
+
+text \<open>For \textbf{3 bonus points}, add an empty, insert, and get-min
+ function to the heap implementation, and prove their essential properties.
+\<close>
+
+definition emp :: "'a::linorder list" where "emp \<equiv> undefined"
+definition get_min :: "'a::linorder list \<Rightarrow> 'a" where "get_min h \<equiv> undefined"
+definition ins :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where "ins x h = undefined"
+
+lemma invar_empty: "is_heap emp" good_sorry
+
+lemma mset_empty: "mset emp = {#}" good_sorry
+
+lemma mset_get_min: "is_heap h \<Longrightarrow> mset h \<noteq> {#} \<Longrightarrow> get_min h = Min_mset (mset h)" good_sorry
+
+lemma invar_ins: "is_heap h \<Longrightarrow> is_heap (ins x h)" good_sorry
+
+lemma mset_ins: "mset (ins x h) = mset h + {#x#}" good_sorry
+
+
+(*<*)
+</submission>
+<check-imports>
+Submission
+</check-imports>
+<check>
+
+</check>
+<image>
+HOL-Library
+</image>
+<submission-imports>