src/HOL/MicroJava/BV/BVSpec.ML
changeset 8011 d14c4e9e9c8e
child 8032 1eaae1a2f8ff
equal deleted inserted replaced
8010:69032a618aa9 8011:d14c4e9e9c8e
       
     1 (*  Title:      HOL/MicroJava/BV/BVSpec.ML
       
     2     ID:         $Id$
       
     3     Author:     Cornelia Pusch
       
     4     Copyright   1999 Technische Universitaet Muenchen
       
     5 *)
       
     6 
       
     7 Addsimps [wt_MR_def];
       
     8 
       
     9 Goalw [wt_jvm_prog_def] "wt_jvm_prog G phi \\<Longrightarrow> (\\<exists>wt. wf_prog wt G)";
       
    10 by(Blast_tac 1);
       
    11 qed "wt_jvm_progD";
       
    12 
       
    13 
       
    14 Goalw [wt_jvm_prog_def]
       
    15 "\\<lbrakk> wt_jvm_prog G phi; \
       
    16 \   cmethd (G,C) sig = Some (C,rT,maxl,ins); \
       
    17 \   pc < length ins \\<rbrakk> \
       
    18 \\\<Longrightarrow> wt_instr (ins!pc) G rT (phi C sig) (length ins) pc";
       
    19 by(forward_tac [rotate_prems 2 cmethd_wf_mdecl] 1);
       
    20  by (Asm_full_simp_tac 1);
       
    21 by(asm_full_simp_tac (simpset() addsimps [wf_mdecl_def,wt_method_def]) 1);
       
    22 qed "wt_jvm_prog_impl_wt_instr";
       
    23 
       
    24 Goalw [wt_jvm_prog_def]
       
    25 "\\<lbrakk> wt_jvm_prog G phi; \
       
    26 \   cmethd (G,C) sig = Some (C,rT,maxl,ins) \\<rbrakk> \\<Longrightarrow> \
       
    27 \ 0 < (length ins) \\<and> wt_start G C (snd sig) maxl (phi C sig)";
       
    28 by(forward_tac [rotate_prems 2 cmethd_wf_mdecl] 1);
       
    29  by (Asm_full_simp_tac 1);
       
    30 by(asm_full_simp_tac (simpset() addsimps [wf_mdecl_def,wt_method_def]) 1);
       
    31 qed "wt_jvm_prog_impl_wt_start";