--- a/src/HOL/Map.thy Tue Aug 04 14:06:24 2015 +0200
+++ b/src/HOL/Map.thy Tue Aug 04 14:29:45 2015 +0200
@@ -11,65 +11,65 @@
imports List
begin
-type_synonym ('a, 'b) "map" = "'a => 'b option" (infixr "~=>" 0)
+type_synonym ('a, 'b) "map" = "'a \<Rightarrow> 'b option" (infixr "~=>" 0)
type_notation (xsymbols)
"map" (infixr "\<rightharpoonup>" 0)
abbreviation
empty :: "'a \<rightharpoonup> 'b" where
- "empty == %x. None"
+ "empty \<equiv> \<lambda>x. None"
definition
- map_comp :: "('b \<rightharpoonup> 'c) => ('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> 'c)" (infixl "o'_m" 55) where
+ map_comp :: "('b \<rightharpoonup> 'c) \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'c)" (infixl "o'_m" 55) where
"f o_m g = (\<lambda>k. case g k of None \<Rightarrow> None | Some v \<Rightarrow> f v)"
notation (xsymbols)
map_comp (infixl "\<circ>\<^sub>m" 55)
definition
- map_add :: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> 'b)" (infixl "++" 100) where
- "m1 ++ m2 = (\<lambda>x. case m2 x of None => m1 x | Some y => Some y)"
+ map_add :: "('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b)" (infixl "++" 100) where
+ "m1 ++ m2 = (\<lambda>x. case m2 x of None \<Rightarrow> m1 x | Some y \<Rightarrow> Some y)"
definition
- restrict_map :: "('a \<rightharpoonup> 'b) => 'a set => ('a \<rightharpoonup> 'b)" (infixl "|`" 110) where
- "m|`A = (\<lambda>x. if x : A then m x else None)"
+ restrict_map :: "('a \<rightharpoonup> 'b) \<Rightarrow> 'a set \<Rightarrow> ('a \<rightharpoonup> 'b)" (infixl "|`" 110) where
+ "m|`A = (\<lambda>x. if x \<in> A then m x else None)"
notation (latex output)
restrict_map ("_\<restriction>\<^bsub>_\<^esub>" [111,110] 110)
definition
- dom :: "('a \<rightharpoonup> 'b) => 'a set" where
- "dom m = {a. m a ~= None}"
+ dom :: "('a \<rightharpoonup> 'b) \<Rightarrow> 'a set" where
+ "dom m = {a. m a \<noteq> None}"
definition
- ran :: "('a \<rightharpoonup> 'b) => 'b set" where
- "ran m = {b. EX a. m a = Some b}"
+ ran :: "('a \<rightharpoonup> 'b) \<Rightarrow> 'b set" where
+ "ran m = {b. \<exists>a. m a = Some b}"
definition
- map_le :: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50) where
- "(m\<^sub>1 \<subseteq>\<^sub>m m\<^sub>2) = (\<forall>a \<in> dom m\<^sub>1. m\<^sub>1 a = m\<^sub>2 a)"
+ map_le :: "('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> bool" (infix "\<subseteq>\<^sub>m" 50) where
+ "(m\<^sub>1 \<subseteq>\<^sub>m m\<^sub>2) \<longleftrightarrow> (\<forall>a \<in> dom m\<^sub>1. m\<^sub>1 a = m\<^sub>2 a)"
nonterminal maplets and maplet
syntax
- "_maplet" :: "['a, 'a] => maplet" ("_ /|->/ _")
- "_maplets" :: "['a, 'a] => maplet" ("_ /[|->]/ _")
- "" :: "maplet => maplets" ("_")
- "_Maplets" :: "[maplet, maplets] => maplets" ("_,/ _")
- "_MapUpd" :: "['a \<rightharpoonup> 'b, maplets] => 'a \<rightharpoonup> 'b" ("_/'(_')" [900,0]900)
- "_Map" :: "maplets => 'a \<rightharpoonup> 'b" ("(1[_])")
+ "_maplet" :: "['a, 'a] \<Rightarrow> maplet" ("_ /|->/ _")
+ "_maplets" :: "['a, 'a] \<Rightarrow> maplet" ("_ /[|->]/ _")
+ "" :: "maplet \<Rightarrow> maplets" ("_")
+ "_Maplets" :: "[maplet, maplets] \<Rightarrow> maplets" ("_,/ _")
+ "_MapUpd" :: "['a \<rightharpoonup> 'b, maplets] \<Rightarrow> 'a \<rightharpoonup> 'b" ("_/'(_')" [900,0]900)
+ "_Map" :: "maplets \<Rightarrow> 'a \<rightharpoonup> 'b" ("(1[_])")
syntax (xsymbols)
- "_maplet" :: "['a, 'a] => maplet" ("_ /\<mapsto>/ _")
- "_maplets" :: "['a, 'a] => maplet" ("_ /[\<mapsto>]/ _")
+ "_maplet" :: "['a, 'a] \<Rightarrow> maplet" ("_ /\<mapsto>/ _")
+ "_maplets" :: "['a, 'a] \<Rightarrow> maplet" ("_ /[\<mapsto>]/ _")
translations
- "_MapUpd m (_Maplets xy ms)" == "_MapUpd (_MapUpd m xy) ms"
- "_MapUpd m (_maplet x y)" == "m(x := CONST Some y)"
- "_Map ms" == "_MapUpd (CONST empty) ms"
- "_Map (_Maplets ms1 ms2)" <= "_MapUpd (_Map ms1) ms2"
- "_Maplets ms1 (_Maplets ms2 ms3)" <= "_Maplets (_Maplets ms1 ms2) ms3"
+ "_MapUpd m (_Maplets xy ms)" \<rightleftharpoons> "_MapUpd (_MapUpd m xy) ms"
+ "_MapUpd m (_maplet x y)" \<rightleftharpoons> "m(x := CONST Some y)"
+ "_Map ms" \<rightleftharpoons> "_MapUpd (CONST empty) ms"
+ "_Map (_Maplets ms1 ms2)" \<leftharpoondown> "_MapUpd (_Map ms1) ms2"
+ "_Maplets ms1 (_Maplets ms2 ms3)" \<leftharpoondown> "_Maplets (_Maplets ms1 ms2) ms3"
primrec
map_of :: "('a \<times> 'b) list \<Rightarrow> 'a \<rightharpoonup> 'b" where
@@ -81,9 +81,9 @@
"map_upds m xs ys = m ++ map_of (rev (zip xs ys))"
translations
- "_MapUpd m (_maplets x y)" == "CONST map_upds m x y"
+ "_MapUpd m (_maplets x y)" \<rightleftharpoons> "CONST map_upds m x y"
-lemma map_of_Cons_code [code]:
+lemma map_of_Cons_code [code]:
"map_of [] k = None"
"map_of ((l, v) # ps) k = (if l = k then Some v else map_of ps k)"
by simp_all
@@ -92,15 +92,15 @@
subsection \<open>@{term [source] empty}\<close>
lemma empty_upd_none [simp]: "empty(x := None) = empty"
-by (rule ext) simp
+ by (rule ext) simp
subsection \<open>@{term [source] map_upd}\<close>
-lemma map_upd_triv: "t k = Some x ==> t(k\<mapsto>x) = t"
-by (rule ext) simp
+lemma map_upd_triv: "t k = Some x \<Longrightarrow> t(k\<mapsto>x) = t"
+ by (rule ext) simp
-lemma map_upd_nonempty [simp]: "t(k\<mapsto>x) ~= empty"
+lemma map_upd_nonempty [simp]: "t(k\<mapsto>x) \<noteq> empty"
proof
assume "t(k \<mapsto> x) = empty"
then have "(t(k \<mapsto> x)) k = None" by simp
@@ -122,7 +122,7 @@
lemma image_map_upd [simp]: "x \<notin> A \<Longrightarrow> m(x \<mapsto> y) ` A = m ` A"
by auto
-lemma finite_range_updI: "finite (range f) ==> finite (range (f(a\<mapsto>b)))"
+lemma finite_range_updI: "finite (range f) \<Longrightarrow> finite (range (f(a\<mapsto>b)))"
unfolding image_def
apply (simp (no_asm_use) add:full_SetCompr_eq)
apply (rule finite_subset)
@@ -152,7 +152,7 @@
lemma Some_eq_map_of_iff [simp]:
"distinct(map fst xys) \<Longrightarrow> (Some y = map_of xys x) = ((x,y) \<in> set xys)"
-by (auto simp del:map_of_eq_Some_iff simp add: map_of_eq_Some_iff [symmetric])
+by (auto simp del: map_of_eq_Some_iff simp: map_of_eq_Some_iff [symmetric])
lemma map_of_is_SomeI [simp]: "\<lbrakk> distinct(map fst xys); (x,y) \<in> set xys \<rbrakk>
\<Longrightarrow> map_of xys x = Some y"
@@ -200,7 +200,9 @@
and dist: "distinct xs"
and map_of: "map_of (zip xs ys) = map_of (zip xs zs)"
shows "ys = zs"
-using assms(1) assms(2)[symmetric] using dist map_of proof (induct ys xs zs rule: list_induct3)
+ using assms(1) assms(2)[symmetric]
+ using dist map_of
+proof (induct ys xs zs rule: list_induct3)
case Nil show ?case by simp
next
case (Cons y ys x xs z zs)
@@ -230,11 +232,11 @@
by (induct xs) (simp, atomize (full), auto)
lemma map_of_mapk_SomeI:
- "inj f ==> map_of t k = Some x ==>
- map_of (map (split (%k. Pair (f k))) t) (f k) = Some x"
-by (induct t) (auto simp add: inj_eq)
+ "inj f \<Longrightarrow> map_of t k = Some x \<Longrightarrow>
+ map_of (map (split (\<lambda>k. Pair (f k))) t) (f k) = Some x"
+by (induct t) (auto simp: inj_eq)
-lemma weak_map_of_SomeI: "(k, x) : set l ==> \<exists>x. map_of l k = Some x"
+lemma weak_map_of_SomeI: "(k, x) \<in> set l \<Longrightarrow> \<exists>x. map_of l k = Some x"
by (induct l) auto
lemma map_of_filter_in:
@@ -243,7 +245,7 @@
lemma map_of_map:
"map_of (map (\<lambda>(k, v). (k, f v)) xs) = map_option f \<circ> map_of xs"
- by (induct xs) (auto simp add: fun_eq_iff)
+ by (induct xs) (auto simp: fun_eq_iff)
lemma dom_map_option:
"dom (\<lambda>k. map_option (f k) (m k)) = dom m"
@@ -269,20 +271,20 @@
lemma map_comp_empty [simp]:
"m \<circ>\<^sub>m empty = empty"
"empty \<circ>\<^sub>m m = empty"
-by (auto simp add: map_comp_def split: option.splits)
+by (auto simp: map_comp_def split: option.splits)
lemma map_comp_simps [simp]:
"m2 k = None \<Longrightarrow> (m1 \<circ>\<^sub>m m2) k = None"
"m2 k = Some k' \<Longrightarrow> (m1 \<circ>\<^sub>m m2) k = m1 k'"
-by (auto simp add: map_comp_def)
+by (auto simp: map_comp_def)
lemma map_comp_Some_iff:
"((m1 \<circ>\<^sub>m m2) k = Some v) = (\<exists>k'. m2 k = Some k' \<and> m1 k' = Some v)"
-by (auto simp add: map_comp_def split: option.splits)
+by (auto simp: map_comp_def split: option.splits)
lemma map_comp_None_iff:
"((m1 \<circ>\<^sub>m m2) k = None) = (m2 k = None \<or> (\<exists>k'. m2 k = Some k' \<and> m1 k' = None)) "
-by (auto simp add: map_comp_def split: option.splits)
+by (auto simp: map_comp_def split: option.splits)
subsection \<open>@{text "++"}\<close>
@@ -304,7 +306,7 @@
"(m ++ n) k = Some x \<Longrightarrow> n k = Some x \<or> n k = None \<and> m k = Some x"
by (rule map_add_Some_iff [THEN iffD1])
-lemma map_add_find_right [simp]: "!!xx. n k = Some xx ==> (m ++ n) k = Some xx"
+lemma map_add_find_right [simp]: "n k = Some xx \<Longrightarrow> (m ++ n) k = Some xx"
by (subst map_add_Some_iff) fast
lemma map_add_None [iff]: "((m ++ n) k = None) = (n k = None & m k = None)"
@@ -328,7 +330,7 @@
done
lemma finite_range_map_of_map_add:
- "finite (range f) ==> finite (range (f ++ map_of l))"
+ "finite (range f) \<Longrightarrow> finite (range (f ++ map_of l))"
apply (induct l)
apply (auto simp del: fun_upd_apply)
apply (erule finite_range_updI)
@@ -349,7 +351,7 @@
lemma map_add_map_of_foldr:
"m ++ map_of ps = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) ps m"
- by (induct ps) (auto simp add: fun_eq_iff map_add_def)
+ by (induct ps) (auto simp: fun_eq_iff map_add_def)
subsection \<open>@{term [source] restrict_map}\<close>
@@ -358,7 +360,7 @@
by (simp add: restrict_map_def)
lemma restrict_map_insert: "f |` (insert a A) = (f |` A)(a := f a)"
-by (auto simp add: restrict_map_def)
+by (auto simp: restrict_map_def)
lemma restrict_map_empty [simp]: "empty|`D = empty"
by (simp add: restrict_map_def)
@@ -386,7 +388,7 @@
by (simp add: restrict_map_def fun_eq_iff)
lemma fun_upd_None_restrict [simp]:
- "(m|`D)(x := None) = (if x:D then m|`(D - {x}) else m|`D)"
+ "(m|`D)(x := None) = (if x \<in> D then m|`(D - {x}) else m|`D)"
by (simp add: restrict_map_def fun_eq_iff)
lemma fun_upd_restrict: "(m|`D)(x := y) = (m|`(D-{x}))(x := y)"
@@ -416,9 +418,9 @@
lemma map_upds_Cons [simp]: "m(a#as [\<mapsto>] b#bs) = (m(a\<mapsto>b))(as[\<mapsto>]bs)"
by (simp add:map_upds_def)
-lemma map_upds_append1 [simp]: "\<And>ys m. size xs < size ys \<Longrightarrow>
+lemma map_upds_append1 [simp]: "size xs < size ys \<Longrightarrow>
m(xs@[x] [\<mapsto>] ys) = m(xs [\<mapsto>] ys)(x \<mapsto> ys!size xs)"
-apply(induct xs)
+apply(induct xs arbitrary: ys m)
apply (clarsimp simp add: neq_Nil_conv)
apply (case_tac ys)
apply simp
@@ -436,7 +438,7 @@
lemma map_upd_upds_conv_if:
"(f(x\<mapsto>y))(xs [\<mapsto>] ys) =
- (if x : set(take (length ys) xs) then f(xs [\<mapsto>] ys)
+ (if x \<in> set(take (length ys) xs) then f(xs [\<mapsto>] ys)
else (f(xs [\<mapsto>] ys))(x\<mapsto>y))"
apply (induct xs arbitrary: x y ys f)
apply simp
@@ -445,11 +447,11 @@
done
lemma map_upds_twist [simp]:
- "a ~: set as ==> m(a\<mapsto>b)(as[\<mapsto>]bs) = m(as[\<mapsto>]bs)(a\<mapsto>b)"
+ "a \<notin> set as \<Longrightarrow> m(a\<mapsto>b)(as[\<mapsto>]bs) = m(as[\<mapsto>]bs)(a\<mapsto>b)"
using set_take_subset by (fastforce simp add: map_upd_upds_conv_if)
lemma map_upds_apply_nontin [simp]:
- "x ~: set xs ==> (f(xs[\<mapsto>]ys)) x = f x"
+ "x \<notin> set xs \<Longrightarrow> (f(xs[\<mapsto>]ys)) x = f x"
apply (induct xs arbitrary: ys)
apply simp
apply (case_tac ys)
@@ -490,22 +492,22 @@
lemma dom_eq_empty_conv [simp]: "dom f = {} \<longleftrightarrow> f = empty"
by (auto simp: dom_def)
-lemma domI: "m a = Some b ==> a : dom m"
-by(simp add:dom_def)
+lemma domI: "m a = Some b \<Longrightarrow> a \<in> dom m"
+ by (simp add: dom_def)
(* declare domI [intro]? *)
-lemma domD: "a : dom m ==> \<exists>b. m a = Some b"
-by (cases "m a") (auto simp add: dom_def)
+lemma domD: "a \<in> dom m \<Longrightarrow> \<exists>b. m a = Some b"
+ by (cases "m a") (auto simp add: dom_def)
-lemma domIff [iff, simp del]: "(a : dom m) = (m a ~= None)"
-by(simp add:dom_def)
+lemma domIff [iff, simp del]: "a \<in> dom m \<longleftrightarrow> m a \<noteq> None"
+ by (simp add: dom_def)
lemma dom_empty [simp]: "dom empty = {}"
-by(simp add:dom_def)
+ by (simp add: dom_def)
lemma dom_fun_upd [simp]:
- "dom(f(x := y)) = (if y=None then dom f - {x} else insert x (dom f))"
-by(auto simp add:dom_def)
+ "dom(f(x := y)) = (if y = None then dom f - {x} else insert x (dom f))"
+ by (auto simp: dom_def)
lemma dom_if:
"dom (\<lambda>x. if P x then f x else g x) = dom f \<inter> {x. P x} \<union> dom g \<inter> {x. \<not> P x}"
@@ -515,36 +517,36 @@
"dom (map_of xys) = fst ` set xys"
by (induct xys) (auto simp add: dom_if)
-lemma dom_map_of_zip [simp]: "length xs = length ys ==> dom (map_of (zip xs ys)) = set xs"
-by (induct rule: list_induct2) (auto simp add: dom_if)
+lemma dom_map_of_zip [simp]: "length xs = length ys \<Longrightarrow> dom (map_of (zip xs ys)) = set xs"
+ by (induct rule: list_induct2) (auto simp: dom_if)
lemma finite_dom_map_of: "finite (dom (map_of l))"
-by (induct l) (auto simp add: dom_def insert_Collect [symmetric])
+ by (induct l) (auto simp: dom_def insert_Collect [symmetric])
lemma dom_map_upds [simp]:
- "dom(m(xs[\<mapsto>]ys)) = set(take (length ys) xs) Un dom m"
+ "dom(m(xs[\<mapsto>]ys)) = set(take (length ys) xs) \<union> dom m"
apply (induct xs arbitrary: m ys)
apply simp
apply (case_tac ys)
apply auto
done
-lemma dom_map_add [simp]: "dom(m++n) = dom n Un dom m"
-by(auto simp:dom_def)
+lemma dom_map_add [simp]: "dom (m ++ n) = dom n \<union> dom m"
+ by (auto simp: dom_def)
lemma dom_override_on [simp]:
- "dom(override_on f g A) =
- (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}"
-by(auto simp: dom_def override_on_def)
+ "dom (override_on f g A) =
+ (dom f - {a. a \<in> A - dom g}) \<union> {a. a \<in> A \<inter> dom g}"
+ by (auto simp: dom_def override_on_def)
-lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1"
-by (rule ext) (force simp: map_add_def dom_def split: option.split)
+lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1 ++ m2 = m2 ++ m1"
+ by (rule ext) (force simp: map_add_def dom_def split: option.split)
lemma map_add_dom_app_simps:
- "\<lbrakk> m\<in>dom l2 \<rbrakk> \<Longrightarrow> (l1++l2) m = l2 m"
- "\<lbrakk> m\<notin>dom l1 \<rbrakk> \<Longrightarrow> (l1++l2) m = l2 m"
- "\<lbrakk> m\<notin>dom l2 \<rbrakk> \<Longrightarrow> (l1++l2) m = l1 m"
-by (auto simp add: map_add_def split: option.split_asm)
+ "m \<in> dom l2 \<Longrightarrow> (l1 ++ l2) m = l2 m"
+ "m \<notin> dom l1 \<Longrightarrow> (l1 ++ l2) m = l2 m"
+ "m \<notin> dom l2 \<Longrightarrow> (l1 ++ l2) m = l1 m"
+ by (auto simp add: map_add_def split: option.split_asm)
lemma dom_const [simp]:
"dom (\<lambda>x. Some (f x)) = UNIV"
@@ -554,7 +556,7 @@
lemma finite_map_freshness:
"finite (dom (f :: 'a \<rightharpoonup> 'b)) \<Longrightarrow> \<not> finite (UNIV :: 'a set) \<Longrightarrow>
\<exists>x. f x = None"
-by(bestsimp dest:ex_new_if_finite)
+ by (bestsimp dest: ex_new_if_finite)
lemma dom_minus:
"f x = None \<Longrightarrow> dom f - insert x A = dom f - A"
@@ -575,12 +577,14 @@
proof (rule ext)
fix k show "map_of xs k = map_of ys k"
proof (cases "map_of xs k")
- case None then have "k \<notin> set (map fst xs)" by (simp add: map_of_eq_None_iff)
+ case None
+ then have "k \<notin> set (map fst xs)" by (simp add: map_of_eq_None_iff)
with set_eq have "k \<notin> set (map fst ys)" by simp
then have "map_of ys k = None" by (simp add: map_of_eq_None_iff)
with None show ?thesis by simp
next
- case (Some v) then have "k \<in> set (map fst xs)" by (auto simp add: dom_map_of_conv_image_fst [symmetric])
+ case (Some v)
+ then have "k \<in> set (map fst xs)" by (auto simp add: dom_map_of_conv_image_fst [symmetric])
with map_eq show ?thesis by auto
qed
qed
@@ -594,45 +598,48 @@
qed
lemma finite_set_of_finite_maps:
-assumes "finite A" "finite B"
-shows "finite {m. dom m = A \<and> ran m \<subseteq> B}" (is "finite ?S")
+ assumes "finite A" "finite B"
+ shows "finite {m. dom m = A \<and> ran m \<subseteq> B}" (is "finite ?S")
proof -
let ?S' = "{m. \<forall>x. (x \<in> A \<longrightarrow> m x \<in> Some ` B) \<and> (x \<notin> A \<longrightarrow> m x = None)}"
have "?S = ?S'"
proof
- show "?S \<subseteq> ?S'" by(auto simp: dom_def ran_def image_def)
+ show "?S \<subseteq> ?S'" by (auto simp: dom_def ran_def image_def)
show "?S' \<subseteq> ?S"
proof
fix m assume "m \<in> ?S'"
hence 1: "dom m = A" by force
- hence 2: "ran m \<subseteq> B" using \<open>m \<in> ?S'\<close> by(auto simp: dom_def ran_def)
+ hence 2: "ran m \<subseteq> B" using \<open>m \<in> ?S'\<close> by (auto simp: dom_def ran_def)
from 1 2 show "m \<in> ?S" by blast
qed
qed
with assms show ?thesis by(simp add: finite_set_of_finite_funs)
qed
+
subsection \<open>@{term [source] ran}\<close>
-lemma ranI: "m a = Some b ==> b : ran m"
-by(auto simp: ran_def)
+lemma ranI: "m a = Some b \<Longrightarrow> b \<in> ran m"
+ by (auto simp: ran_def)
(* declare ranI [intro]? *)
lemma ran_empty [simp]: "ran empty = {}"
-by(auto simp: ran_def)
+ by (auto simp: ran_def)
-lemma ran_map_upd [simp]: "m a = None ==> ran(m(a\<mapsto>b)) = insert b (ran m)"
-unfolding ran_def
+lemma ran_map_upd [simp]: "m a = None \<Longrightarrow> ran(m(a\<mapsto>b)) = insert b (ran m)"
+ unfolding ran_def
apply auto
-apply (subgoal_tac "aa ~= a")
+apply (subgoal_tac "aa \<noteq> a")
apply auto
done
-lemma ran_distinct:
- assumes dist: "distinct (map fst al)"
+lemma ran_distinct:
+ assumes dist: "distinct (map fst al)"
shows "ran (map_of al) = snd ` set al"
-using assms proof (induct al)
- case Nil then show ?case by simp
+ using assms
+proof (induct al)
+ case Nil
+ then show ?case by simp
next
case (Cons kv al)
then have "ran (map_of al) = snd ` set al" by simp
@@ -642,24 +649,25 @@
qed
lemma ran_map_option: "ran (\<lambda>x. map_option f (m x)) = f ` ran m"
-by(auto simp add: ran_def)
+ by (auto simp add: ran_def)
+
subsection \<open>@{text "map_le"}\<close>
lemma map_le_empty [simp]: "empty \<subseteq>\<^sub>m g"
-by (simp add: map_le_def)
+ by (simp add: map_le_def)
lemma upd_None_map_le [simp]: "f(x := None) \<subseteq>\<^sub>m f"
-by (force simp add: map_le_def)
+ by (force simp add: map_le_def)
lemma map_le_upd[simp]: "f \<subseteq>\<^sub>m g ==> f(a := b) \<subseteq>\<^sub>m g(a := b)"
-by (fastforce simp add: map_le_def)
+ by (fastforce simp add: map_le_def)
lemma map_le_imp_upd_le [simp]: "m1 \<subseteq>\<^sub>m m2 \<Longrightarrow> m1(x := None) \<subseteq>\<^sub>m m2(x \<mapsto> y)"
-by (force simp add: map_le_def)
+ by (force simp add: map_le_def)
lemma map_le_upds [simp]:
- "f \<subseteq>\<^sub>m g ==> f(as [\<mapsto>] bs) \<subseteq>\<^sub>m g(as [\<mapsto>] bs)"
+ "f \<subseteq>\<^sub>m g \<Longrightarrow> f(as [\<mapsto>] bs) \<subseteq>\<^sub>m g(as [\<mapsto>] bs)"
apply (induct as arbitrary: f g bs)
apply simp
apply (case_tac bs)
@@ -667,13 +675,13 @@
done
lemma map_le_implies_dom_le: "(f \<subseteq>\<^sub>m g) \<Longrightarrow> (dom f \<subseteq> dom g)"
-by (fastforce simp add: map_le_def dom_def)
+ by (fastforce simp add: map_le_def dom_def)
lemma map_le_refl [simp]: "f \<subseteq>\<^sub>m f"
-by (simp add: map_le_def)
+ by (simp add: map_le_def)
lemma map_le_trans[trans]: "\<lbrakk> m1 \<subseteq>\<^sub>m m2; m2 \<subseteq>\<^sub>m m3\<rbrakk> \<Longrightarrow> m1 \<subseteq>\<^sub>m m3"
-by (auto simp add: map_le_def dom_def)
+ by (auto simp add: map_le_def dom_def)
lemma map_le_antisym: "\<lbrakk> f \<subseteq>\<^sub>m g; g \<subseteq>\<^sub>m f \<rbrakk> \<Longrightarrow> f = g"
unfolding map_le_def
@@ -682,17 +690,17 @@
apply (case_tac "x \<in> dom g", simp, fastforce)
done
-lemma map_le_map_add [simp]: "f \<subseteq>\<^sub>m (g ++ f)"
-by (fastforce simp add: map_le_def)
+lemma map_le_map_add [simp]: "f \<subseteq>\<^sub>m g ++ f"
+ by (fastforce simp: map_le_def)
-lemma map_le_iff_map_add_commute: "(f \<subseteq>\<^sub>m f ++ g) = (f++g = g++f)"
-by(fastforce simp: map_add_def map_le_def fun_eq_iff split: option.splits)
+lemma map_le_iff_map_add_commute: "f \<subseteq>\<^sub>m f ++ g \<longleftrightarrow> f ++ g = g ++ f"
+ by (fastforce simp: map_add_def map_le_def fun_eq_iff split: option.splits)
-lemma map_add_le_mapE: "f++g \<subseteq>\<^sub>m h \<Longrightarrow> g \<subseteq>\<^sub>m h"
-by (fastforce simp add: map_le_def map_add_def dom_def)
+lemma map_add_le_mapE: "f ++ g \<subseteq>\<^sub>m h \<Longrightarrow> g \<subseteq>\<^sub>m h"
+ by (fastforce simp: map_le_def map_add_def dom_def)
-lemma map_add_le_mapI: "\<lbrakk> f \<subseteq>\<^sub>m h; g \<subseteq>\<^sub>m h \<rbrakk> \<Longrightarrow> f++g \<subseteq>\<^sub>m h"
-by (clarsimp simp add: map_le_def map_add_def dom_def split: option.splits)
+lemma map_add_le_mapI: "\<lbrakk> f \<subseteq>\<^sub>m h; g \<subseteq>\<^sub>m h \<rbrakk> \<Longrightarrow> f ++ g \<subseteq>\<^sub>m h"
+ by (auto simp: map_le_def map_add_def dom_def split: option.splits)
lemma dom_eq_singleton_conv: "dom f = {x} \<longleftrightarrow> (\<exists>v. f = [x \<mapsto> v])"
proof(rule iffI)
@@ -714,8 +722,10 @@
lemma set_map_of_compr:
assumes distinct: "distinct (map fst xs)"
shows "set xs = {(k, v). map_of xs k = Some v}"
-using assms proof (induct xs)
- case Nil then show ?case by simp
+ using assms
+proof (induct xs)
+ case Nil
+ then show ?case by simp
next
case (Cons x xs)
obtain k v where "x = (k, v)" by (cases x) blast
@@ -742,7 +752,8 @@
assume ?rhs show ?lhs
proof
fix k
- show "map_of xs k = map_of ys k" proof (cases "map_of xs k")
+ show "map_of xs k = map_of ys k"
+ proof (cases "map_of xs k")
case None
with \<open>?rhs\<close> have "map_of ys k = None"
by (simp add: map_of_eq_None_iff)