Theory WellType
section ‹Well-typedness Constraints›
theory WellType imports Term WellForm begin
text ‹
the formulation of well-typedness of method calls given below (as well as
the Java Specification 1.0) is a little too restrictive: Is does not allow
methods of class Object to be called upon references of interface type.
\begin{description}
\item[simplifications:]\ \\
\begin{itemize}
\item the type rules include all static checks on expressions and statements, 
  e.g.\ definedness of names (of parameters, locals, fields, methods)
\end{itemize}
\end{description}
›
text "local variables, including method parameters and This:"
type_synonym lenv = "vname ⇀ ty"
type_synonym 'c env = "'c prog × lenv"
abbreviation (input)
  prg :: "'c env => 'c prog"
  where "prg == fst"
abbreviation (input)
  localT :: "'c env => (vname ⇀ ty)"
  where "localT == snd"
definition more_spec :: "'c prog ⇒ (ty × 'x) × ty list ⇒ (ty × 'x) × ty list ⇒ bool"
  where "more_spec G == λ((d,h),pTs). λ((d',h'),pTs'). G⊢d≼d' ∧
                                list_all2 (λT T'. G⊢T≼T') pTs pTs'"
definition appl_methds :: "'c prog ⇒ cname ⇒ sig ⇒ ((ty × ty) × ty list) set"
  
  where "appl_methds G C == λ(mn, pTs).
                     {((Class md,rT),pTs') |md rT mb pTs'.
                      method (G,C)  (mn, pTs') = Some (md,rT,mb) ∧
                      list_all2 (λT T'. G⊢T≼T') pTs pTs'}"
definition max_spec :: "'c prog ⇒ cname ⇒ sig ⇒ ((ty × ty) × ty list) set"
  
  where "max_spec G C sig == {m. m ∈appl_methds G C sig ∧ 
                                       (∀m'∈appl_methds G C sig.
                                         more_spec G m' m --> m' = m)}"
lemma max_spec2appl_meths: 
  "x ∈ max_spec G C sig ==> x ∈ appl_methds G C sig"
apply (unfold max_spec_def)
apply (fast)
done
lemma appl_methsD: 
"((md,rT),pTs')∈appl_methds G C (mn, pTs) ==>  
  ∃D b. md = Class D ∧ method (G,C) (mn, pTs') = Some (D,rT,b)  
  ∧ list_all2 (λT T'. G⊢T≼T') pTs pTs'"
apply (unfold appl_methds_def)
apply (fast)
done
lemmas max_spec2mheads = insertI1 [THEN [2] equalityD2 [THEN subsetD], 
                         THEN max_spec2appl_meths, THEN appl_methsD]
primrec typeof :: "(loc => ty option) => val => ty option"
where
  "typeof dt  Unit    = Some (PrimT Void)"
| "typeof dt  Null    = Some NT"
| "typeof dt (Bool b) = Some (PrimT Boolean)"
| "typeof dt (Intg i) = Some (PrimT Integer)"
| "typeof dt (Addr a) = dt a"
lemma is_type_typeof [rule_format (no_asm), simp]: 
  "(∀a. v ≠ Addr a) --> (∃T. typeof t v = Some T ∧ is_type G T)"
apply (rule val.induct)
apply     auto
done
lemma typeof_empty_is_type [rule_format (no_asm)]: 
  "typeof (λa. None) v = Some T ⟶ is_type G T"
apply (rule val.induct)
apply     auto
done
lemma typeof_default_val: "∃T. (typeof dt (default_val ty) = Some T) ∧ G⊢ T ≼ ty"
apply (case_tac ty)
apply (rename_tac prim_ty, case_tac prim_ty)
apply auto
done
type_synonym
  java_mb = "vname list × (vname × ty) list × stmt × expr"
  
inductive
  ty_expr :: "'c env => expr => ty => bool" (‹_ ⊢ _ :: _› [51, 51, 51] 50)
  and ty_exprs :: "'c env => expr list => ty list => bool" (‹_ ⊢ _ [::] _› [51, 51, 51] 50)
  and wt_stmt :: "'c env => stmt => bool" (‹_ ⊢ _ √› [51, 51] 50)
where
  
  NewC: "[| is_class (prg E) C |] ==>
         E⊢NewC C::Class C"  
  
| Cast: "[| E⊢e::C; is_class (prg E) D;
            prg E⊢C≼? Class D |] ==>
         E⊢Cast D e:: Class D"
  
| Lit:    "[| typeof (λv. None) x = Some T |] ==>
         E⊢Lit x::T"
  
  
| LAcc: "[| localT E v = Some T; is_type (prg E) T |] ==>
         E⊢LAcc v::T"
| BinOp:"[| E⊢e1::T;
            E⊢e2::T;
            if bop = Eq then T' = PrimT Boolean
                        else T' = T ∧ T = PrimT Integer|] ==>
            E⊢BinOp bop e1 e2::T'"
  
| LAss: "[| v ~= This;
            E⊢LAcc v::T;
            E⊢e::T';
            prg E⊢T'≼T |] ==>
         E⊢v::=e::T'"
  
| FAcc: "[| E⊢a::Class C; 
            field (prg E,C) fn = Some (fd,fT) |] ==>
            E⊢{fd}a..fn::fT"
  
| FAss: "[| E⊢{fd}a..fn::T;
            E⊢v        ::T';
            prg E⊢T'≼T |] ==>
         E⊢{fd}a..fn:=v::T'"
  
| Call: "[| E⊢a::Class C;
            E⊢ps[::]pTs;
            max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')} |] ==>
         E⊢{C}a..mn({pTs'}ps)::rT"
  
| Nil: "E⊢[][::][]"
  
| Cons:"[| E⊢e::T;
           E⊢es[::]Ts |] ==>
        E⊢e#es[::]T#Ts"
| Skip:"E⊢Skip√"
| Expr:"[| E⊢e::T |] ==>
        E⊢Expr e√"
| Comp:"[| E⊢s1√; 
           E⊢s2√ |] ==>
        E⊢s1;; s2√"
  
| Cond:"[| E⊢e::PrimT Boolean;
           E⊢s1√;
           E⊢s2√ |] ==>
         E⊢If(e) s1 Else s2√"
  
| Loop:"[| E⊢e::PrimT Boolean;
           E⊢s√ |] ==>
        E⊢While(e) s√"
definition wf_java_mdecl :: "'c prog => cname => java_mb mdecl => bool" where
"wf_java_mdecl G C == λ((mn,pTs),rT,(pns,lvars,blk,res)).
  length pTs = length pns ∧
  distinct pns ∧
  unique lvars ∧
        This ∉ set pns ∧ This ∉ set (map fst lvars) ∧ 
  (∀pn∈set pns. map_of lvars pn = None) ∧
  (∀(vn,T)∈set lvars. is_type G T) &
  (let E = (G,(map_of lvars)(pns[↦]pTs, This↦Class C)) in
   E⊢blk√ ∧ (∃T. E⊢res::T ∧ G⊢T≼rT))"
abbreviation "wf_java_prog == wf_prog wf_java_mdecl"
lemma wf_java_prog_wf_java_mdecl: "⟦ 
  wf_java_prog G; (C, D, fds, mths) ∈ set G; jmdcl ∈ set mths ⟧
  ⟹ wf_java_mdecl G C jmdcl"
apply (simp only: wf_prog_def) 
apply (erule conjE)+
apply (drule bspec, assumption)
apply (simp add: wf_cdecl_mdecl_def split_beta)
done
lemma wt_is_type: "(E⊢e::T ⟶ ws_prog (prg E) ⟶ is_type (prg E) T) ∧  
       (E⊢es[::]Ts ⟶ ws_prog (prg E) ⟶ Ball (set Ts) (is_type (prg E))) ∧ 
       (E⊢c √ ⟶ True)"
apply (rule ty_expr_ty_exprs_wt_stmt.induct)
apply auto
apply (   erule typeof_empty_is_type)
apply (  simp split: if_split_asm)
apply ( drule field_fields)
apply ( drule (1) fields_is_type)
apply (  simp (no_asm_simp))
apply  (assumption)
apply (auto dest!: max_spec2mheads method_wf_mhead is_type_rTI 
            simp add: wf_mdecl_def)
done
lemmas ty_expr_is_type = wt_is_type [THEN conjunct1,THEN mp, rule_format]
lemma expr_class_is_class: "
  ⟦ws_prog (prg E); E ⊢ e :: Class C⟧ ⟹ is_class (prg E) C"
  by (frule ty_expr_is_type, assumption, simp)
end