Theory WWellForm

(*  Title:       CoreC++

    Author:      Tobias Nipkow
    Maintainer:  Daniel Wasserrab <wasserra at fmi.uni-passau.de>
*)

section ‹Weak well-formedness of CoreC++ programs›

theory WWellForm imports WellForm Expr begin


definition wwf_mdecl :: "prog  cname  mdecl  bool" where
  "wwf_mdecl P C    λ(M,Ts,T,(pns,body)).
  length Ts = length pns  distinct pns  this  set pns  fv body  {this}  set pns"

lemma wwf_mdecl[simp]:
  "wwf_mdecl P C (M,Ts,T,pns,body) =
  (length Ts = length pns  distinct pns  this  set pns  fv body  {this}  set pns)"
by(simp add:wwf_mdecl_def)


abbreviation
  wwf_prog :: "prog  bool" where
  "wwf_prog == wf_prog wwf_mdecl"

end