src/HOL/MicroJava/JVM/JVMExec.thy
author kleing
Fri, 22 Sep 2000 13:12:19 +0200
changeset 10057 8c8d2d0d3ef8
parent 10042 7164dc0d24d8
child 10060 4522e59b7d84
permissions -rw-r--r--
converted to Isar, tuned

(*  Title:      HOL/MicroJava/JVM/JVMExec.thy
    ID:         $Id$
    Author:     Cornelia Pusch
    Copyright   1999 Technische Universitaet Muenchen
*)

header {* Program Execution in the JVM *}

theory JVMExec = JVMExecInstr :


consts
  exec :: "jvm_prog \<times> jvm_state => jvm_state option"


(** exec is not recursive. recdef is just used for pattern matching **)
recdef exec "{}"
  "exec (G, xp, hp, []) = None"

  "exec (G, None, hp, (stk,loc,C,sig,pc)#frs) =
  (let 
     i = snd(snd(snd(the(method (G,C) sig)))) ! pc
   in
     Some (exec_instr i G hp stk loc C sig pc frs))"

  "exec (G, Some xp, hp, frs) = None" 


constdefs
  exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool"  
              ("_ \<turnstile> _ -jvm-> _" [61,61,61]60)
  "G \<turnstile> s -jvm-> t == (s,t) \<in> {(s,t). exec(G,s) = Some t}^*"


syntax (HTML output)
  exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool"  
              ("_ |- _ -jvm-> _" [61,61,61]60)

end