Theory J1WellForm

(*  Title:      JinjaDCI/Compiler/J1WellForm.thy

    Author:     Tobias Nipkow, Susannah Mansky
    Copyright   2003 Technische Universitaet Muenchen, 2019-20 UIUC

    Based on the Jinja theory Compiler/J1WellForm.thy by Tobias Nipkow
*)

section ‹ Well-Formedness of Intermediate Language ›

theory J1WellForm
imports "../J/JWellForm" J1
begin

subsection "Well-Typedness"

type_synonym 
  env1  = "ty list"   ― ‹type environment indexed by variable number›

inductive
  WT1 :: "[J1_prog,env1, expr1     , ty     ]  bool"
         ("(_,_ 1/ _ :: _)"   [51,51,51]50)
  and WTs1 :: "[J1_prog,env1, expr1 list, ty list]  bool"
         ("(_,_ 1/ _ [::] _)" [51,51,51]50)
  for P :: J1_prog
where
  
  WTNew1:
  "is_class P C  
  P,E 1 new C :: Class C"

| WTCast1:
  " P,E 1 e :: Class D;  is_class P C;  P  C * D  P  D * C 
   P,E 1 Cast C e :: Class C"

| WTVal1:
  "typeof v = Some T 
  P,E 1 Val v :: T"

| WTVar1:
  " E!i = T; i < size E 
   P,E 1 Var i :: T"

| WTBinOp1:
  " P,E 1 e1 :: T1;  P,E 1 e2 :: T2;
     case bop of Eq  (P  T1  T2  P  T2  T1)  T = Boolean
               | Add  T1 = Integer  T2 = Integer  T = Integer 
   P,E 1 e1 «bop» e2 :: T"

| WTLAss1:
  " E!i = T;  i < size E; P,E 1 e :: T';  P  T'  T 
   P,E 1 i:=e :: Void"

| WTFAcc1:
  " P,E 1 e :: Class C;  P  C sees F,NonStatic:T in D 
   P,E 1 eF{D} :: T"

| WTSFAcc1:
  " P  C sees F,Static:T in D 
   P,E 1 CsF{D} :: T"

| WTFAss1:
  " P,E 1 e1 :: Class C;  P  C sees F,NonStatic:T in D;  P,E 1 e2 :: T';  P  T'  T 
   P,E 1 e1F{D} := e2 :: Void"

| WTSFAss1:
  "  P  C sees F,Static:T in D;  P,E 1 e2 :: T';  P  T'  T 
   P,E 1 CsF{D}:=e2 :: Void"

| WTCall1:
  " P,E 1 e :: Class C; P  C sees M,NonStatic:Ts'  T = m in D;
    P,E 1 es [::] Ts;  P  Ts [≤] Ts' 
   P,E 1 eM(es) :: T"

| WTSCall1:
  " P  C sees M,Static:Ts  T = m in D;
     P,E 1 es [::] Ts';  P  Ts' [≤] Ts; M  clinit 
   P,E 1 CsM(es) :: T"

| WTBlock1:
  " is_type P T; P,E@[T] 1 e::T' 
    P,E 1 {i:T; e} :: T'"

| WTSeq1:
  " P,E 1 e1::T1;  P,E 1 e2::T2 
    P,E 1 e1;;e2 :: T2"

| WTCond1:
  " P,E 1 e :: Boolean;  P,E 1 e1::T1;  P,E 1 e2::T2;
    P  T1  T2  P  T2  T1;  P  T1  T2  T = T2; P  T2  T1  T = T1 
   P,E 1 if (e) e1 else e2 :: T"

| WTWhile1:
  " P,E 1 e :: Boolean;  P,E 1 c::T 
   P,E 1 while (e) c :: Void"

| WTThrow1:
  "P,E 1 e :: Class C  
  P,E 1 throw e :: Void"

| WTTry1:
  " P,E 1 e1 :: T;  P,E@[Class C] 1 e2 :: T; is_class P C 
   P,E 1 try e1 catch(C i) e2 :: T"

| WTNil1:
  "P,E 1 [] [::] []"

| WTCons1:
  " P,E 1 e :: T; P,E 1 es [::] Ts 
    P,E 1 e#es [::] T#Ts"

(*<*)
declare  WT1_WTs1.intros[intro!]
declare WTNil1[iff]

lemmas WT1_WTs1_induct = WT1_WTs1.induct [split_format (complete)]
  and WT1_WTs1_inducts = WT1_WTs1.inducts [split_format (complete)]

inductive_cases eee[elim!]:
  "P,E 1 Val v :: T"
  "P,E 1 Var i :: T"
  "P,E 1 Cast D e :: T"
  "P,E 1 i:=e :: T"
  "P,E 1 {i:U; e} :: T"
  "P,E 1 e1;;e2 :: T"
  "P,E 1 if (e) e1 else e2 :: T"
  "P,E 1 while (e) c :: T"
  "P,E 1 throw e :: T"
  "P,E 1 try e1 catch(C i) e2 :: T"
  "P,E 1 eF{D} :: T"
  "P,E 1 CsF{D} :: T"
  "P,E 1 e1F{D}:=e2 :: T"
  "P,E 1 CsF{D}:=e2 :: T"
  "P,E 1 e1 «bop» e2 :: T"
  "P,E 1 new C :: T"
  "P,E 1 eM(es) :: T"
  "P,E 1 CsM(es) :: T"
  "P,E 1 [] [::] Ts"
  "P,E 1 e#es [::] Ts"
(*>*)

lemma init_nWT1 [simp]:"¬P,E 1 INIT C (Cs,b)  e :: T"
 by(auto elim:WT1.cases)
lemma rinit_nWT1 [simp]:"¬P,E 1 RI(C,e);Cs  e' :: T"
 by(auto elim:WT1.cases)

lemma WTs1_same_size: "Ts. P,E 1 es [::] Ts  size es = size Ts"
(*<*)by (induct es type:list) auto(*>*)


lemma WT1_unique:
  "P,E 1 e :: T1  (T2. P,E 1 e :: T2  T1 = T2)" and
  WTs1_unique: "P,E 1 es [::] Ts1  (Ts2. P,E 1 es [::] Ts2  Ts1 = Ts2)"
(*<*)
proof(induct rule:WT1_WTs1.inducts)
  case WTVal1 then show ?case by clarsimp
next
  case (WTBinOp1 E e1 T1 e2 T2 bop T)
  then show ?case by(case_tac bop) force+
next
  case WTFAcc1 then show ?case
    by (blast dest:sees_field_idemp sees_field_fun)
next
  case WTSFAcc1 then show ?case by (blast dest:sees_field_fun)
next
  case WTSFAss1 then show ?case by (blast dest:sees_field_fun)
next
  case WTCall1 then show ?case
    by (blast dest:sees_method_idemp sees_method_fun)
next
  case WTSCall1 then show ?case by (blast dest:sees_method_fun)
qed blast+
(*>*)


lemma assumes wf: "wf_prog p P"
shows WT1_is_type: "P,E 1 e :: T  set E  types P  is_type P T"
and "P,E 1 es [::] Ts  True"
(*<*)
proof(induct rule:WT1_WTs1.inducts)
  case WTVal1 then show ?case by (simp add:typeof_lit_is_type)
next
  case WTVar1 then show ?case by (blast intro:nth_mem)
next
  case WTBinOp1 then show ?case by (simp split:bop.splits)
next
  case WTFAcc1 then show ?case
    by (simp add:sees_field_is_type[OF _ wf])
next
  case WTSFAcc1 then show ?case
    by (simp add:sees_field_is_type[OF _ wf])
next
  case WTCall1 then show ?case
    by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
next
  case WTSCall1 then show ?case
    by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
next
  case WTCond1 then show ?case by blast
qed simp+
(*>*)

lemma WT1_nsub_RI: "P,E 1 e :: T  ¬sub_RI e"
 and WTs1_nsub_RIs: "P,E 1 es [::] Ts  ¬sub_RIs es"
proof(induct rule: WT1_WTs1.inducts) qed(simp_all)

subsection‹ Runtime Well-Typedness ›

inductive
  WTrt1 :: "J1_prog  heap  sheap  env1  expr1  ty  bool"
  and WTrts1 :: "J1_prog  heap  sheap  env1  expr1 list  ty list  bool"
  and WTrt21 :: "[J1_prog,env1,heap,sheap,expr1,ty]  bool"
        ("_,_,_,_ 1 _ : _"   [51,51,51,51]50)
  and WTrts21 :: "[J1_prog,env1,heap,sheap,expr1 list, ty list]  bool"
        ("_,_,_,_ 1 _ [:] _" [51,51,51,51]50)
  for P :: J1_prog and h :: heap and sh :: sheap
where
  
  "P,E,h,sh 1 e : T  WTrt1 P h sh E e T"
| "P,E,h,sh 1 es[:]Ts  WTrts1 P h sh E es Ts"

| WTrtNew1:
  "is_class P C  
  P,E,h,sh 1 new C : Class C"

| WTrtCast1:
  " P,E,h,sh 1 e : T; is_refT T; is_class P C 
   P,E,h,sh 1 Cast C e : Class C"

| WTrtVal1:
  "typeof⇘hv = Some T 
  P,E,h,sh 1 Val v : T"

| WTrtVar1:
  " E!i = T; i < size E   
  P,E,h,sh 1 Var i : T"

| WTrtBinOpEq1:
  " P,E,h,sh 1 e1 : T1;  P,E,h,sh 1 e2 : T2 
   P,E,h,sh 1 e1 «Eq» e2 : Boolean"

| WTrtBinOpAdd1:
  " P,E,h,sh 1 e1 : Integer;  P,E,h,sh 1 e2 : Integer 
   P,E,h,sh 1 e1 «Add» e2 : Integer"

| WTrtLAss1:
  " E!i = T; i < size E; P,E,h,sh 1 e : T';  P  T'  T 
    P,E,h,sh 1 i:=e : Void"

| WTrtFAcc1:
  " P,E,h,sh 1 e : Class C; P  C has F,NonStatic:T in D  
  P,E,h,sh 1 eF{D} : T"

| WTrtFAccNT1:
  "P,E,h,sh 1 e : NT 
  P,E,h,sh 1 eF{D} : T"

| WTrtSFAcc1:
  " P  C has F,Static:T in D  
  P,E,h,sh 1 CsF{D} : T"

| WTrtFAss1:
  " P,E,h,sh 1 e1 : Class C;  P  C has F,NonStatic:T in D; P,E,h,sh 1 e2 : T2;  P  T2  T 
   P,E,h,sh 1 e1F{D}:=e2 : Void"

| WTrtFAssNT1:
  " P,E,h,sh 1 e1:NT; P,E,h,sh 1 e2 : T2 
   P,E,h,sh 1 e1F{D}:=e2 : Void"

| WTrtSFAss1:
  " P,E,h,sh 1 e2 : T2; P  C has F,Static:T in D; P  T2  T 
   P,E,h,sh 1 CsF{D}:=e2 : Void"

| WTrtCall1:
  " P,E,h,sh 1 e : Class C; P  C sees M,NonStatic:Ts  T = m in D;
     P,E,h,sh 1 es [:] Ts'; P  Ts' [≤] Ts 
   P,E,h,sh 1 eM(es) : T"

| WTrtCallNT1:
  " P,E,h,sh 1 e : NT; P,E,h,sh 1 es [:] Ts 
   P,E,h,sh 1 eM(es) : T"

| WTrtSCall1:
  " P  C sees M,Static:Ts  T = m in D;
     P,E,h,sh 1 es [:] Ts'; P  Ts' [≤] Ts;
     M = clinit  sh D = (sfs,Processing)  es = map Val vs 
   P,E,h,sh 1 CsM(es) : T"

| WTrtBlock1:
  "P,E@[T],h,sh 1 e : T'  
  P,E,h,sh 1 {i:T; e} : T'"

| WTrtSeq1:
  " P,E,h,sh 1 e1:T1;  P,E,h,sh 1 e2:T2 
   P,E,h,sh 1 e1;;e2 : T2"

| WTrtCond1:
  " P,E,h,sh 1 e : Boolean;  P,E,h,sh 1 e1:T1;  P,E,h,sh 1 e2:T2;
     P  T1  T2  P  T2  T1; P  T1  T2  T = T2; P  T2  T1  T = T1 
   P,E,h,sh 1 if (e) e1 else e2 : T"

| WTrtWhile1:
  " P,E,h,sh 1 e : Boolean;  P,E,h,sh 1 c:T 
    P,E,h,sh 1 while(e) c : Void"

| WTrtThrow1:
  " P,E,h,sh 1 e : Tr; is_refT Tr  
  P,E,h,sh 1 throw e : T"

| WTrtTry1:
  " P,E,h,sh 1 e1 : T1;  P,E@[Class C],h,sh 1 e2 : T2; P  T1  T2 
   P,E,h,sh 1 try e1 catch(C i) e2 : T2"

| WTrtInit1:
  " P,E,h,sh 1 e : T; C'  set (C#Cs). is_class P C'; ¬sub_RI e;
     C'  set (tl Cs). sfs. sh C' = (sfs,Processing);
     b  (C'  set Cs. sfs. sh C' = (sfs,Processing));
     distinct Cs; supercls_lst P Cs 
   P,E,h,sh 1 INIT C (Cs, b)  e : T"

| WTrtRI1:
  " P,E,h,sh 1 e : T; P,E,h,sh 1 e' : T'; C'  set (C#Cs). is_class P C'; ¬sub_RI e';
     C'  set (C#Cs). not_init C' e;
     C'  set Cs. sfs. sh C' = (sfs,Processing);
     sfs. sh C = (sfs, Processing)  (sh C = (sfs, Error)  e = THROW NoClassDefFoundError);
     distinct (C#Cs); supercls_lst P (C#Cs) 
   P,E,h,sh 1 RI(C, e);Cs  e' : T'"

― ‹well-typed expression lists›

| WTrtNil1:
  "P,E,h,sh 1 [] [:] []"

| WTrtCons1:
  " P,E,h,sh 1 e : T;  P,E,h,sh 1 es [:] Ts 
    P,E,h,sh 1 e#es [:] T#Ts"

(*<*)
declare WTrt1_WTrts1.intros[intro!] WTrtNil1[iff]
declare
  WTrtFAcc1[rule del] WTrtFAccNT1[rule del] WTrtSFAcc1[rule del]
  WTrtFAss1[rule del] WTrtFAssNT1[rule del] WTrtSFAss1[rule del]
  WTrtCall1[rule del] WTrtCallNT1[rule del] WTrtSCall1[rule del]

lemmas WTrt1_induct = WTrt1_WTrts1.induct [split_format (complete)]
  and WTrt1_inducts = WTrt1_WTrts1.inducts [split_format (complete)]
(*>*)

(*<*)
inductive_cases WTrt1_elim_cases[elim!]:
  "P,E,h,sh 1 Val v : T"
  "P,E,h,sh 1 Var i : T"
  "P,E,h,sh 1 v :=e : T"
  "P,E,h,sh 1 {i:U; e} : T"
  "P,E,h,sh 1 e1;;e2 : T2"
  "P,E,h,sh 1 if (e) e1 else e2 : T"
  "P,E,h,sh 1 while(e) c : T"
  "P,E,h,sh 1 throw e : T"
  "P,E,h,sh 1 try e1 catch(C V) e2 : T"
  "P,E,h,sh 1 Cast D e : T"
  "P,E,h,sh 1 eF{D} : T"
  "P,E,h,sh 1 CsF{D} : T"
  "P,E,h,sh 1 eF{D} := v : T"
  "P,E,h,sh 1 CsF{D} := v : T"
  "P,E,h,sh 1 e1 «bop» e2 : T"
  "P,E,h,sh 1 new C : T"
  "P,E,h,sh 1 eM{D}(es) : T"
  "P,E,h,sh 1 CsM{D}(es) : T"
  "P,E,h,sh 1 INIT C (Cs,b)  e : T"
  "P,E,h,sh 1 RI(C,e);Cs  e' : T"
  "P,E,h,sh 1 [] [:] Ts"
  "P,E,h,sh 1 e#es [:] Ts"
(*>*)

lemma WT1_implies_WTrt1: "P,E 1 e :: T  P,E,h,sh 1 e : T"
and WTs1_implies_WTrts1: "P,E 1 es [::] Ts  P,E,h,sh 1 es [:] Ts"
(*<*)
proof(induct rule: WT1_WTs1_inducts)
  case WTVal1 then show ?case by (fastforce dest:typeof_lit_typeof)
next
  case (WTBinOp1 E e1 T1 e2 T2 bop T)
  then show ?case by (case_tac bop) fastforce+
next
  case WTFAcc1 then show ?case
    by (fastforce simp: WTrtFAcc1 has_visible_field)
next
  case WTSFAcc1 then show ?case
    by (fastforce simp: WTrtSFAcc1 has_visible_field)
next
  case WTFAss1 then show ?case by (meson WTrtFAss1 has_visible_field)
next
  case WTSFAss1 then show ?case by (meson WTrtSFAss1 has_visible_field)
next
  case WTCall1 then show ?case by (fastforce simp: WTrtCall1)
next
  case WTSCall1 then show ?case by (fastforce simp: WTrtSCall1)
qed fastforce+
(*>*)

subsection‹ Well-formedness›

― ‹Indices in blocks increase by 1›

primrec  :: "expr1  nat  bool"
  and ℬs :: "expr1 list  nat  bool" where
" (new C) i = True" |
" (Cast C e) i =  e i" |
" (Val v) i = True" |
" (e1 «bop» e2) i = ( e1 i   e2 i)" |
" (Var j) i = True" |
" (eF{D}) i =  e i" |
" (CsF{D}) i = True" |
" (j:=e) i =  e i" |
" (e1F{D} := e2) i = ( e1 i   e2 i)" |
" (CsF{D} := e2) i =  e2 i" |
" (eM(es)) i = ( e i  ℬs es i)" |
" (CsM(es)) i = ℬs es i" |
" ({j:T ; e}) i = (i = j   e (i+1))" |
" (e1;;e2) i = ( e1 i   e2 i)" |
" (if (e) e1 else e2) i = ( e i   e1 i   e2 i)" |
" (throw e) i =  e i" |
" (while (e) c) i = ( e i   c i)" |
" (try e1 catch(C j) e2) i = ( e1 i  i=j   e2 (i+1))" |
" (INIT C (Cs,b)  e) i =  e i" |
" (RI(C,e);Cs  e') i = ( e i   e' i)" |

"ℬs [] i = True" |
"ℬs (e#es) i = ( e i  ℬs es i)"


definition wf_J1_mdecl :: "J1_prog  cname  expr1 mdecl  bool"
where
  "wf_J1_mdecl P C    λ(M,b,Ts,T,body).
    ¬sub_RI body 
 (case b of
    NonStatic 
        (T'. P,Class C#Ts 1 body :: T'  P  T'  T) 
        𝒟 body {..size Ts}   body (size Ts + 1)
  | Static  (T'. P,Ts 1 body :: T'  P  T'  T) 
        𝒟 body {..<size Ts}   body (size Ts))"

lemma wf_J1_mdecl_NonStatic[simp]:
  "wf_J1_mdecl P C (M,NonStatic,Ts,T,body) 
    (¬sub_RI body 
    (T'. P,Class C#Ts 1 body :: T'  P  T'  T) 
     𝒟 body {..size Ts}   body (size Ts + 1))"
(*<*)by (simp add:wf_J1_mdecl_def)(*>*)

lemma wf_J1_mdecl_Static[simp]:
  "wf_J1_mdecl P C (M,Static,Ts,T,body) 
    (¬sub_RI body 
    (T'. P,Ts 1 body :: T'  P  T'  T) 
     𝒟 body {..<size Ts}   body (size Ts))"
(*<*)by (simp add:wf_J1_mdecl_def)(*>*)

abbreviation "wf_J1_prog == wf_prog wf_J1_mdecl"

lemma sees_wf1_nsub_RI:
assumes wf: "wf_J1_prog P" and cM: "P  C sees M,b : TsT = body in D"
shows "¬sub_RI body"
using sees_wf_mdecl[OF wf cM] by(simp add: wf_J1_mdecl_def wf_mdecl_def)

lemma wf1_types_clinit:
assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = (sfs, Processing)"
shows "P,E,h,sh 1 Csclinit([]) : Void"
proof -
  from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
  then have sP: "(C, D, fs, ms)  set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
  then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
  then obtain m where sm: "(clinit, Static, [], Void, m)  set ms"
    by(unfold wf_clinit_def) auto
  then have "P  C sees clinit,Static:[]  Void = m in C"
    using mdecl_visible[OF wf sP sm] by simp
  then show ?thesis using WTrtSCall1 proc by blast
qed


lemma assumes wf: "wf_J1_prog P"
shows eval1_proc_pres: "P 1 e,(h,l,sh)  e',(h',l',sh')
   not_init C e  sfs. sh C = (sfs, Processing)  sfs'. sh' C = (sfs', Processing)"
  and evals1_proc_pres: "P 1 es,(h,l,sh) [⇒] es',(h',l',sh')
   not_inits C es  sfs. sh C = (sfs, Processing)  sfs'. sh' C = (sfs', Processing)"
(*<*)
proof(induct rule:eval1_evals1_inducts)
  case Call1 then show ?case using sees_wf1_nsub_RI[OF wf Call1.hyps(6)] nsub_RI_not_init by auto
next
  case (SCallInit1 ps h l sh vs h1 l1 sh1 C' M Ts T body D v' h2 l2 sh2 l2' e' h3 l3 sh3)
  then show ?case
    using SCallInit1 sees_wf1_nsub_RI[OF wf SCallInit1.hyps(3)] nsub_RI_not_init[of body] by auto
next
  case SCall1 then show ?case using sees_wf1_nsub_RI[OF wf SCall1.hyps(3)] nsub_RI_not_init by auto
next
  case (InitNone1 sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
next
  case (InitDone1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
next
  case (InitProcessing1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
next
  case (InitError1 sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto
next
  case (InitObject1 sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
next
  case (InitNonObject1 sh C1 sfs D a b sh' C' Cs h l e' a a b)
  then show ?case by(cases "C = C1") auto
next
  case (RInit1 e a a b v h' l' sh' C sfs i sh'' C' Cs e1 a a b) then show ?case by(cases Cs, auto)
next
  case (RInitInitFail1 e h l sh a h' l' sh' C1 sfs i sh'' D Cs e1 h1 l1 sh1)
  then show ?case using eval1_final by fastforce
qed(auto)
(*>*)

lemma clinit1_proc_pres:
  " wf_J1_prog P; P 1 C0sclinit([]),(h,l,sh)  e',(h',l',sh');
     sh C' = (sfs,Processing) 
   sfs. sh' C' = (sfs,Processing)"
 by(auto dest: eval1_proc_pres)

end