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