# HG changeset patch # User haftmann # Date 1266402119 -3600 # Node ID 58b9503a7f9adc5ee9536c9582adcd16b5f5536f # Parent ea99593b44a50f659a99fa6bd590c6d789350407# Parent 8e3b8b5f1e960ce61cbc2292e321b1c333c98435 merged diff -r ea99593b44a5 -r 58b9503a7f9a src/HOL/ex/Execute_Choice.thy --- a/src/HOL/ex/Execute_Choice.thy Wed Feb 17 11:01:01 2010 +0100 +++ b/src/HOL/ex/Execute_Choice.thy Wed Feb 17 11:21:59 2010 +0100 @@ -6,9 +6,18 @@ imports Main AssocList begin -definition valuesum :: "('a, 'b :: comm_monoid_add) mapping \ 'b" where +text {* + A trivial example: +*} + +definition valuesum :: "('a, 'b :: ab_group_add) mapping \ 'b" where "valuesum m = (\k \ Mapping.keys m. the (Mapping.lookup m k))" +text {* + Not that instead of defining @{term valuesum} with choice, we define it + directly and derive a description involving choice afterwards: +*} + lemma valuesum_rec: assumes fin: "finite (dom (Mapping.lookup m))" shows "valuesum m = (if Mapping.is_empty m then 0 else @@ -35,30 +44,59 @@ then show ?thesis by (simp add: keys_def valuesum_def is_empty_def) qed +text {* + In the context of the else-branch we can show that the exact choice is + irrelvant; in practice, finding this point where choice becomes irrelevant is the + most difficult thing! +*} + +lemma valuesum_choice: + "finite (Mapping.keys M) \ x \ Mapping.keys M \ y \ Mapping.keys M \ + the (Mapping.lookup M x) + valuesum (Mapping.delete x M) = + the (Mapping.lookup M y) + valuesum (Mapping.delete y M)" + by (simp add: valuesum_def keys_def setsum_diff) + +text {* + Given @{text valuesum_rec} as initial description, we stepwise refine it to something executable; + first, we formally insert the constructor @{term AList} and split the one equation into two, + where the second one provides the necessary context: +*} + lemma valuesum_rec_AList: - "valuesum (AList []) = 0" - "valuesum (AList (x # xs)) = (let l = (SOME l. l \ Mapping.keys (AList (x # xs))) in + shows [code]: "valuesum (AList []) = 0" + and "valuesum (AList (x # xs)) = (let l = (SOME l. l \ Mapping.keys (AList (x # xs))) in the (Mapping.lookup (AList (x # xs)) l) + valuesum (Mapping.delete l (AList (x # xs))))" by (simp_all add: valuesum_rec finite_dom_map_of is_empty_AList) -axioms - FIXME: "x \ A \ y \ A \ C x = C y" +text {* + As a side effect the precondition disappears (but note this has nothing to do with choice!). + The first equation deals with the uncritical empty case and can already be used for code generation. -lemma aux: "(SOME l. l \ Mapping.keys (AList (x # xs))) = fst (hd (x # xs))" -proof (rule FIXME) - show "fst (hd (x # xs)) \ Mapping.keys (AList (x # xs))" - by (simp add: keys_AList) - show "(SOME l. l \ Mapping.keys (AList (x # xs))) \ Mapping.keys (AList (x # xs))" - apply (rule someI) apply (simp add: keys_AList) apply auto - done -qed + Using @{text valuesum_choice}, we are able to prove an executable version of @{term valuesum}: +*} lemma valuesum_rec_exec [code]: - "valuesum (AList []) = 0" "valuesum (AList (x # xs)) = (let l = fst (hd (x # xs)) in the (Mapping.lookup (AList (x # xs)) l) + valuesum (Mapping.delete l (AList (x # xs))))" - by (simp_all add: valuesum_rec_AList aux) +proof - + let ?M = "AList (x # xs)" + let ?l1 = "(SOME l. l \ Mapping.keys ?M)" + let ?l2 = "fst (hd (x # xs))" + have "finite (Mapping.keys ?M)" by (simp add: keys_AList) + moreover have "?l1 \ Mapping.keys ?M" + by (rule someI) (auto simp add: keys_AList) + moreover have "?l2 \ Mapping.keys ?M" + by (simp add: keys_AList) + ultimately have "the (Mapping.lookup ?M ?l1) + valuesum (Mapping.delete ?l1 ?M) = + the (Mapping.lookup ?M ?l2) + valuesum (Mapping.delete ?l2 ?M)" + by (rule valuesum_choice) + then show ?thesis by (simp add: valuesum_rec_AList) +qed + +text {* + See how it works: +*} -value "valuesum (AList [(''abc'', (42::nat)), (''def'', 1705)])" +value "valuesum (AList [(''abc'', (42::int)), (''def'', 1705)])" end diff -r ea99593b44a5 -r 58b9503a7f9a src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Wed Feb 17 11:01:01 2010 +0100 +++ b/src/HOL/ex/ROOT.ML Wed Feb 17 11:21:59 2010 +0100 @@ -66,7 +66,8 @@ "Refute_Examples", "Quickcheck_Examples", "Landau", - "Execute_Choice" + "Execute_Choice", + "Summation" ]; HTML.with_charset "utf-8" (no_document use_thys) diff -r ea99593b44a5 -r 58b9503a7f9a src/HOL/ex/Summation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Summation.thy Wed Feb 17 11:21:59 2010 +0100 @@ -0,0 +1,107 @@ +(* Author: Florian Haftmann, TU Muenchen *) + +header {* Some basic facts about discrete summation *} + +theory Summation +imports Main +begin + +text {* Auxiliary. *} + +lemma add_setsum_orient: + "setsum f {k.. k < l \ setsum f {j.. :: "(int \ 'a\ab_group_add) \ int \ 'a" where + "\ f k = f (k + 1) - f k" + +lemma \_shift: + "\ (\k. l + f k) = \ f" + by (simp add: \_def expand_fun_eq) + +lemma \_same_shift: + assumes "\ f = \ g" + shows "\l. (op +) l \ f = g" +proof - + fix k + from assms have "\k. \ f k = \ g k" by simp + then have k_incr: "\k. f (k + 1) - g (k + 1) = f k - g k" + by (simp add: \_def algebra_simps) + then have "\k. f ((k - 1) + 1) - g ((k - 1) + 1) = f (k - 1) - g (k - 1)" + by blast + then have k_decr: "\k. f (k - 1) - g (k - 1) = f k - g k" + by simp + have "\k. f k - g k = f 0 - g 0" + proof - + fix k + show "f k - g k = f 0 - g 0" + by (induct k rule: int_induct) (simp_all add: k_incr k_decr) + qed + then have "\k. ((op +) (g 0 - f 0) \ f) k = g k" + by (simp add: algebra_simps) + then have "(op +) (g 0 - f 0) \ f = g" .. + then show ?thesis .. +qed + +text {* The formal sum operator. *} + +definition \ :: "(int \ 'a\ab_group_add) \ int \ int \ 'a" where + "\ f j l = (if j < l then setsum f {j.. l then - setsum f {l.._same [simp]: + "\ f j j = 0" + by (simp add: \_def) + +lemma \_positive: + "j < l \ \ f j l = setsum f {j.._def) + +lemma \_negative: + "j > l \ \ f j l = - \ f l j" + by (simp add: \_def) + +lemma add_\: + "\ f j k + \ f k l = \ f j l" + by (simp add: \_def algebra_simps add_setsum_int) + (simp_all add: add_setsum_orient [of f k j l] + add_setsum_orient [of f j l k] + add_setsum_orient [of f j k l] add_setsum_int) + +lemma \_incr_upper: + "\ f j (l + 1) = \ f j l + f l" +proof - + have "{l.. f l (l + 1) = f l" by (simp add: \_def) + moreover have "\ f j (l + 1) = \ f j l + \ f l (l + 1)" by (simp add: add_\) + ultimately show ?thesis by simp +qed + +text {* Fundamental lemmas: The relation between @{term \} and @{term \}. *} + +lemma \_\: + "\ (\ f j) = f" +proof + fix k + show "\ (\ f j) k = f k" + by (simp add: \_def \_incr_upper) +qed + +lemma \_\: + "\ (\ f) j l = f l - f j" +proof - + from \_\ have "\ (\ (\ f) j) = \ f" . + then obtain k where "(op +) k \ \ (\ f) j = f" by (blast dest: \_same_shift) + then have "\q. f q = k + \ (\ f) j q" by (simp add: expand_fun_eq) + then show ?thesis by simp +qed + +end