src/HOL/MicroJava/J/Value.thy
author ballarin
Mon, 02 Aug 2004 10:15:37 +0200
changeset 15097 b807858b97bd
parent 12911 704713ca07ea
child 16417 9bc16273c2d4
permissions -rw-r--r--
Modifications for trancl_tac (new solver in simplifier).
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
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     2
    ID:         $Id$
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     3
    Author:     David von Oheimb
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     4
    Copyright   1999 Technische Universitaet Muenchen
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
     5
*)
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     6
12911
704713ca07ea new document
kleing
parents: 12517
diff changeset
     7
header {* \isaheader{Java Values} *}
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
     8
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
     9
theory Value = Type:
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    10
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    11
typedecl loc_ -- "locations, i.e. abstract references on objects" 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    12
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    13
datatype loc 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    14
  = XcptRef xcpt -- "special locations for pre-allocated system exceptions"
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    15
  | Loc loc_     -- "usual locations (references on objects)"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    16
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    17
datatype val
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    18
  = Unit        -- "dummy result value of void methods"
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    19
  | Null        -- "null reference"
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    20
  | Bool bool   -- "Boolean value"
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    21
  | Intg int    -- "integer value, name Intg instead of Int because
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    22
                   of clash with HOL/Set.thy" 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    23
  | Addr loc    -- "addresses, i.e. locations of objects "
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    24
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    25
consts
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    26
  the_Bool :: "val => bool"
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    27
  the_Intg :: "val => int"
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    28
  the_Addr :: "val => loc"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    29
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    30
primrec
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    31
  "the_Bool (Bool b) = b"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    32
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    33
primrec
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    34
  "the_Intg (Intg i) = i"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    35
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    36
primrec
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    37
  "the_Addr (Addr a) = a"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    38
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    39
consts
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    40
  defpval :: "prim_ty => val"  -- "default value for primitive types"
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 12338
diff changeset
    41
  default_val :: "ty => val"   -- "default value for all types"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    42
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    43
primrec
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    44
  "defpval Void    = Unit"
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    45
  "defpval Boolean = Bool False"
11908
82f68fd05094 eliminated old numerals;
wenzelm
parents: 11701
diff changeset
    46
  "defpval Integer = Intg 0"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    47
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    48
primrec
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    49
  "default_val (PrimT pt) = defpval pt"
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    50
  "default_val (RefT  r ) = Null"
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    51
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff changeset
    52
end