src/HOL/Bali/Name.thy
author wenzelm
Wed, 10 Feb 2010 00:45:16 +0100
changeset 35067 af4c18c30593
parent 32960 69916a850301
child 35069 09154b995ed8
permissions -rw-r--r--
modernized syntax translations, using mostly abbreviation/notation; minor tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12857
a4386cc9b1c3 tuned header;
wenzelm
parents: 12854
diff changeset
     1
(*  Title:      HOL/Bali/Name.thy
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     2
    ID:         $Id$
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     3
    Author:     David von Oheimb
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     4
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     5
header {* Java names *}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14981
diff changeset
     7
theory Name imports Basis begin
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     8
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     9
(* cf. 6.5 *) 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    10
typedecl tnam   --{* ordinary type name, i.e. class or interface name *}
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    11
typedecl pname  --{* package name *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    12
typedecl mname  --{* method name *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    13
typedecl vname  --{* variable or field name *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    14
typedecl label  --{* label as destination of break or continue *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    15
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    16
datatype ename        --{* expression name *} 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    17
        = VNam vname 
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    18
        | Res         --{* special name to model the return value of methods *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    19
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    20
datatype lname        --{* names for local variables and the This pointer *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    21
        = EName ename 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    22
        | This
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
    23
abbreviation VName   :: "vname \<Rightarrow> lname"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
    24
      where "VName n == EName (VNam n)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    25
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
    26
abbreviation Result :: lname
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
    27
      where "Result == EName Res"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    28
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    29
datatype xname          --{* names of standard exceptions *}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    30
        = Throwable
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    31
        | NullPointer | OutOfMemory | ClassCast   
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    32
        | NegArrSize  | IndOutBound | ArrStore
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    33
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    34
lemma xn_cases: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    35
  "xn = Throwable   \<or> xn = NullPointer \<or>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    36
         xn = OutOfMemory \<or> xn = ClassCast \<or> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    37
         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    38
apply (induct xn)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    39
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    40
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    41
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    42
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    43
datatype tname  --{* type names for standard classes and other type names *}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    44
        = Object'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    45
        | SXcpt'   xname
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24783
diff changeset
    46
        | TName   tnam
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    47
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    48
record   qtname = --{* qualified tname cf. 6.5.3, 6.5.4*}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 12859
diff changeset
    49
          pid :: pname  
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    50
          tid :: tname
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    51
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    52
axclass has_pname < "type"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    53
consts pname::"'a::has_pname \<Rightarrow> pname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    54
12859
wenzelm
parents: 12858
diff changeset
    55
instance pname::has_pname ..
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    56
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    57
defs (overloaded)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    58
pname_pname_def: "pname (p::pname) \<equiv> p"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    59
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    60
axclass has_tname < "type"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    61
consts tname::"'a::has_tname \<Rightarrow> tname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    62
12859
wenzelm
parents: 12858
diff changeset
    63
instance tname::has_tname ..
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    64
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    65
defs (overloaded)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    66
tname_tname_def: "tname (t::tname) \<equiv> t"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    67
17160
fb65eda72fc7 removed obsolete arities;
wenzelm
parents: 16417
diff changeset
    68
axclass has_qtname < type
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    69
consts qtname:: "'a::has_qtname \<Rightarrow> qtname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    70
17160
fb65eda72fc7 removed obsolete arities;
wenzelm
parents: 16417
diff changeset
    71
instance qtname_ext_type :: (type) has_qtname ..
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    72
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    73
defs (overloaded)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    74
qtname_qtname_def: "qtname (q::qtname) \<equiv> q"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    75
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    76
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    77
  "mname"  <= "Name.mname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    78
  "xname"  <= "Name.xname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    79
  "tname"  <= "Name.tname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    80
  "ename"  <= "Name.ename"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    81
  "qtname" <= (type) "\<lparr>pid::pname,tid::tname\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    82
  (type) "'a qtname_scheme" <= (type) "\<lparr>pid::pname,tid::tname,\<dots>::'a\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    83
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    84
19726
df95778b4c2f axiomatization java_lang;
wenzelm
parents: 17160
diff changeset
    85
axiomatization java_lang::pname --{* package java.lang *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    86
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    87
consts 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    88
  Object :: qtname
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    89
  SXcpt  :: "xname \<Rightarrow> qtname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    90
defs
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 19726
diff changeset
    91
  Object_def: "Object  \<equiv> \<lparr>pid = java_lang, tid = Object'\<rparr>"
5a3e336a2e37 avoid internal names;
wenzelm
parents: 19726
diff changeset
    92
  SXcpt_def:  "SXcpt   \<equiv> \<lambda>x.  \<lparr>pid = java_lang, tid = SXcpt' x\<rparr>"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    93
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    94
lemma Object_neq_SXcpt [simp]: "Object \<noteq> SXcpt xn"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    95
by (simp add: Object_def SXcpt_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    96
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    97
lemma SXcpt_inject [simp]: "(SXcpt xn = SXcpt xm) = (xn = xm)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    98
by (simp add: SXcpt_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    99
end