Theory Correctness1

Up to index of Isabelle/HOL/Jinja

theory Correctness1 = J1WellForm + Compiler1:

(*  Title:      Jinja/Compiler/Correctness1.thy
    ID:         $Id: Correctness1.html 1910 2004-05-19 04:46:04Z kleing $
    Author:     Tobias Nipkow
    Copyright   TUM 2003
*)

header {* \isaheader{Correctness of Stage 1} *}

theory Correctness1 = J1WellForm + Compiler1:

section{*Correctness of program compilation *}

consts
  unmod :: "expr1 => nat => bool"
  unmods :: "expr1 list => nat => bool"
primrec
"unmod (new C) i = True"
"unmod (Cast C e) i = unmod e i"
"unmod (Val v) i = True"
"unmod (e1 «bop» e2) i = (unmod e1 i ∧ unmod e2 i)"
"unmod (Var i) j = True"
"unmod (i:=e) j = (i ≠ j ∧ unmod e j)"
"unmod (e\<bullet>F{D}) i = unmod e i"
"unmod (e1\<bullet>F{D}:=e2) i = (unmod e1 i ∧ unmod e2 i)"
"unmod (e\<bullet>M(es)) i = (unmod e i ∧ unmods es i)"
"unmod {j:T; e} i = unmod e i"
"unmod (e1;;e2) i = (unmod e1 i ∧  unmod e2 i)"
"unmod (if (e) e1 else e2) i = (unmod e i ∧ unmod e1 i ∧ unmod e2 i)"
"unmod (while (e) c) i = (unmod e i ∧ unmod c i)"
"unmod (throw e) i = unmod e i"
"unmod (try e1 catch(C i) e2) j = (unmod e1 j ∧ (if i=j then False else unmod e2 j))"

"unmods ([]) i = True"
"unmods (e#es) i = (unmod e i ∧ unmods es i)"

lemma hidden_unmod: "!!Vs. hidden Vs i ==> unmod (compE1 Vs e) i" and
 "!!Vs. hidden Vs i ==> unmods (compEs1 Vs es) i"
(*<*)
apply(induct e and es)
apply (simp_all add:hidden_inacc)
apply(auto simp add:hidden_def)
done
(*>*)


lemma eval1_preserves_unmod:
  "[| P \<turnstile>1 ⟨e,(h,ls)⟩ => ⟨e',(h',ls')⟩; unmod e i; i < size ls |]
  ==> ls ! i = ls' ! i"
and "[| P \<turnstile>1 ⟨es,(h,ls)⟩ [=>] ⟨es',(h',ls')⟩; unmods es i; i < size ls |]
      ==> ls ! i = ls' ! i"
(*<*)
apply(induct rule:eval1_evals1_induct)
apply(auto dest!:eval1_preserves_len split:split_if_asm)
done
(*>*)


lemma LAss_lem:
  "[|x ∈ set xs; size xs ≤ size ys |]
  ==> m1m m2(xs[\<mapsto>]ys) ==> m1(x\<mapsto>y) ⊆m m2(xs[\<mapsto>]ys[index xs x := y])"
(*<*)
apply(simp add:map_le_def);
apply(simp add:fun_upds_apply index_less_aux eq_sym_conv)
done
(*>*)


lemma Block_lem:
assumes 0: "l ⊆m [Vs [\<mapsto>] ls]"
    and 1: "l' ⊆m [Vs [\<mapsto>] ls', V\<mapsto>v]"
    and hidden: "V ∈ set Vs ==> ls ! index Vs V = ls' ! index Vs V"
    and size: "size ls = size ls'"    "size Vs < size ls'"
shows "l'(V := l V) ⊆m [Vs [\<mapsto>] ls']"
(*<*)
proof -
  have "l'(V := l V) ⊆m [Vs [\<mapsto>] ls', V\<mapsto>v](V := l V)"
    using 1 by(rule map_le_upd)
  also have "… = [Vs [\<mapsto>] ls'](V := l V)" by simp
  also have "… ⊆m [Vs [\<mapsto>] ls']"
  proof (cases "l V")
    case None thus ?thesis by simp
  next
    case (Some w)
    hence "[Vs [\<mapsto>] ls] V = Some w"
      using 0 by(force simp add: map_le_def split:if_splits)
    hence VinVs: "V ∈ set Vs" and w: "w = ls ! index Vs V"
      using size by(auto simp add:fun_upds_apply split:if_splits)
    hence "w = ls' ! index Vs V" using hidden[OF VinVs] by simp
    hence "[Vs [\<mapsto>] ls'](V := l V) = [Vs [\<mapsto>] ls']"
      using Some size VinVs by(simp add:index_less_aux map_upds_upd_conv_index)
    thus ?thesis by simp
  qed
  finally show ?thesis .
qed
(*>*)

(*<*)
declare fun_upd_apply[simp del]
(*>*)


text{*\noindent The main theorem: *}

theorem assumes wf: "wwf_J_prog P"
shows eval1_eval: "P \<turnstile> ⟨e,(h,l)⟩ => ⟨e',(h',l')⟩
  ==> (!!Vs ls. [| fv e ⊆ set Vs;  l ⊆m [Vs[\<mapsto>]ls]; size Vs + max_vars e ≤ size ls |]
       ==> ∃ls'. compP1 P \<turnstile>1 ⟨compE1 Vs e,(h,ls)⟩ => ⟨fin1 e',(h',ls')⟩ ∧ l' ⊆m [Vs[\<mapsto>]ls'])"
(*<*)
  (is "_ ==> (!!Vs ls. PROP ?P e h l e' h' l' Vs ls)"
   is "_ ==> (!!Vs ls. [| _; _; _ |] ==> ∃ls'. ?Post e h l e' h' l' Vs ls ls')")
(*>*)

and evals1_evals: "P \<turnstile> ⟨es,(h,l)⟩ [=>] ⟨es',(h',l')⟩
    ==> (!!Vs ls. [| fvs es ⊆ set Vs;  l ⊆m [Vs[\<mapsto>]ls]; size Vs + max_varss es ≤ size ls |]
         ==> ∃ls'. compP1 P \<turnstile>1 ⟨compEs1 Vs es,(h,ls)⟩ [=>] ⟨compEs1 Vs es',(h',ls')⟩ ∧
                      l' ⊆m [Vs[\<mapsto>]ls'])"
(*<*)
  (is "_ ==> (!!Vs ls. PROP ?Ps es h l es' h' l' Vs ls)"
   is "_ ==> (!!Vs ls. [| _; _; _|] ==> ∃ls'. ?Posts es h l es' h' l' Vs ls ls')")
proof (induct rule:eval_evals_induct)
  case Nil thus ?case by(fastsimp intro!:Nil1)
next
  case (Cons e es es' h l h' l' h2 l2 v)
  have "PROP ?P e h l (Val v) h' l' Vs ls".
  with Cons.prems
  obtain ls' where 1: "?Post e h l (Val v) h' l' Vs ls ls'"
    "size ls = size ls'" by(auto intro!:eval1_preserves_len)
  have "PROP ?Ps es h' l' es' h2 l2 Vs ls'".
  with 1 Cons.prems
  obtain ls2 where 2: "?Posts es h' l' es' h2 l2 Vs ls' ls2" by(auto)
  from 1 2 Cons show ?case by(auto intro!:Cons1)
next
  case (ConsThrow e e' es h l h' l') thus ?case
    by(fastsimp intro!:ConsThrow1 dest: eval_final)
next
  case (Block T V e e' h h' l l')
  let ?Vs = "Vs @ [V]"
  have IH:
  "[|fv e ⊆ set ?Vs; l(V := None) ⊆m [?Vs [\<mapsto>] ls];
    size ?Vs + max_vars e ≤ size ls|]
   ==> ∃ls'. compP1 P \<turnstile>1 ⟨compE1 ?Vs e,(h,ls)⟩ => ⟨fin1 e',(h', ls')⟩ ∧
             l' ⊆m [?Vs [\<mapsto>] ls']" and
  fv: "fv {V:T; e} ⊆ set Vs" and rel: "l ⊆m [Vs [\<mapsto>] ls]" and
  len: "length Vs + max_vars {V:T; e} ≤ length ls" .
  have len': "length Vs < length ls" using len by(auto)
  have "fv e ⊆ set ?Vs" using fv by auto
  moreover have "l(V := None) ⊆m [?Vs [\<mapsto>] ls]" using rel len' by simp
  moreover have "size ?Vs + max_vars e ≤ size ls" using len by simp
  ultimately obtain ls' where
   1: "compP1 P \<turnstile>1 ⟨compE1 ?Vs e,(h,ls)⟩ => ⟨fin1 e',(h',ls')⟩"
   and rel': "l' ⊆m [?Vs [\<mapsto>] ls']" using IH by blast
  have [simp]: "length ls = length ls'" by(rule eval1_preserves_len[OF 1])
  show "∃ls'. compP1 P \<turnstile>1 ⟨compE1 Vs {V:T; e},(h,ls)⟩ => ⟨fin1 e',(h',ls')⟩
              ∧ l'(V := l V) ⊆m [Vs [\<mapsto>] ls']" (is "∃ls'. ?R ls'")
  proof
    show "?R ls'"
    proof
      show "compP1 P \<turnstile>1 ⟨compE1 Vs {V:T; e},(h,ls)⟩ => ⟨fin1 e',(h',ls')⟩"
        using 1 by(simp add:Block1)
    next
      show "l'(V := l V) ⊆m [Vs [\<mapsto>] ls']"
      proof -
        have "l' ⊆m [Vs [\<mapsto>] ls', V \<mapsto> ls' ! length Vs]"
          using len' rel' by simp
        moreover
        { assume VinVs: "V ∈ set Vs"
          hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index)
          hence "unmod (compE1 (Vs @ [V]) e) (index Vs V)"
            by(rule hidden_unmod)
          moreover have "index Vs V < length ls"
            using len' VinVs by(simp add:index_less_aux)
          ultimately have "ls ! index Vs V = ls' ! index Vs V"
            by(rule eval1_preserves_unmod[OF 1])
        }
        ultimately show ?thesis using Block_lem[OF rel] len' by auto
      qed
    qed
  qed
next
  case (TryThrow C D V a e' e2 fs h' l' h l)
  have "PROP ?P e' h l (Throw a) h' l' Vs ls".
  with TryThrow.prems
  obtain ls' where 1: "?Post e' h l (Throw a) h' l' Vs ls ls'"  by(auto)
  show ?case using 1 TryThrow.hyps by(auto intro!:eval1_evals1.TryThrow1)
next
  case (TryCatch C D V a e1 e2 e' fs h1 h2 l1 l2 h l)
  let ?e = "try e1 catch(C V) e2"
  have IH1: "[|fv e1 ⊆ set Vs; l ⊆m [Vs [\<mapsto>] ls];
              size Vs + max_vars e1 ≤ length ls|]
          ==> ∃ls1. compP1 P \<turnstile>1 ⟨compE1 Vs e1,(h,ls)⟩ =>
                                ⟨fin1 (Throw a),(h1,ls1)⟩ ∧
                    l1m [Vs [\<mapsto>] ls1]" and
    fv: "fv ?e ⊆ set Vs" and
    rel: "l ⊆m [Vs [\<mapsto>] ls]" and
    len: "length Vs + max_vars ?e ≤ length ls".
  have "fv e1 ⊆ set Vs" using fv by auto
  moreover have "length Vs + max_vars e1 ≤ length ls" using len by(auto)
  ultimately obtain ls1 where
    1: "compP1 P \<turnstile>1 ⟨compE1 Vs e1,(h,ls)⟩ => ⟨Throw a,(h1,ls1)⟩"
    and rel1: "l1m [Vs [\<mapsto>] ls1]" using IH1 rel by fastsimp
  from 1 have [simp]: "size ls = size ls1" by(rule eval1_preserves_len)
  let ?Vs = "Vs @ [V]" let ?ls = "(ls1[size Vs:=Addr a])"
  have IH2: "[|fv e2 ⊆ set ?Vs; l1(V \<mapsto> Addr a) ⊆m [?Vs [\<mapsto>] ?ls];
              length ?Vs + max_vars e2 ≤ length ?ls|] ==> ∃ls2.
       compP1 P \<turnstile>1 ⟨compE1 ?Vs e2,(h1,?ls)⟩ => ⟨fin1 e',(h2, ls2)⟩ ∧
       l2m [?Vs [\<mapsto>] ls2]".
  have len1: "size Vs < size ls1" using len by(auto)
  have "fv e2 ⊆ set ?Vs" using fv by auto
  moreover have "l1(V \<mapsto> Addr a) ⊆m [?Vs [\<mapsto>] ?ls]" using rel1 len1 by simp
  moreover have "length ?Vs + max_vars e2 ≤ length ?ls" using len by(auto)
  ultimately obtain ls2 where
    2: "compP1 P \<turnstile>1 ⟨compE1 ?Vs e2,(h1,?ls)⟩ => ⟨fin1 e',(h2, ls2)⟩"
    and rel2: "l2m [?Vs [\<mapsto>] ls2]"  using IH2 by blast
  from 2 have [simp]: "size ls1 = size ls2"
    by(fastsimp dest: eval1_preserves_len)
  show "∃ls2. compP1 P \<turnstile>1 ⟨compE1 Vs ?e,(h,ls)⟩ => ⟨fin1 e',(h2,ls2)⟩ ∧
              l2(V := l1 V) ⊆m [Vs [\<mapsto>] ls2]"  (is "∃ls2. ?R ls2")
  proof
    show "?R ls2"
    proof
      have hp: "h1 a = Some (D, fs)".
      have "P \<turnstile> D \<preceq>* C". hence caught: "compP1 P \<turnstile> D \<preceq>* C" by simp
      from TryCatch1[OF 1 _ caught len1 2, OF hp]
      show "compP1 P \<turnstile>1 ⟨compE1 Vs ?e,(h,ls)⟩ => ⟨fin1 e',(h2,ls2)⟩" by simp
    next
      show "l2(V := l1 V) ⊆m [Vs [\<mapsto>] ls2]"
      proof -
        have "l2m [Vs [\<mapsto>] ls2, V \<mapsto> ls2 ! length Vs]"
          using len1 rel2 by simp
        moreover
        { assume VinVs: "V ∈ set Vs"
          hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index)
          hence "unmod (compE1 (Vs @ [V]) e2) (index Vs V)"
            by(rule hidden_unmod)
          moreover have "index Vs V < length ?ls"
            using len1 VinVs by(simp add:index_less_aux)
          ultimately have "?ls ! index Vs V = ls2 ! index Vs V"
            by(rule eval1_preserves_unmod[OF 2])
          moreover have "index Vs V < size Vs" using VinVs by simp
          ultimately have "ls1 ! index Vs V = ls2 ! index Vs V"
            using len1 by(simp del:size_index_conv)
        }
        ultimately show ?thesis using Block_lem[OF rel1] len1  by simp
      qed
    qed
  qed
next
  case Try thus ?case by(fastsimp intro!:Try1)
next
  case Throw thus ?case by(fastsimp intro!:Throw1)
next
  case ThrowNull thus ?case by(fastsimp intro!:ThrowNull1)
next
  case ThrowThrow thus ?case  by(fastsimp intro!:ThrowThrow1)
next
  case (CondT e e' e1 e2 h l h1 l1 h2 l2)
  have "PROP ?P e h l true h1 l1 Vs ls".
  with CondT.prems
  obtain ls1 where 1: "?Post e h l true h1 l1 Vs ls ls1"
    "size ls = size ls1"  by(auto intro!:eval1_preserves_len)
  have "PROP ?P e1 h1 l1 e' h2 l2 Vs ls1".
  with 1 CondT.prems
  obtain ls2 where 2: "?Post e1 h1 l1 e' h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 show ?case by(auto intro!:CondT1)
next
  case (CondF e e' e1 e2 h l h1 l1 h2 l2 Vs ls)
  have "PROP ?P e h l false h1 l1 Vs ls".
  with CondF.prems
  obtain ls1 where 1: "?Post e h l false h1 l1 Vs ls ls1"
    "size ls = size ls1"  by(auto intro!:eval1_preserves_len)
  have "PROP ?P e2 h1 l1 e' h2 l2 Vs ls1".
  with 1 CondF.prems
  obtain ls2 where 2: "?Post e2 h1 l1 e' h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 show ?case by(auto intro!:CondF1)
next
  case CondThrow thus ?case by(fastsimp intro!:CondThrow1)
next
  case (Seq e e1 e' h l h1 l1 h2 l2 v)
  have "PROP ?P e h l (Val v) h1 l1 Vs ls".
  with Seq.prems
  obtain ls1 where 1: "?Post e h l (Val v) h1 l1 Vs ls ls1"
    "size ls = size ls1"  by(auto intro!:eval1_preserves_len)
  have "PROP ?P e1 h1 l1 e' h2 l2 Vs ls1".
  with 1 Seq.prems
  obtain ls2 where 2: "?Post e1 h1 l1 e' h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 Seq show ?case by(auto intro!:Seq1)
next
  case SeqThrow thus ?case by(fastsimp intro!:SeqThrow1)
next
  case WhileF thus ?case by(fastsimp intro!:eval1_evals1.intros)
next
  case (WhileT c e e' h l h1 l1 h2 l2 h3 l3 v)
  have "PROP ?P e h l true h1 l1 Vs ls".
  with WhileT.prems
  obtain ls1 where 1: "?Post e h l true h1 l1 Vs ls ls1"
    "size ls = size ls1"   by(auto intro!:eval1_preserves_len)
  have "PROP ?P c h1 l1 (Val v) h2 l2 Vs ls1".
  with 1 WhileT.prems
  obtain ls2 where 2: "?Post c h1 l1 (Val v) h2 l2 Vs ls1 ls2"
    "size ls1 = size ls2"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P (While (e) c) h2 l2 e' h3 l3 Vs ls2".
  with 1 2 WhileT.prems
  obtain ls3 where 3: "?Post (While (e) c) h2 l2 e' h3 l3 Vs ls2 ls3" by(auto)
  from 1 2 3 show ?case by(auto intro!:WhileT1)
next
  case (WhileBodyThrow c e e' h l h1 l1 h2 l2)
  have "PROP ?P e h l true h1 l1 Vs ls".
  with WhileBodyThrow.prems
  obtain ls1 where 1: "?Post e h l true h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P c h1 l1 (throw e') h2 l2 Vs ls1".
  with 1 WhileBodyThrow.prems
  obtain ls2 where 2: "?Post c h1 l1 (throw e') h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 show ?case by(auto intro!:WhileBodyThrow1)
next
  case WhileCondThrow thus ?case by(fastsimp intro!:WhileCondThrow1)
next
  case New thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case NewFail thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case Cast thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case CastNull thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case CastThrow thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case (CastFail C D a e fs h1 l1 h l)
  have "PROP ?P e h l (addr a) h1 l1 Vs ls".
  with CastFail.prems
  obtain ls1 where 1: "?Post e h l (addr a) h1 l1 Vs ls ls1" by auto
  show ?case using 1 CastFail.hyps
    by(auto intro!:CastFail1[where D=D])
next
  case Val thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case (BinOp bop e e1 h l h1 l1 h2 l2 v v1 v2)
  have "PROP ?P e h l (Val v1) h1 l1 Vs ls".
  with BinOp.prems
  obtain ls1 where 1: "?Post e h l (Val v1) h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P e1 h1 l1 (Val v2) h2 l2 Vs ls1".
  with 1 BinOp.prems
  obtain ls2 where 2: "?Post e1 h1 l1 (Val v2) h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 BinOp show ?case by(auto intro!:BinOp1)
next
  case (BinOpThrow2 bop e e0 e1 h l h1 l1 h2 l2 v1)
  have "PROP ?P e0 h l (Val v1) h1 l1 Vs ls".
  with BinOpThrow2.prems
  obtain ls1 where 1: "?Post e0 h l (Val v1) h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P e1 h1 l1 (throw e) h2 l2 Vs ls1".
  with 1 BinOpThrow2.prems
  obtain ls2 where 2: "?Post e1 h1 l1 (throw e) h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow21)
next
  case BinOpThrow1 thus ?case  by(fastsimp intro!:eval1_evals1.intros)
next
  case Var thus ?case
    by(force intro!:Var1 simp add:index_less_aux map_le_def fun_upds_apply)
next
  case LAss thus ?case
    by(fastsimp simp add: index_less_aux LAss_lem intro:eval1_evals1.intros
                dest:eval1_preserves_len)
next
  case LAssThrow thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case FAcc thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case FAccNull thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case FAccThrow thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case (FAss C D F a e1 e2 fs fs' h2 h2' l2 h l h1 l1 v)
  have "PROP ?P e1 h l (addr a) h1 l1 Vs ls".
  with FAss.prems
  obtain ls1 where 1: "?Post e1 h l (addr a) h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P e2 h1 l1 (Val v) h2 l2 Vs ls1".
  with 1 FAss.prems
  obtain ls2 where 2: "?Post e2 h1 l1 (Val v) h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 FAss show ?case by(auto intro!:FAss1)
next
  case (FAssNull D F e1 e2 h l h1 l1 h2 l2 v)
  have "PROP ?P e1 h l null h1 l1 Vs ls".
  with FAssNull.prems
  obtain ls1 where 1: "?Post e1 h l null h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?P e2 h1 l1 (Val v) h2 l2 Vs ls1".
  with 1 FAssNull.prems
  obtain ls2 where 2: "?Post e2 h1 l1 (Val v) h2 l2 Vs ls1 ls2" by(auto)
  from 1 2 FAssNull show ?case by(auto intro!:FAssNull1)
next
  case FAssThrow1 thus ?case by(fastsimp intro:eval1_evals1.intros)
next
  case (FAssThrow2 D F e e1 e2 h l h1 l1 h2 l2 v)
  have "PROP ?P e1 h l (Val v) h1 l1 Vs ls".
  with FAssThrow2.prems
  obtain ls1 where 1: "?Post e1 h l (Val v) h1 l1 Vs ls ls1"
    "size ls = size ls1"   by(auto intro!:eval1_preserves_len)
  have "PROP ?P e2 h1 l1 (throw e) h2 l2 Vs ls1".
  with 1 FAssThrow2.prems
  obtain ls2 where 2: "?Post e2 h1 l1 (throw e) h2 l2 Vs ls1 ls2"  by(auto)
  from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow21)
next
  case (CallNull M e es h l h1 l1 h2 l2 vs)
  have "PROP ?P e h l null h1 l1 Vs ls".
  with CallNull.prems
  obtain ls1 where 1: "?Post e h l null h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?Ps es h1 l1 (map Val vs) h2 l2 Vs ls1".
  with 1 CallNull.prems
  obtain ls2 where 2: "?Posts es h1 l1 (map Val vs) h2 l2 Vs ls1 ls2" by(auto)
  from 1 2 CallNull show ?case
    by (auto simp add: map_compose[symmetric] comp_def elim!: CallNull1)
next
  case CallObjThrow thus ?case  by(fastsimp intro:eval1_evals1.intros)
next
  case (CallParamsThrow M e es es' ex h l h1 l1 h2 l2 v vs)
  have "PROP ?P e h l (Val v) h1 l1 Vs ls".
  with CallParamsThrow.prems
  obtain ls1 where 1: "?Post e h l (Val v) h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?Ps es h1 l1 (map Val vs @ throw ex # es') h2 l2 Vs ls1".
  with 1 CallParamsThrow.prems
  obtain ls2 where 2: "?Posts es h1 l1 (map Val vs @ throw ex # es') h2 l2 Vs ls1 ls2" by(auto)
  from 1 2 CallParamsThrow show ?case
    by (auto simp add: map_compose[symmetric] comp_def
             elim!: CallParamsThrow1 dest!:evals_final)
next
  case (Call C D M T Ts a body e b' fs h2 h3 l2 l2' l3 pns es h l h1 l1 vs)
  have "PROP ?P e h l (addr a) h1 l1 Vs ls".
  with Call.prems
  obtain ls1 where 1: "?Post e h l (addr a) h1 l1 Vs ls ls1"
    "size ls = size ls1"    by(auto intro!:eval1_preserves_len)
  have "PROP ?Ps es h1 l1 (map Val vs) h2 l2 Vs ls1".
  with 1 Call.prems
  obtain ls2 where 2: "?Posts es h1 l1 (map Val vs) h2 l2 Vs ls1 ls2"
    "size ls1 = size ls2"    by(auto intro!:evals1_preserves_len)
  let ?Vs = "this#pns"
  let ?ls = "Addr a # vs @ replicate (max_vars body) arbitrary"
  have mdecl: "P \<turnstile> C sees M: Ts->T = (pns, body) in D".
  have fv_body: "fv body ⊆ set ?Vs" and wf_size: "size Ts = size pns"
    using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
  have mdecl1: "compP1 P \<turnstile> C sees M: Ts->T = (compE1 ?Vs body) in D"
    using sees_method_compP[OF mdecl, of "λ(pns,e). compE1 (this#pns) e"]
    by(simp)
  have [simp]: "l2' = [this \<mapsto> Addr a, pns [\<mapsto>] vs]".
  have Call_size: "size vs = size pns".
  have "PROP ?P body h2 l2' b' h3 l3 ?Vs ?ls".
  with 1 2 fv_body Call_size Call.prems
  obtain ls3 where 3: "?Post body h2 l2' b' h3 l3 ?Vs ?ls ls3"  by(auto)
  have hp: "h2 a = Some (C, fs)".
  from 1 2 3 hp mdecl1 wf_size Call_size show ?case
    by(fastsimp simp add: map_compose[symmetric] comp_def
                intro!: Call1 dest!:evals_final)
qed
(*>*)


subsection{*Preservation of well-formedness*}

text{* The compiler preserves well-formedness. Is less trivial than it
may appear. We start with two simple properties: preservation of
well-typedness *}

lemma compE1_pres_wt: "!!Vs Ts U.
  [| P,[Vs[\<mapsto>]Ts] \<turnstile> e :: U; size Ts = size Vs |]
  ==> compP f P,Ts \<turnstile>1 compE1 Vs e :: U"
and  "!!Vs Ts Us.
  [| P,[Vs[\<mapsto>]Ts] \<turnstile> es [::] Us; size Ts = size Vs |]
  ==> compP f P,Ts \<turnstile>1 compEs1 Vs es [::] Us"
(*<*)
apply(induct e and es)
apply clarsimp
apply(fastsimp)
apply clarsimp
apply(fastsimp split:bop.splits)
apply (fastsimp simp:map_upds_apply_eq_Some split:split_if_asm)
apply (fastsimp simp:map_upds_apply_eq_Some split:split_if_asm)
apply (fastsimp)
apply (fastsimp)
apply (fastsimp dest!: sees_method_compP[where f = f])
apply (fastsimp simp:nth_append)
apply (fastsimp)
apply (fastsimp)
apply (fastsimp)
apply (fastsimp)
apply (fastsimp simp:nth_append)
apply simp
apply (fastsimp)
done
(*>*)

text{*\noindent and the correct block numbering: *}

lemma \<B>: "!!Vs n. size Vs = n ==> \<B> (compE1 Vs e) n"
and \<B>s: "!!Vs n. size Vs = n ==> \<B>s (compEs1 Vs es) n"
(*<*)by(induct e and es) simp_all(*>*)

text{* The main complication is preservation of definite assignment
@{term"\<D>"}. *}

lemma image_index: "A ⊆ set(xs@[x]) ==> index (xs @ [x]) ` A =
  (if x ∈ A then insert (size xs) (index xs ` (A-{x})) else index xs ` A)"
(*<*)
apply(auto simp:image_def)
   apply(rule bexI)
    prefer 2 apply blast
   apply simp
  apply(rule ccontr)
  apply(erule_tac x=xa in ballE)
   prefer 2 apply blast
  apply(fastsimp simp add:neq_commute)
 apply(subgoal_tac "x ≠ xa")
  prefer 2 apply blast
 apply(fastsimp simp add:neq_commute)
apply(subgoal_tac "x ≠ xa")
 prefer 2 apply blast
apply(force)
done
(*>*)


lemma A_compE1_None[simp]:
      "!!Vs. \<A> e = None ==> \<A> (compE1 Vs e) = None"
and "!!Vs. \<A>s es = None ==> \<A>s (compEs1 Vs es) = None"
(*<*)by(induct e and es)(auto simp:hyperset_defs)(*>*)


lemma A_compE1:
      "!!A Vs. [| \<A> e = ⌊A⌋; fv e ⊆ set Vs |] ==> \<A> (compE1 Vs e) = ⌊index Vs ` A⌋"
and "!!A Vs. [| \<A>s es = ⌊A⌋; fvs es ⊆ set Vs |] ==> \<A>s (compEs1 Vs es) = ⌊index Vs ` A⌋"
(*<*)
proof(induct e and es)
  case (Block V' T e)
  hence "fv e ⊆ set (Vs@[V'])" by fastsimp
  moreover obtain B where "\<A> e = ⌊B⌋"
    using Block.prems by(simp add: hyperset_defs)
  moreover from calculation have "B ⊆ set (Vs@[V'])" by(auto dest!:A_fv)
  ultimately show ?case using Block
    by(auto simp add: hyperset_defs image_index)
next
  case (TryCatch e1 C V' e2)
  hence fve2: "fv e2 ⊆ set (Vs@[V'])" by auto
  show ?case
  proof (cases "\<A> e1")
    assume A1: "\<A> e1 = None"
    then obtain A2 where A2: "\<A> e2 = ⌊A2⌋" using TryCatch
      by(simp add:hyperset_defs)
    hence "A2 ⊆ set (Vs@[V'])" using TryCatch.prems A_fv[OF A2] by simp blast
    thus ?thesis using TryCatch fve2 A1 A2
      by(auto simp add:hyperset_defs image_index)
  next
    fix A1 assume A1: "\<A> e1 =  ⌊A1⌋"
    show ?thesis
    proof (cases  "\<A> e2")
      assume A2: "\<A> e2 = None"
      then show ?case using TryCatch A1 by(simp add:hyperset_defs)
    next
      fix A2 assume A2: "\<A> e2 = ⌊A2⌋"
      have "A1 ⊆ set Vs" using TryCatch.prems A_fv[OF A1] by simp blast
      moreover
      have "A2 ⊆ set (Vs@[V'])" using TryCatch.prems A_fv[OF A2] by simp blast
      ultimately show ?thesis using TryCatch A1 A2
        by(fastsimp simp add: hyperset_defs image_index
          Diff_subset_conv inj_on_image_Int[OF inj_on_index])
    qed
  qed
next
  case (Cond e e1 e2)
  { assume "\<A> e = None ∨ \<A> e1 = None ∨ \<A> e2 = None"
    hence ?case using Cond by(auto simp add:hyperset_defs image_Un)
  }
  moreover
  { fix A A1 A2
    assume "\<A> e = ⌊A⌋" and A1: "\<A> e1 = ⌊A1⌋" and A2: "\<A> e2 = ⌊A2⌋"
    moreover
    have "A1 ⊆ set Vs" using Cond.prems A_fv[OF A1] by simp blast
    moreover
    have "A2 ⊆ set Vs" using Cond.prems A_fv[OF A2] by simp blast
    ultimately have ?case using Cond
      by(auto simp add:hyperset_defs image_Un
          inj_on_image_Int[OF inj_on_index])
  }
  ultimately show ?case by fastsimp
qed (auto simp add:hyperset_defs)
(*>*)


lemma D_None[iff]: "\<D> (e::'a exp) None" and [iff]: "\<D>s (es::'a exp list) None"
(*<*)by(induct e and es)(simp_all)(*>*)


lemma D_index_compE1:
      "!!A Vs. [| A ⊆ set Vs; fv e ⊆ set Vs |] ==>
                \<D> e ⌊A⌋ ==> \<D> (compE1 Vs e) ⌊index Vs ` A⌋"
and "!!A Vs. [| A ⊆ set Vs; fvs es ⊆ set Vs |] ==>
                \<D>s es ⌊A⌋ ==> \<D>s (compEs1 Vs es) ⌊index Vs ` A⌋"
(*<*)
proof(induct e and es)
  case (BinOp e1 bop e2)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using BinOp by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] BinOp.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using BinOp.prems A_fv[OF Some] by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` (A ∪ A1)⌋" using BinOp Some by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
next
  case (FAss e1 F D e2)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using FAss by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] FAss.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using FAss.prems A_fv[OF Some] by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` (A ∪ A1)⌋" using FAss Some by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
next
  case (Call e1 M es)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using Call by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] Call.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using Call.prems A_fv[OF Some] by auto
    hence "\<D>s (compEs1 Vs es) ⌊index Vs ` (A ∪ A1)⌋" using Call Some by auto
    hence "\<D>s (compEs1 Vs es) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
next
  case (TryCatch e1 C V e2)
  have "[| A∪{V} ⊆ set(Vs@[V]); fv e2 ⊆ set(Vs@[V]); \<D> e2 ⌊A∪{V}⌋|] ==>
        \<D> (compE1 (Vs@[V]) e2) ⌊index (Vs@[V]) ` (A∪{V})⌋".
  hence "\<D> (compE1 (Vs@[V]) e2) ⌊index (Vs@[V]) ` (A∪{V})⌋"
    using TryCatch.prems by(simp add:Diff_subset_conv)
  moreover have "index (Vs@[V]) ` A ⊆ index Vs ` A ∪ {size Vs}"
    using TryCatch.prems by(auto simp add: image_index split:split_if_asm)
  ultimately show ?case using TryCatch by(auto simp:hyperset_defs elim!:D_mono')
next
  case (Seq e1 e2)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using Seq by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] Seq.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using Seq.prems A_fv[OF Some] by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` (A ∪ A1)⌋" using Seq Some by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
next
  case (Cond e e1 e2)
  hence IH1: "\<D> (compE1 Vs e) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e")
    case None thus ?thesis using Cond by simp
  next
    case (Some B)
    have indexB: "\<A> (compE1 Vs e) = ⌊index Vs ` B⌋"
      using A_compE1[OF Some] Cond.prems  by auto
    have "A ∪ B ⊆ set Vs" using Cond.prems A_fv[OF Some] by auto
    hence "\<D> (compE1 Vs e1) ⌊index Vs ` (A ∪ B)⌋"
      and "\<D> (compE1 Vs e2) ⌊index Vs ` (A ∪ B)⌋"
      using Cond Some by auto
    hence "\<D> (compE1 Vs e1) ⌊index Vs ` A ∪ index Vs ` B⌋"
      and "\<D> (compE1 Vs e2) ⌊index Vs ` A ∪ index Vs ` B⌋"
      by(simp add: image_Un)+
    thus ?thesis using IH1 indexB by auto
  qed
next
  case (While e1 e2)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using While by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] While.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using While.prems A_fv[OF Some] by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` (A ∪ A1)⌋" using While Some by auto
    hence "\<D> (compE1 Vs e2) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
next
  case (Block V T e)
  have "[| fv e ⊆ set(Vs@[V]); A-{V} ⊆ set(Vs@[V]); \<D> e ⌊A-{V}⌋ |] ==>
        \<D> (compE1 (Vs@[V]) e) ⌊index (Vs@[V]) ` (A-{V})⌋".
  hence "\<D> (compE1 (Vs@[V]) e) ⌊index (Vs@[V]) ` (A-{V})⌋"
    using Block.prems by(simp add:Diff_subset_conv)
  moreover have "size Vs ∉ index Vs ` A"
    using Block.prems by(auto simp add:image_def)
  ultimately show ?case using Block
    by(auto simp add: image_index Diff_subset_conv hyperset_defs elim!: D_mono')
next
  case (Cons_exp e1 es)
  hence IH1: "\<D> (compE1 Vs e1) ⌊index Vs ` A⌋" by simp
  show ?case
  proof (cases "\<A> e1")
    case None thus ?thesis using Cons_exp by simp
  next
    case (Some A1)
    have indexA1: "\<A> (compE1 Vs e1) = ⌊index Vs ` A1⌋"
      using A_compE1[OF Some] Cons_exp.prems  by auto
    have "A ∪ A1 ⊆ set Vs" using Cons_exp.prems A_fv[OF Some] by auto
    hence "\<D>s (compEs1 Vs es) ⌊index Vs ` (A ∪ A1)⌋" using Cons_exp Some by auto
    hence "\<D>s (compEs1 Vs es) ⌊index Vs ` A ∪ index Vs ` A1⌋"
      by(simp add: image_Un)
    thus ?thesis using IH1 indexA1 by auto
  qed
qed (simp_all add:hyperset_defs)
(*>*)


lemma index_image_set: "distinct xs ==> index xs ` set xs = {..size xs(}"
(*<*)by(induct xs rule:rev_induct) (auto simp add: image_index)(*>*)


lemma D_compE1:
  "[| \<D> e ⌊set Vs⌋; fv e ⊆ set Vs; distinct Vs |] ==> \<D> (compE1 Vs e) ⌊{..length Vs(}⌋"
(*<*)by(fastsimp dest!: D_index_compE1[OF subset_refl] simp add:index_image_set)(*>*)


lemma D_compE1':
assumes "\<D> e ⌊set(V#Vs)⌋" and "fv e ⊆ set(V#Vs)" and "distinct(V#Vs)"
shows "\<D> (compE1 (V#Vs) e) ⌊{..length Vs}⌋"
(*<*)
proof -
  have "{..size Vs} = {..size(V#Vs)(}" by auto
  thus ?thesis using prems by (simp only:)(rule D_compE1)
qed
(*>*)


lemma compP1_pres_wf: "wf_J_prog P ==> wf_J1_prog (compP1 P)"
(*<*)
apply simp
apply(rule wf_prog_compPI)
 prefer 2 apply assumption
apply(case_tac m)
apply(simp add:wf_mdecl_def wf_J1_mdecl_def wf_J_mdecl)
apply(clarify)
apply(frule WT_fv)
apply(fastsimp intro!: compE1_pres_wt D_compE1' \<B>)
done
(*>*)


end

Correctness of program compilation

lemma hidden_unmod:

  hidden Vs i ==> unmod (compE1 Vs e) i

and

  hidden Vs i ==> unmods (compEs1 Vs es) i

lemma eval1_preserves_unmod:

  [| eval1 P e (h, ls) e' (h', ls'); unmod e i; i < length ls |]
  ==> ls ! i = ls' ! i

and

  [| evals1 P es (h, ls) es' (h', ls'); unmods es i; i < length ls |]
  ==> ls ! i = ls' ! i

lemma LAss_lem:

  [| x : set xs; length xs <= length ys; m1m m2(xs [|->] ys) |]
  ==> m1(x |-> y) ⊆m m2(xs [|->] ys[index xs x := y])

lemma

  [| lm [Vs [|->] ls]; l'm [Vs [|->] ls', V |-> v];
     V : set Vs ==> ls ! index Vs V = ls' ! index Vs V; length ls = length ls';
     length Vs < length ls' |]
  ==> l'(V := l V) ⊆m [Vs [|->] ls']

theorem eval1_eval:

  [| wwf_J_prog P; eval P e (h, l) e' (h', l'); fv e <= set Vs;
     lm [Vs [|->] ls]; length Vs + max_vars e <= length ls |]
  ==> EX ls'.
         eval1 (compP1 P) (compE1 Vs e) (h, ls) (fin1 e') (h', ls') &
         l'm [Vs [|->] ls']

and evals1_evals:

  [| wwf_J_prog P; evals P es (h, l) es' (h', l'); fvs es <= set Vs;
     lm [Vs [|->] ls]; length Vs + max_varss es <= length ls |]
  ==> EX ls'.
         evals1 (compP1 P) (compEs1 Vs es) (h, ls) (compEs1 Vs es') (h', ls') &
         l'm [Vs [|->] ls']

Preservation of well-formedness

lemma compE1_pres_wt:

  [| WT P [Vs [|->] Ts] e U; length Ts = length Vs |]
  ==> WT1 (compP f P) Ts (compE1 Vs e) U

and

  [| WTs P [Vs [|->] Ts] es Us; length Ts = length Vs |]
  ==> WTs1 (compP f P) Ts (compEs1 Vs es) Us

lemma \<B>:

  length Vs = n ==> \<B> (compE1 Vs e) n

and \<B>s:

  length Vs = n ==> \<B>s (compEs1 Vs es) n

lemma image_index:

  A <= set (xs @ [x])
  ==> index (xs @ [x]) ` A =
      (if x : A then insert (length xs) (index xs ` (A - {x})) else index xs ` A)

lemma A_compE1_None:

  \<A> e = None ==> \<A> (compE1 Vs e) = None

and

  \<A>s es = None ==> \<A>s (compEs1 Vs es) = None

lemma A_compE1:

  [| \<A> e = ⌊A⌋; fv e <= set Vs |] ==> \<A> (compE1 Vs e) = ⌊index Vs ` A

and

  [| \<A>s es = ⌊A⌋; fvs es <= set Vs |]
  ==> \<A>s (compEs1 Vs es) = ⌊index Vs ` A

lemma D_None:

  \<D> e None

and

  \<D>s es None

lemma D_index_compE1:

  [| A <= set Vs; fv e <= set Vs; \<D> eA⌋ |]
  ==> \<D> (compE1 Vs e) ⌊index Vs ` A

and

  [| A <= set Vs; fvs es <= set Vs; \<D>s esA⌋ |]
  ==> \<D>s (compEs1 Vs es) ⌊index Vs ` A

lemma index_image_set:

  distinct xs ==> index xs ` set xs = {..length xs(}

lemma D_compE1:

  [| \<D> e ⌊set Vs⌋; fv e <= set Vs; distinct Vs |]
  ==> \<D> (compE1 Vs e) ⌊{..length Vs(}⌋

lemma

  [| \<D> e ⌊set (V # Vs)⌋; fv e <= set (V # Vs); distinct (V # Vs) |]
  ==> \<D> (compE1 (V # Vs) e) ⌊{..length Vs}⌋

lemma compP1_pres_wf:

  wf_J_prog P ==> wf_J1_prog (compP1 P)