src/HOL/MicroJava/J/Value.thy
author Manuel Eberl <manuel@pruvisto.org>
Tue, 15 Apr 2025 17:38:20 +0200
changeset 82518 da14e77a48b2
parent 67443 3abf6a722518
permissions -rw-r--r--
lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
     1
(*  Title:      HOL/MicroJava/J/Value.thy
41589
bbd861837ebc tuned headers;
wenzelm
parents: 39758
diff changeset
     2
    Author:     David von Oheimb, Technische Universitaet Muenchen
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
     3
*)
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     4
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
     5
section \<open>Java Values\<close>
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12911
diff changeset
     7
theory Value imports Type begin
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     8
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
     9
typedecl loc' \<comment> \<open>locations, i.e. abstract references on objects\<close> 
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    10
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    11
datatype loc 
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    12
  = XcptRef xcpt \<comment> \<open>special locations for pre-allocated system exceptions\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    13
  | Loc loc'     \<comment> \<open>usual locations (references on objects)\<close>
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    14
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    15
datatype val
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    16
  = Unit        \<comment> \<open>dummy result value of void methods\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    17
  | Null        \<comment> \<open>null reference\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    18
  | Bool bool   \<comment> \<open>Boolean value\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    19
  | Intg int    \<comment> \<open>integer value, name Intg instead of Int because
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    20
                   of clash with HOL/Set.thy\<close> 
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    21
  | Addr loc    \<comment> \<open>addresses, i.e. locations of objects\<close>
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    22
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    23
primrec the_Bool :: "val => bool" where
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    24
  "the_Bool (Bool b) = b"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    25
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    26
primrec the_Intg :: "val => int" where
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    27
  "the_Intg (Intg i) = i"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    28
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    29
primrec the_Addr :: "val => loc" where
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    30
  "the_Addr (Addr a) = a"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    31
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    32
primrec defpval :: "prim_ty => val"  \<comment> \<open>default value for primitive types\<close> where
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    33
  "defpval Void    = Unit"
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    34
| "defpval Boolean = Bool False"
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    35
| "defpval Integer = Intg 0"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    36
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62042
diff changeset
    37
primrec default_val :: "ty => val"   \<comment> \<open>default value for all types\<close> where
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    38
  "default_val (PrimT pt) = defpval pt"
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 24783
diff changeset
    39
| "default_val (RefT  r ) = Null"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    40
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    41
end