diff -r 087d82632852 -r e04123f4bacc thys/FMap.thy --- a/thys/FMap.thy Tue Apr 29 15:26:48 2014 +0100 +++ b/thys/FMap.thy Fri May 30 12:04:49 2014 +0100 @@ -15,6 +15,8 @@ lift_definition lookup :: "'key f\ 'value \ 'key \ 'value option" is "(\ x. x)" .. +notation lookup (infixl "$" 999) + abbreviation the_lookup (infix "f!" 55) where "m f! x \ the (lookup m x)" @@ -23,7 +25,7 @@ lemma fempty_fdom[simp]: "fdom f\ = {}" by (transfer, auto) -lemma fdomIff: "(a : fdom m) = (lookup m a \ None)" +lemma fdomIff: "a \ fdom m \ lookup m a \ None" by (transfer, auto) lemma lookup_not_fdom: "x \ fdom m \ lookup m x = None" @@ -345,14 +347,44 @@ by simp qed qed + +fun alookup +where + "alookup k [] = None" +| "alookup k ((x, y) # xs) = (if (k = x) then Some y else alookup k xs)" + -definition fmapdom_to_list :: "('a :: linorder) f\ 'b \ 'a list" +definition fdom_to_list :: "('a :: linorder) f\ 'b \ 'a list" where - "fmapdom_to_list f = (THE xs. set xs = fdom f \ sorted xs \ distinct xs)" + "fdom_to_list f = (THE xs. set xs = fdom f \ sorted xs \ distinct xs)" + +lemma + "fdom f = set (fdom_to_list f)" +unfolding fdom_to_list_def +apply(subgoal_tac "finite (fdom f)") +apply(drule finite_sorted_distinct_unique) +apply(drule theI') +apply(simp_all) +done + + definition fmap_to_alist :: "('a :: linorder) f\ 'b \ ('a \ 'b) list" where - "fmap_to_alist f = [(x, f f! x). x \ fmapdom_to_list f]" + "fmap_to_alist f = [(x, f f! x). x \ fdom_to_list f]" + +lemma + shows "(f $ k) = alookup k (fmap_to_alist f)" +thm finite_dom_map_of +apply(transfer) +apply(simp add: fmap_to_alist_def) + +thm lookup_def + +apply(rule theI2) +apply(induct k xs\"fmap_to_alist f" rule: alookup.induct) +apply(simp add: fmap_to_alist_def fmapdom_to_list_def) +thm theD1 end