Theory JWellForm

Up to index of Isabelle/HOL/Jinja

theory JWellForm
imports WWellForm WellType DefAss
(*  Title:      Jinja/J/JWellForm.thy

Author: Tobias Nipkow
Copyright 2003 Technische Universitaet Muenchen
*)


header {* \isaheader{Well-formedness Constraints} *}

theory JWellForm
imports "../Common/WellForm" WWellForm WellType DefAss
begin

definition wf_J_mdecl :: "J_prog => cname => J_mb mdecl => bool"
where
"wf_J_mdecl P C ≡ λ(M,Ts,T,(pns,body)).
length Ts = length pns ∧
distinct pns ∧
this ∉ set pns ∧
(∃T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' ∧ P \<turnstile> T' ≤ T) ∧
\<D> body ⌊{this} ∪ set pns⌋"


lemma wf_J_mdecl[simp]:
"wf_J_mdecl P C (M,Ts,T,pns,body) ≡
(length Ts = length pns ∧
distinct pns ∧
this ∉ set pns ∧
(∃T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' ∧ P \<turnstile> T' ≤ T) ∧
\<D> body ⌊{this} ∪ set pns⌋)"

(*<*)by(simp add:wf_J_mdecl_def)(*>*)


abbreviation
wf_J_prog :: "J_prog => bool" where
"wf_J_prog == wf_prog wf_J_mdecl"

lemma wf_J_prog_wf_J_mdecl:
"[| wf_J_prog P; (C, D, fds, mths) ∈ set P; jmdcl ∈ set mths |]
==> wf_J_mdecl P C jmdcl"

(*<*)
apply (simp add: wf_prog_def)
apply (simp add: wf_cdecl_def)
apply (erule conjE)+
apply (drule bspec, assumption)
apply simp
apply (erule conjE)+
apply (drule bspec, assumption)
apply (simp add: wf_mdecl_def split_beta)
done
(*>*)


lemma wf_mdecl_wwf_mdecl: "wf_J_mdecl P C Md ==> wwf_J_mdecl P C Md"
(*<*)by(fastforce simp:wwf_J_mdecl_def dest!:WT_fv)(*>*)


lemma wf_prog_wwf_prog: "wf_J_prog P ==> wwf_J_prog P"
(*<*)
apply(simp add:wf_prog_def wf_cdecl_def wf_mdecl_def)
apply(fast intro:wf_mdecl_wwf_mdecl)
done
(*>*)


end