src/HOL/MicroJava/J/Term.thy
changeset 61361 8b5f00202e1a
parent 58886 8a6cac7c7247
child 62042 6c6ccf573479
equal deleted inserted replaced
61360:a273bdac0934 61361:8b5f00202e1a
     1 (*  Title:      HOL/MicroJava/J/Term.thy
     1 (*  Title:      HOL/MicroJava/J/Term.thy
     2     Author:     David von Oheimb, Technische Universitaet Muenchen
     2     Author:     David von Oheimb, Technische Universitaet Muenchen
     3 *)
     3 *)
     4 
     4 
     5 section {* Expressions and Statements *}
     5 section \<open>Expressions and Statements\<close>
     6 
     6 
     7 theory Term imports Value begin
     7 theory Term imports Value begin
     8 
     8 
     9 datatype binop = Eq | Add    -- "function codes for binary operation"
     9 datatype binop = Eq | Add    -- "function codes for binary operation"
    10 
    10