merged
authorhaftmann
Wed, 17 Feb 2010 09:51:46 +0100
changeset 35161 be96405ccaf3
parent 35155 3011b2149089 (current diff)
parent 35160 6eb2b6c1d2d5 (diff)
child 35162 ea99593b44a5
child 35163 2e0966d6f951
child 35182 4c39632b811f
merged
--- a/src/HOL/Library/AssocList.thy	Tue Feb 16 20:42:44 2010 +0100
+++ b/src/HOL/Library/AssocList.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -5,7 +5,7 @@
 header {* Map operations implemented on association lists*}
 
 theory AssocList 
-imports Main
+imports Main Mapping
 begin
 
 text {*
@@ -656,4 +656,48 @@
     (map_of xs k = None \<or> (\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = None)) " 
   by (simp add: compose_conv map_comp_None_iff)
 
+
+subsection {* Implementation of mappings *}
+
+definition AList :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) mapping" where
+  "AList xs = Mapping (map_of xs)"
+
+code_datatype AList
+
+lemma lookup_AList [simp, code]:
+  "Mapping.lookup (AList xs) = map_of xs"
+  by (simp add: AList_def)
+
+lemma empty_AList [code]:
+  "Mapping.empty = AList []"
+  by (rule mapping_eqI) simp
+
+lemma is_empty_AList [code]:
+  "Mapping.is_empty (AList xs) \<longleftrightarrow> null xs"
+  by (cases xs) (simp_all add: is_empty_def)
+
+lemma update_AList [code]:
+  "Mapping.update k v (AList xs) = AList (update k v xs)"
+  by (rule mapping_eqI) (simp add: update_conv')
+
+lemma delete_AList [code]:
+  "Mapping.delete k (AList xs) = AList (delete k xs)"
+  by (rule mapping_eqI) (simp add: delete_conv')
+
+lemma keys_AList [code]:
+  "Mapping.keys (AList xs) = set (map fst xs)"
+  by (simp add: keys_def dom_map_of_conv_image_fst)
+
+lemma size_AList [code]:
+  "Mapping.size (AList xs) = length (remdups (map fst xs))"
+  by (simp add: size_def length_remdups_card_conv dom_map_of_conv_image_fst)
+
+lemma tabulate_AList [code]:
+  "Mapping.tabulate ks f = AList (map (\<lambda>k. (k, f k)) ks)"
+  by (rule mapping_eqI) (simp add: map_of_map_restrict)
+
+lemma bulkload_AList [code]:
+  "Mapping.bulkload vs = AList (map (\<lambda>n. (n, vs ! n)) [0..<length vs])"
+  by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
+
 end
--- a/src/HOL/Library/Mapping.thy	Tue Feb 16 20:42:44 2010 +0100
+++ b/src/HOL/Library/Mapping.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -3,50 +3,58 @@
 header {* An abstract view on maps for code generation. *}
 
 theory Mapping
-imports Map Main
+imports Main
 begin
 
 subsection {* Type definition and primitive operations *}
 
-datatype ('a, 'b) map = Map "'a \<rightharpoonup> 'b"
+datatype ('a, 'b) mapping = Mapping "'a \<rightharpoonup> 'b"
 
-definition empty :: "('a, 'b) map" where
-  "empty = Map (\<lambda>_. None)"
-
-primrec lookup :: "('a, 'b) map \<Rightarrow> 'a \<rightharpoonup> 'b" where
-  "lookup (Map f) = f"
+definition empty :: "('a, 'b) mapping" where
+  "empty = Mapping (\<lambda>_. None)"
 
-primrec update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
-  "update k v (Map f) = Map (f (k \<mapsto> v))"
+primrec lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<rightharpoonup> 'b" where
+  "lookup (Mapping f) = f"
 
-primrec delete :: "'a \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
-  "delete k (Map f) = Map (f (k := None))"
+primrec update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
+  "update k v (Mapping f) = Mapping (f (k \<mapsto> v))"
 
-primrec keys :: "('a, 'b) map \<Rightarrow> 'a set" where
-  "keys (Map f) = dom f"
+primrec delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
+  "delete k (Mapping f) = Mapping (f (k := None))"
 
 
 subsection {* Derived operations *}
 
-definition size :: "('a, 'b) map \<Rightarrow> nat" where
-  "size m = (if finite (keys m) then card (keys m) else 0)"
+definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set" where
+  "keys m = dom (lookup m)"
 
-definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
+definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool" where
+  "is_empty m \<longleftrightarrow> dom (lookup m) = {}"
+
+definition size :: "('a, 'b) mapping \<Rightarrow> nat" where
+  "size m = (if finite (dom (lookup m)) then card (dom (lookup m)) else 0)"
+
+definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
   "replace k v m = (if lookup m k = None then m else update k v m)"
 
-definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) map" where
-  "tabulate ks f = Map (map_of (map (\<lambda>k. (k, f k)) ks))"
+definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping" where
+  "tabulate ks f = Mapping (map_of (map (\<lambda>k. (k, f k)) ks))"
 
-definition bulkload :: "'a list \<Rightarrow> (nat, 'a) map" where
-  "bulkload xs = Map (\<lambda>k. if k < length xs then Some (xs ! k) else None)"
+definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping" where
+  "bulkload xs = Mapping (\<lambda>k. if k < length xs then Some (xs ! k) else None)"
 
 
 subsection {* Properties *}
 
-lemma lookup_inject:
+lemma lookup_inject [simp]:
   "lookup m = lookup n \<longleftrightarrow> m = n"
   by (cases m, cases n) simp
 
+lemma mapping_eqI:
+  assumes "lookup m = lookup n"
+  shows "m = n"
+  using assms by simp
+
 lemma lookup_empty [simp]:
   "lookup empty = Map.empty"
   by (simp add: empty_def)
@@ -55,98 +63,82 @@
   "lookup (update k v m) = (lookup m) (k \<mapsto> v)"
   by (cases m) simp
 
-lemma lookup_delete:
-  "lookup (delete k m) k = None"
-  "k \<noteq> l \<Longrightarrow> lookup (delete k m) l = lookup m l"
-  by (cases m, simp)+
+lemma lookup_delete [simp]:
+  "lookup (delete k m) = (lookup m) (k := None)"
+  by (cases m) simp
 
-lemma lookup_tabulate:
+lemma lookup_tabulate [simp]:
   "lookup (tabulate ks f) = (Some o f) |` set ks"
   by (induct ks) (auto simp add: tabulate_def restrict_map_def expand_fun_eq)
 
-lemma lookup_bulkload:
+lemma lookup_bulkload [simp]:
   "lookup (bulkload xs) = (\<lambda>k. if k < length xs then Some (xs ! k) else None)"
-  unfolding bulkload_def by simp
+  by (simp add: bulkload_def)
 
 lemma update_update:
   "update k v (update k w m) = update k v m"
   "k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"
-  by (cases m, simp add: expand_fun_eq)+
+  by (rule mapping_eqI, simp add: fun_upd_twist)+
 
-lemma replace_update:
-  "lookup m k = None \<Longrightarrow> replace k v m = m"
-  "lookup m k \<noteq> None \<Longrightarrow> replace k v m = update k v m"
-  by (auto simp add: replace_def)
-
-lemma delete_empty [simp]:
-  "delete k empty = empty"
-  by (simp add: empty_def)
+lemma update_delete [simp]:
+  "update k v (delete k m) = update k v m"
+  by (rule mapping_eqI) simp
 
 lemma delete_update:
   "delete k (update k v m) = delete k m"
   "k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"
-  by (cases m, simp add: expand_fun_eq)+
-
-lemma update_delete [simp]:
-  "update k v (delete k m) = update k v m"
-  by (cases m) simp
-
-lemma keys_empty [simp]:
-  "keys empty = {}"
-  unfolding empty_def by simp
+  by (rule mapping_eqI, simp add: fun_upd_twist)+
 
-lemma keys_update [simp]:
-  "keys (update k v m) = insert k (keys m)"
-  by (cases m) simp
+lemma delete_empty [simp]:
+  "delete k empty = empty"
+  by (rule mapping_eqI) simp
 
-lemma keys_delete [simp]:
-  "keys (delete k m) = keys m - {k}"
-  by (cases m) simp
-
-lemma keys_tabulate [simp]:
-  "keys (tabulate ks f) = set ks"
-  by (auto simp add: tabulate_def dest: map_of_SomeD intro!: weak_map_of_SomeI)
+lemma replace_update:
+  "k \<notin> dom (lookup m) \<Longrightarrow> replace k v m = m"
+  "k \<in> dom (lookup m) \<Longrightarrow> replace k v m = update k v m"
+  by (rule mapping_eqI, auto simp add: replace_def fun_upd_twist)+
 
 lemma size_empty [simp]:
   "size empty = 0"
-  by (simp add: size_def keys_empty)
+  by (simp add: size_def)
 
 lemma size_update:
-  "finite (keys m) \<Longrightarrow> size (update k v m) =
-    (if k \<in> keys m then size m else Suc (size m))"
-  by (simp add: size_def keys_update)
-    (auto simp only: card_insert card_Suc_Diff1)
+  "finite (dom (lookup m)) \<Longrightarrow> size (update k v m) =
+    (if k \<in> dom (lookup m) then size m else Suc (size m))"
+  by (auto simp add: size_def insert_dom)
 
 lemma size_delete:
-  "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
-  by (simp add: size_def keys_delete)
+  "size (delete k m) = (if k \<in> dom (lookup m) then size m - 1 else size m)"
+  by (simp add: size_def)
 
 lemma size_tabulate:
   "size (tabulate ks f) = length (remdups ks)"
-  by (simp add: size_def keys_tabulate distinct_card [of "remdups ks", symmetric])
+  by (simp add: size_def distinct_card [of "remdups ks", symmetric] comp_def)
 
 lemma bulkload_tabulate:
   "bulkload xs = tabulate [0..<length xs] (nth xs)"
-  by (rule sym)
-    (auto simp add: bulkload_def tabulate_def expand_fun_eq map_of_eq_None_iff comp_def)
+  by (rule mapping_eqI) (simp add: expand_fun_eq)
 
 
 subsection {* Some technical code lemmas *}
 
 lemma [code]:
-  "map_case f m = f (Mapping.lookup m)"
+  "mapping_case f m = f (Mapping.lookup m)"
   by (cases m) simp
 
 lemma [code]:
-  "map_rec f m = f (Mapping.lookup m)"
+  "mapping_rec f m = f (Mapping.lookup m)"
   by (cases m) simp
 
 lemma [code]:
-  "Nat.size (m :: (_, _) map) = 0"
+  "Nat.size (m :: (_, _) mapping) = 0"
   by (cases m) simp
 
 lemma [code]:
-  "map_size f g m = 0"
+  "mapping_size f g m = 0"
   by (cases m) simp
 
+
+hide (open) const empty is_empty lookup update delete keys size replace tabulate bulkload
+
 end
\ No newline at end of file
--- a/src/HOL/Library/Tree.thy	Tue Feb 16 20:42:44 2010 +0100
+++ b/src/HOL/Library/Tree.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -96,11 +96,11 @@
 
 subsection {* Trees as mappings *}
 
-definition Tree :: "('a\<Colon>linorder, 'b) tree \<Rightarrow> ('a, 'b) map" where
-  "Tree t = Map (Tree.lookup t)"
+definition Tree :: "('a\<Colon>linorder, 'b) tree \<Rightarrow> ('a, 'b) mapping" where
+  "Tree t = Mapping (Tree.lookup t)"
 
 lemma [code, code del]:
-  "(eq_class.eq :: (_, _) map \<Rightarrow> _) = eq_class.eq" ..
+  "(eq_class.eq :: (_, _) mapping \<Rightarrow> _) = eq_class.eq" ..
 
 lemma [code, code del]:
   "Mapping.delete k m = Mapping.delete k m" ..
@@ -115,13 +115,18 @@
   "Mapping.lookup (Tree t) = lookup t"
   by (simp add: Tree_def)
 
+lemma is_empty_Tree [code]:
+  "Mapping.is_empty (Tree Empty) \<longleftrightarrow> True"
+  "Mapping.is_empty (Tree (Branch v k l r)) \<longleftrightarrow> False"
+  by (simp_all only: is_empty_def lookup_Tree dom_lookup) auto
+
 lemma update_Tree [code]:
   "Mapping.update k v (Tree t) = Tree (update k v t)"
   by (simp add: Tree_def lookup_update)
 
 lemma keys_Tree [code]:
   "Mapping.keys (Tree t) = set (filter (\<lambda>k. lookup t k \<noteq> None) (remdups (keys t)))"
-  by (simp add: Tree_def dom_lookup)
+  by (simp add: Mapping.keys_def lookup_Tree dom_lookup)
 
 lemma size_Tree [code]:
   "Mapping.size (Tree t) = size t"
@@ -135,8 +140,10 @@
   "Mapping.tabulate ks f = Tree (bulkload (sort ks) f)"
 proof -
   have "Mapping.lookup (Mapping.tabulate ks f) = Mapping.lookup (Tree (bulkload (sort ks) f))"
-    by (simp add: lookup_Tree lookup_bulkload lookup_tabulate)
-  then show ?thesis by (simp add: lookup_inject)
+    by (simp add: lookup_bulkload lookup_Tree)
+  then show ?thesis by (simp only: lookup_inject)
 qed
 
+hide (open) const Empty Branch lookup update keys size bulkload Tree
+
 end
--- a/src/HOL/Map.thy	Tue Feb 16 20:42:44 2010 +0100
+++ b/src/HOL/Map.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -389,6 +389,10 @@
   "x \<in> D \<Longrightarrow> (m|`D)(x := y) = (m|`(D-{x}))(x := y)"
 by (simp add: restrict_map_def expand_fun_eq)
 
+lemma map_of_map_restrict:
+  "map_of (map (\<lambda>k. (k, f k)) ks) = (Some \<circ> f) |` set ks"
+  by (induct ks) (simp_all add: expand_fun_eq restrict_map_insert)
+
 
 subsection {* @{term [source] map_upds} *}
 
@@ -534,7 +538,7 @@
 by (auto simp add: map_add_def split: option.split_asm)
 
 lemma dom_const [simp]:
-  "dom (\<lambda>x. Some y) = UNIV"
+  "dom (\<lambda>x. Some (f x)) = UNIV"
   by auto
 
 (* Due to John Matthews - could be rephrased with dom *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Execute_Choice.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -0,0 +1,64 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* A simple cookbook example how to eliminate choice in programs. *}
+
+theory Execute_Choice
+imports Main AssocList
+begin
+
+definition valuesum :: "('a, 'b :: comm_monoid_add) mapping \<Rightarrow> 'b" where
+  "valuesum m = (\<Sum>k \<in> Mapping.keys m. the (Mapping.lookup m k))"
+
+lemma valuesum_rec:
+  assumes fin: "finite (dom (Mapping.lookup m))"
+  shows "valuesum m = (if Mapping.is_empty m then 0 else
+    let l = (SOME l. l \<in> Mapping.keys m) in the (Mapping.lookup m l) + valuesum (Mapping.delete l m))"
+proof (cases "Mapping.is_empty m")
+  case True then show ?thesis by (simp add: is_empty_def keys_def valuesum_def)
+next
+  case False
+  then have l: "\<exists>l. l \<in> dom (Mapping.lookup m)" by (auto simp add: is_empty_def expand_fun_eq mem_def)
+  then have "(let l = SOME l. l \<in> dom (Mapping.lookup m) in
+     the (Mapping.lookup m l) + (\<Sum>k \<in> dom (Mapping.lookup m) - {l}. the (Mapping.lookup m k))) =
+       (\<Sum>k \<in> dom (Mapping.lookup m). the (Mapping.lookup m k))"
+  proof (rule someI2_ex)
+    fix l
+    note fin
+    moreover assume "l \<in> dom (Mapping.lookup m)"
+    moreover obtain A where "A = dom (Mapping.lookup m) - {l}" by simp
+    ultimately have "dom (Mapping.lookup m) = insert l A" and "finite A" and "l \<notin> A" by auto
+    then show "(let l = l
+        in the (Mapping.lookup m l) + (\<Sum>k \<in> dom (Mapping.lookup m) - {l}. the (Mapping.lookup m k))) =
+        (\<Sum>k \<in> dom (Mapping.lookup m). the (Mapping.lookup m k))"
+      by simp
+   qed
+  then show ?thesis by (simp add: keys_def valuesum_def is_empty_def)
+qed
+
+lemma valuesum_rec_AList:
+  "valuesum (AList []) = 0"
+  "valuesum (AList (x # xs)) = (let l = (SOME l. l \<in> 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 \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> C x = C y"
+
+lemma aux: "(SOME l. l \<in> Mapping.keys (AList (x # xs))) = fst (hd (x # xs))"
+proof (rule FIXME)
+  show "fst (hd (x # xs)) \<in> Mapping.keys (AList (x # xs))"
+    by (simp add: keys_AList)
+  show "(SOME l. l \<in> Mapping.keys (AList (x # xs))) \<in> Mapping.keys (AList (x # xs))"
+    apply (rule someI) apply (simp add: keys_AList) apply auto
+    done
+qed
+
+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)
+
+value "valuesum (AList [(''abc'', (42::nat)), (''def'', 1705)])"
+
+end
--- a/src/HOL/ex/ROOT.ML	Tue Feb 16 20:42:44 2010 +0100
+++ b/src/HOL/ex/ROOT.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -65,7 +65,8 @@
   "HarmonicSeries",
   "Refute_Examples",
   "Quickcheck_Examples",
-  "Landau"
+  "Landau",
+  "Execute_Choice"
 ];
 
 HTML.with_charset "utf-8" (no_document use_thys)