diff -r 8c83139a1433 -r b85bfa89a387 src/HOL/MicroJava/Comp/AuxLemmas.thy --- a/src/HOL/MicroJava/Comp/AuxLemmas.thy Mon Oct 04 14:46:48 2010 +0200 +++ b/src/HOL/MicroJava/Comp/AuxLemmas.thy Mon Oct 04 14:46:48 2010 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/Comp/AuxLemmas.thy - ID: $Id$ Author: Martin Strecker *) @@ -78,30 +77,24 @@ "y \ set xs \ map (the \ f(y\v)) xs = map (the \ f) xs" by (simp add: the_map_upd) -lemma map_map_upds [rule_format (no_asm), simp]: -"\ f vs. (\y\set ys. y \ set xs) \ map (the \ f(ys[\]vs)) xs = map (the \ f) xs" -apply (induct xs) +lemma map_map_upds [simp]: + "(\y\set ys. y \ set xs) \ map (the \ f(ys[\]vs)) xs = map (the \ f) xs" +apply (induct xs arbitrary: f vs) apply simp apply fastsimp done -lemma map_upds_distinct [rule_format (no_asm), simp]: - "\ f vs. length ys = length vs \ distinct ys \ map (the \ f(ys[\]vs)) ys = vs" -apply (induct ys) +lemma map_upds_distinct [simp]: + "distinct ys \ length ys = length vs \ map (the \ f(ys[\]vs)) ys = vs" +apply (induct ys arbitrary: f vs rule: distinct.induct) apply simp -apply (intro strip) apply (case_tac vs) -apply simp -apply (simp only: map_upds_Cons distinct.simps hd.simps tl.simps map.simps) -apply clarify -apply (simp del: o_apply) -apply simp +apply simp_all done - -lemma map_of_map_as_map_upd [rule_format (no_asm)]: "distinct (map f zs) \ -map_of (map (\ p. (f p, g p)) zs) = empty (map f zs [\] map g zs)" -by (induct zs, auto) +lemma map_of_map_as_map_upd: + "distinct (map f zs) \ map_of (map (\ p. (f p, g p)) zs) = empty (map f zs [\] map g zs)" + by (induct zs) auto (* In analogy to Map.map_of_SomeD *) lemma map_upds_SomeD [rule_format (no_asm)]: