Theory AuxLemmas

Up to index of Isabelle/HOL/MicroJava

theory AuxLemmas
imports JBasis
begin

(*  Title:      HOL/MicroJava/Comp/AuxLemmas.thy
    ID:         $Id: AuxLemmas.thy,v 1.5 2005/04/27 14:39:44 paulson Exp $
    Author:     Martin Strecker
*)


(* Auxiliary Lemmas *)

theory AuxLemmas
imports "../J/JBasis"
begin

(**********************************************************************)
(* List.thy *)
(**********************************************************************)



lemma app_nth_greater_len [rule_format (no_asm), simp]:
 "∀ ind. length pre ≤ ind --> (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind"
apply (induct "pre")
apply auto
apply (case_tac ind)
apply auto
done

lemma length_takeWhile: "v ∈ set xs ==> length (takeWhile (%z. z~=v) xs) < length xs"
by (induct xs, auto)

lemma nth_length_takeWhile [simp]: 
  "v ∈ set xs ==> xs ! (length (takeWhile (%z. z~=v) xs)) = v"
by (induct xs, auto)


lemma map_list_update [simp]:
  "[| x ∈ set xs; distinct xs|] ==> 
  (map f xs) [length (takeWhile (λz. z ≠ x) xs) := v] = 
  map (f(x:=v)) xs"
apply (induct xs)
apply simp
apply (case_tac "x=a")
apply auto
done


(**********************************************************************)
(* Product_Type.thy *)
(**********************************************************************)


lemma split_compose: "(split f) o (λ (a,b). ((fa a), (fb b))) =
  (λ (a,b). (f (fa a) (fb b)))"
by (simp add: split_def o_def)

lemma split_iter: "(λ (a,b,c). ((g1 a), (g2 b), (g3 c))) = 
        (λ (a,p). ((g1 a), (λ (b, c). ((g2 b), (g3 c))) p))"
by (simp add: split_def o_def)


(**********************************************************************)
(* Set.thy *)
(**********************************************************************)

lemma singleton_in_set: "A = {a} ==> a ∈ A" by simp

(**********************************************************************)
(* Map.thy *)
(**********************************************************************)

lemma the_map_upd: "(the o f(x\<mapsto>v)) = (the o f)(x:=v)"
by (simp add: expand_fun_eq)

lemma map_of_in_set: 
  "(map_of xs x = None) = (x ∉ set (map fst xs))"
by (induct xs, auto)

lemma map_map_upd [simp]: 
  "y ∉ set xs ==> map (the o f(y\<mapsto>v)) xs = map (the o 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 o f(ys[\<mapsto>]vs)) xs = map (the o f) xs"
apply (induct xs)
 apply simp
apply fastsimp
done

lemma map_upds_distinct [rule_format (no_asm), simp]: 
  "∀ f vs. length ys = length vs --> distinct ys --> map (the o f(ys[\<mapsto>]vs)) ys = vs"
apply (induct ys)
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
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 [\<mapsto>] map g zs)"
by (induct zs, auto)

(* In analogy to Map.map_of_SomeD *)
lemma map_upds_SomeD [rule_format (no_asm)]: 
  "∀ m ys. (m(xs[\<mapsto>]ys)) k = Some y --> k ∈ (set xs) ∨ (m k = Some y)"
apply (induct xs)
apply simp
apply auto
apply(case_tac ys)
apply auto
done

lemma map_of_upds_SomeD: "(map_of m (xs[\<mapsto>]ys)) k = Some y 
  ==> k ∈ (set (xs @ map fst m))"
by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma)


lemma map_of_map_prop [rule_format (no_asm)]: 
  "(map_of (map f xs) k = Some v) --> 
  (∀ x ∈ set xs. (P1 x)) --> 
  (∀ x. (P1 x) --> (P2 (f x))) -->
  (P2(k, v))"
by (induct xs,auto)

lemma map_of_map2: "∀ x ∈ set xs. (fst (f x)) = (fst x) ==>
  map_of (map f xs) a = option_map (λ b. (snd (f (a, b)))) (map_of xs a)"
by (induct xs, auto)

lemma option_map_of [simp]: "(option_map f (map_of xs k) = None) = ((map_of xs k) = None)"
by (simp add: option_map_def split: option.split)



end

lemma app_nth_greater_len:

  length pre  ind ==> (pre @ a # post) ! Suc ind = (pre @ post) ! ind

lemma length_takeWhile:

  v ∈ set xs ==> length (takeWhile (λz. z  v) xs) < length xs

lemma nth_length_takeWhile:

  v ∈ set xs ==> xs ! length (takeWhile (λz. z  v) xs) = v

lemma map_list_update:

  [| x ∈ set xs; distinct xs |]
  ==> map f xs[length (takeWhile (λz. z  x) xs) := v] = map (f(x := v)) xs

lemma split_compose:

  (λ(x, y). f x y) o (λ(a, b). (fa a, fb b)) = (λ(a, b). f (fa a) (fb b))

lemma split_iter:

  (λ(a, b, c). (g1.0 a, g2.0 b, g3.0 c)) =
  (λ(a, p). (g1.0 a, (λ(b, c). (g2.0 b, g3.0 c)) p))

lemma singleton_in_set:

  A = {a} ==> aA

lemma the_map_upd:

  the o f(x |-> v) = (the o f)(x := v)

lemma map_of_in_set:

  (map_of xs x = None) = (x  set (map fst xs))

lemma map_map_upd:

  y  set xs ==> map (the o f(y |-> v)) xs = map (the o f) xs

lemma map_map_upds:

  y∈set ys. y  set xs ==> map (the o f(ys [|->] vs)) xs = map (the o f) xs

lemma map_upds_distinct:

  [| length ys = length vs; distinct ys |] ==> map (the o f(ys [|->] vs)) ys = vs

lemma map_of_map_as_map_upd:

  distinct (map f zs)
  ==> map_of (map (λp. (f p, g p)) zs) = [map f zs [|->] map g zs]

lemma map_upds_SomeD:

  (m(xs [|->] ys)) k = Some y ==> k ∈ set xsm k = Some y

lemma map_of_upds_SomeD:

  (map_of m(xs [|->] ys)) k = Some y ==> k ∈ set (xs @ map fst m)

lemma map_of_map_prop:

  [| map_of (map f xs) k = Some v; Ball (set xs) P1.0;
     ∀x. P1.0 x --> P2.0 (f x) |]
  ==> P2.0 (k, v)

lemma map_of_map2:

  x∈set xs. fst (f x) = fst x
  ==> map_of (map f xs) a = option_mapb. snd (f (a, b))) (map_of xs a)

lemma option_map_of:

  (option_map f (map_of xs k) = None) = (map_of xs k = None)