theory LiveDemo04 imports Main "HOL-Library.Multiset" begin section \Gauß' Formula: Use Chains of Equality\ fun sum_nats :: "nat \ nat" where "sum_nats 0 = 0" | "sum_nats (Suc n) = Suc n + sum_nats n" lemma Gauss: "sum_nats n = (n * Suc n) div 2" proof (induction n) case (Suc n) have "sum_nats (Suc n) = Suc n + sum_nats n" by auto also have "\ = Suc n + (n * Suc n div 2)" using Suc by auto also have "\ = 2 * (Suc n) div 2 + (n * Suc n div 2)" by auto also have "\ = (2 * Suc n + n * Suc n) div 2" by auto also have "2 * Suc n + n * Suc n = ((Suc n * Suc (Suc n)))" by auto finally show ?case . qed auto section \Fast Reversal: Use arbitrary variables\ text \Isabelle has already a defined append function @{const append} which usually written in infix notion: @{term "xs @ ys"}. (Haskell's ++ = Isabelle's @)\ fun reverse :: "'a list \ 'a list" where "reverse [] = []" | "reverse (x # xs) = reverse xs @ [x]" fun rev_it :: "'a list \ 'a list \ 'a list" where "rev_it [] ys = ys" | "rev_it (x # xs) ys = rev_it xs (x # ys)" fun fast_rev :: "'a list \ 'a list" where "fast_rev xs = rev_it xs []" lemma fast_rev: "fast_rev xs = reverse xs" proof - have "rev_it xs [] = reverse xs" proof (induction xs) case (Cons x xs) then show ?case proof auto oops lemma rev_it[simp]: "rev_it xs ys = reverse xs @ ys" proof (induction xs arbitrary: ys) case (Cons x xs ys) then show ?case by auto qed auto lemma fast_rev: "fast_rev xs = reverse xs" using [[simp_trace]] by auto text \Since we have proven a faster implementation of reverse, we can now tell the system to use it via a [code] attribute.\ export_code reverse in Haskell lemma implement_rev_by_fast_rev[code]: "reverse xs = fast_rev xs" by (auto simp: fast_rev) definition "foo xs = xs @ reverse xs" export_code foo in Haskell text \Advantage: we can have two versions of an algorithm: \<^item> @{const reverse} has a structure which makes it easy to reason about \<^item> @{const fast_rev} is better in generated code\ section \Controlled Proof Search\ lemma Gauss_2: "sum_nats n = (n * Suc n) div 2" proof (induction n) case (Suc n) show ?case unfolding sum_nats.simps unfolding Suc.IH apply simp done qed auto text \After successful proof search, one might turn apply-style proof into more readable structured proof.\ section \Insertion Sort: Induction with Implications\ fun insert_sorted :: "'a :: linorder \ 'a list \ 'a list" where "insert_sorted x Nil = Cons x Nil" | "insert_sorted x (Cons y ys) = (if x \ y then Cons x (Cons y ys) else Cons y (insert_sorted x ys))" fun ins_sort :: "'a :: linorder list \ 'a list" where "ins_sort Nil = Nil" | "ins_sort (Cons x xs) = insert_sorted x (ins_sort xs)" text \Previous exercise: development of proof that length is preserved.\ text \Now: prove that result is @{const sorted}.\ lemma sorted_insert_sorted: "sorted xs \ sorted (insert_sorted x xs)" proof (induction xs) case (Cons y xs) thm Cons.IH (* induction hypthesis via .IH *) thm Cons.prems (* premises via .prems *) from Cons (* everything with .IH or .prems *) have IH: "sorted (insert_sorted x xs)" by auto show ?case proof (cases "x \ y") case True have "?thesis = sorted (x # y # xs)" using True by auto also have "\" using Cons.prems True by auto finally show ?thesis . next case False have "sorted (insert_sorted x (y # xs)) = sorted (y # insert_sorted x xs)" using False by auto (* next line detected via: also have "... = something" apply simp, and then read proof goal of what "something" has to be *) also have "\ = (((\x\set (insert_sorted x xs). y \ x) \ sorted (insert_sorted x xs)))" by simp also have "\ = (\x\set (insert_sorted x xs). y \ x)" using IH by auto (* detect intermediate lemma *) also have "set (insert_sorted x xs) = set (x # xs)" by (induction xs) auto finally have "sorted (insert_sorted x (y # xs)) = (\x\set (x # xs). y \ x)" . also have \ using Cons.prems False by auto finally show ?thesis by auto qed qed auto text \Now soundness of insertion sort is easy, using @{thm sorted_insert_sorted} as conditional(!) simp rule.\ declare sorted_insert_sorted[simp] lemma sorted_ins_sort: "sorted (ins_sort xs)" by (induction xs) auto text \Conditional simplification rules naturally appear also in arithmetic reasoning, e.g. @{thm mult_le_cancel_left_pos}\ text \Just for illustration purposes: soundness of insert_sorted using purely apply-style.\ lemma set_insert_sorted: "set (insert_sorted x xs) = set (x # xs)" apply (induction xs) apply auto done lemma sorted_insert_sorted_apply_style: "sorted xs \ sorted (insert_sorted x xs)" apply (induction xs) subgoal by simp (* Nil-Case *) subgoal for y xs (* fix variables for Cons-case *) apply (cases "x \ y") subgoal by auto (* trivial case: x \ y *) apply simp (* detect required lemma *) apply (subst set_insert_sorted) apply auto done done text \Another illustration: soundness of insert_sorted using mixture of apply-style and Isar-style.\ lemma sorted_insert_sorted_mixed_style: "sorted xs \ sorted (insert_sorted x xs)" apply (induction xs) subgoal by simp subgoal for y xs apply (cases "x \ y") subgoal by auto apply auto (* bad style to open Isar proof after auto *) proof (goal_cases) (* switch to Isar via goal_cases, each subgoal gets a number *) case (1 z) have eq: "set (insert_sorted x xs) = set (x # xs)" by (induction xs) auto show ?case using 1 unfolding eq by auto qed done text \Preservation of elements is trivially expressed via multisets.\ lemma mset_insert_sorted[simp]: "mset (insert_sorted x xs) = mset (x # xs)" by (induction xs) auto lemma mset_ins_sort[simp]: "mset (ins_sort xs) = mset xs" by (induction xs) auto text \Insertion sort is the same as Isabelle's @{const sort} function (actually every correct sorting algorithm has the same functional behavior as @{const sort}.)\ lemma "sort xs = ins_sort xs" proof (rule properties_for_sort) show "mset (ins_sort xs) = mset xs" by auto show "sorted (ins_sort xs)" by (rule sorted_ins_sort) qed end