Theory Decl

(*  Title:      JinjaDCI/Common/Decl.thy

    Author:     David von Oheimb, Susannah Mansky
    Copyright   1999 Technische Universitaet Muenchen, 2019-20 UIUC

    Based on the Jinja theory Common/Decl.thy by David von Oheimb
*)

section ‹ Class Declarations and Programs ›

theory Decl imports Type begin

type_synonym 
  fdecl    = "vname × staticb × ty"        ― ‹field declaration›
type_synonym
  'm mdecl = "mname × staticb × ty list × ty × 'm"     ― ‹method = name, static flag, arg.\ types, return type, body›
type_synonym
  'm "class" = "cname × fdecl list × 'm mdecl list"       ― ‹class = superclass, fields, methods›
type_synonym
  'm cdecl = "cname × 'm class"  ― ‹class declaration›
type_synonym
  'm prog  = "'m cdecl list"     ― ‹program›

(* replaced all fname, mname, cname in below with `char list' so that
 pretty printing works   -SM *)
(*<*)
translations
  (type) "fdecl"   <= (type) "char list × staticb × ty"
  (type) "'c mdecl" <= (type) "char list × staticb × ty list × ty × 'c"
  (type) "'c class" <= (type) "char list × fdecl list × ('c mdecl) list"
  (type) "'c cdecl" <= (type) "char list × ('c class)"
  (type) "'c prog" <= (type) "('c cdecl) list"
(*>*)

definition "class" :: "'m prog  cname  'm class"
where
  "class    map_of"

(* Not difficult to prove, but useful for directing particular sequences of equality -SM *)
lemma class_cons: " C  fst x   class (x # P) C = class P C"
 by (simp add: class_def)

definition is_class :: "'m prog  cname  bool"
where
  "is_class P C    class P C  None"

lemma finite_is_class: "finite {C. is_class P C}"
(*<*)
proof -
  have "{C. is_class P C} = dom (map_of P)"
   by (simp add: is_class_def class_def dom_def)
  thus ?thesis by (simp add: finite_dom_map_of)
qed
(*>*)

definition is_type :: "'m prog  ty  bool"
where
  "is_type P T  
  (case T of Void  True | Boolean  True | Integer  True | NT  True
   | Class C  is_class P C)"

lemma is_type_simps [simp]:
  "is_type P Void  is_type P Boolean  is_type P Integer 
  is_type P NT  is_type P (Class C) = is_class P C"
(*<*)by(simp add:is_type_def)(*>*)


abbreviation
  "types P == Collect (is_type P)"

lemma class_exists_equiv:
 "(x. fst x = cn  x  set P) = (class P cn  None)"
proof(rule iffI)
 assume "x. fst x = cn  x  set P" then show "class P cn  None"
   by (metis class_def image_eqI map_of_eq_None_iff)
next
 assume "class P cn  None" then show "x. fst x = cn  x  set P"
   by (metis class_def fst_conv map_of_SomeD option.exhaust)
qed

lemma class_exists_equiv2:
 "(x. fst x = cn  x  set (P1 @ P2)) = (class P1 cn  None  class P2 cn  None)"
by (simp only: class_exists_equiv [where P = "P1@P2"], simp add: class_def)

end