theory Demo05 imports Main begin section \<open>Function Definitions Revisited\<close> subsection \<open>Overlapping Patterns\<close> fun drop_last :: "'a list \<Rightarrow> 'a list" where "drop_last (x # y # ys) = x # drop_last (y # ys)" | "drop_last xs = []" text \<open>Simplification rules use disjoint patterns\<close> thm drop_last.simps subsection \<open>Underspecification\<close> fun head1 where "head1 (x # xs) = x" (* observe missing pattern warning and "undefined" *) thm head1.simps (* does not cover undefined-case *) fun head2 where "head2 xs = (case xs of x # _ \<Rightarrow> x)" thm head2.simps (* undefined case is not mentioned *) ML \<open> (* observe the internal usage of @{const undefined} in the case-expression *) Thm.prop_of @{thm head2.simps} \<close> (* even if we specify undefined cases explicitly, these are not shown *) term "case xs of [] \<Rightarrow> undefined | x # _ \<Rightarrow> x" subsection \<open>Computation Induction\<close> text \<open>Computation induction rule uses exactly the three cases\<close> thm drop_last.induct lemma "length (drop_last xs) = length xs - 1" proof (induction xs rule: drop_last.induct) text \<open>Get three subgoals, for each of the cases. All are automatic, the cases are numbered.\<close> case (1 x y ys) thm 1 (* suffixes .IH and .prems after numbers have to be put in double-quotes *) thm "1.IH" thus ?case by auto qed auto text \<open>No way to cheat! All cases have to be covered, even if these have not been specified: induction rule covers empty list.\<close> thm head1.simps thm head1.induct section \<open>Manual Termination Proofs\<close> function gen_list :: "nat \<Rightarrow> nat \<Rightarrow> nat list" where "gen_list n m = (if n \<le> m then n # gen_list (Suc n) m else [])" by pat_completeness auto text \<open>Evaluation is not yet possible, since termination has not been proven! Similarly, there are no simp rules and there is no computation induction rule\<close> value "gen_list 3 5" thm gen_list.simps thm gen_list.induct termination proof oops termination apply (relation "measure (\<lambda> (n,m). Suc m - n)") apply simp apply (simp only: in_measure, unfold split) oops termination by (relation "measure (\<lambda> (n,m). Suc m - n)") auto print_theorems (* only after proving termination we get simp-rules and induction-rule *) value "gen_list 3 10" lemma "gen_list n m = gen_list n (Suc m)" (* \<open>apply simp\<close> loops, since lhs of gen_list does always match *) oops declare gen_list.simps[simp del] lemma "length (gen_list n m) = Suc m - n" proof (induction n m rule: gen_list.induct) case (1 n m) text \<open>locally enable simplification for specific arguments with @{attribute of}\<close> note [simp] = gen_list.simps[of n m] from 1 show ?case by auto qed section \<open>Attributes\<close> text \<open>Here it is perfectly fine to just state the witness\<close> lemma "\<exists> x :: nat. x mod 7 = 0 \<and> x > 42 \<and> x * x < 2500" by (rule exI[of _ 49]) auto text \<open>In the Isar proof one would even have to add type-constraints\<close> lemma "\<exists> x :: nat. x mod 7 = 0 \<and> x > 42 \<and> x * x < 2500" proof show "(49 :: nat) mod 7 = 0 \<and> (49 :: nat) > 42 \<and> (49 :: nat) * 49 < 2500" by auto qed text \<open>Less type-constraints via let\<close> lemma "\<exists> x :: nat. x mod 7 = 0 \<and> x > 42 \<and> x * x < 2500" proof let ?x = "49 :: nat" show "?x mod 7 = 0 \<and> ?x > 42 \<and> ?x * ?x < 2500" by auto qed text \<open>Preview: use pattern matching via \<open>(is "term")\<close> after proposition\<close> lemma "\<exists> x :: nat. x mod 7 = 0 \<and> x > 42 \<and> x * x < 2500" (is "\<exists> x. ?P x") proof term ?P (* abbreviation ?P is created from pattern matching above *) show "?P 49" by auto qed text \<open>In the following example the Isar style is much more readable than using complicated attributes\<close> lemma "\<exists> x :: nat. (\<exists> y. y * y = x) \<and> x > 42" apply (rule exI[of _ 49, OF conjI[OF exI[of _ 7]]]) by auto thm exI[of _ 49, OF conjI[OF exI[of _ 7]]] lemma "\<exists> x :: nat. (\<exists> y. y * y = x) \<and> x > 42" proof (intro exI conjI) (* intro applies introduction rules exhaustively *) show "7 * 7 = (49 :: nat)" by simp qed auto lemma "\<exists> x :: nat. (\<exists> y. y * y = x) \<and> x > 42" proof (intro exI[of _ 7] exI[of _ 49] conjI) (* this does not work out, since always the first instantiated exI-theorem is applied *) show "7 * 7 = (7 :: nat)" oops section \<open>Quicksort\<close> fun split :: "'a :: linorder \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" where "split a [] = ([], [])" | "split a (x # xs) = (case split a xs of (low, high) \<Rightarrow> if x \<le> a then (x # low, high) else (low, x # high))" text \<open>Advantage: reuse existing library functions such as @{const filter} for which a lot of lemmas have been proven.\<close> lemma split[termination_simp]: "split a xs = (filter (\<lambda> x. x \<le> a) xs, filter (\<lambda> x. \<not> x \<le> a) xs)" by (induction xs) auto fun qsort :: "'a :: linorder list \<Rightarrow> 'a list" where "qsort [] = []" | "qsort (x # xs) = (case split x xs of (low, high) \<Rightarrow> qsort low @ [x] @ qsort high)" text \<open>Auxiliary property that is required in next proof.\<close> lemma set_qsort[simp]: "set (qsort xs) = set xs" proof (induction xs rule: qsort.induct) case (2 x xs) then show ?case by (auto simp: split) qed auto lemma sorted_qsort: "sorted (qsort xs)" text \<open>use induction rule coming from qsort-algorithm\<close> proof (induction xs rule: qsort.induct) case (2 x xs) text \<open>avoid writing term @{term "split x xs"} twice by using pattern matching via \<open>(is "term")\<close>\<close> obtain low high where s: "split x xs = (low, high)" (is "?e = _") by (cases ?e) auto thm 2 text \<open>pipe symmetric version of s into IH @{thm 2} via @{attribute OF} and store result as fact IH\<close> note IH = 2[OF s[symmetric]] have "?case = (sorted (qsort low @ [x] @ qsort high))" using s by auto find_theorems "sorted (_ @ _)" also have "\<dots> = ((\<forall> y \<in> set high. x \<le> y) \<and> (\<forall>y\<in>set (qsort low). y \<le> x \<and> (\<forall> z \<in> set high. y \<le> z)))" unfolding sorted_append using IH by auto also have \<dots> using s by (auto simp: split) text \<open>In the step above, we profit from several properties that have been proven for @{const filter}, so that @{method auto} succeeds. These become available since adding @{thm split} switches from @{const split} to @{const filter}.\<close> finally show ?case . qed auto end