src/HOL/MicroJava/J/TypeRel.ML
changeset 8011 d14c4e9e9c8e
child 8034 6fc37b5c5e98
equal deleted inserted replaced
8010:69032a618aa9 8011:d14c4e9e9c8e
       
     1 (*  Title:      HOL/MicroJava/J/TypeRel.ML
       
     2     ID:         $Id$
       
     3     Author:     David von Oheimb
       
     4     Copyright   1999 Technische Universitaet Muenchen
       
     5 *)
       
     6 
       
     7 val subcls1D = prove_goalw thy [subcls1_def] "\\<And>G. G\\<turnstile>C\\<prec>C1D \\<Longrightarrow> \
       
     8 \ \\<exists>fs ms. class G C = Some (Some D,fs,ms)" (K [Auto_tac]);
       
     9 
       
    10 val subcls1I = prove_goalw  thy [subcls1_def] 
       
    11 "\\<And>G. \\<lbrakk> class G C = Some (Some D,rest) \\<rbrakk> \\<Longrightarrow> G\\<turnstile>C\\<prec>C1D" (K [Auto_tac]);
       
    12 
       
    13 val subcls1_def2 = prove_goalw thy [subcls1_def,is_class_def]  "subcls1 G = \
       
    14 \ (SIGMA C:{C. is_class G C} . {D. fst (the (class G C)) = Some D})"
       
    15  (K [Auto_tac]);
       
    16 
       
    17 context Option.thy;
       
    18 Goal "{y. x = Some y} \\<subseteq> {the x}";
       
    19 by Auto_tac;
       
    20 val some_subset_the = result();
       
    21 context thy;
       
    22 
       
    23 Goal "finite (subcls1 G)";
       
    24 by(stac subcls1_def2 1);
       
    25 by( rtac finite_SigmaI 1);
       
    26 by(  rtac finite_is_class 1);
       
    27 by( rtac finite_subset 1);
       
    28 by(  rtac some_subset_the 1);
       
    29 by( Simp_tac 1);
       
    30 qed "finite_subcls1";
       
    31 
       
    32 fun prove_typerel_lemma drules indrule s = prove_goal thy s (fn prems => [
       
    33 	rtac (hd prems RS indrule) 1,
       
    34 	auto_tac (claset() addDs drules, simpset())]);
       
    35 
       
    36 fun prove_typerel s lemmata = prove_goal thy s (fn prems => [
       
    37 	cut_facts_tac prems 1,
       
    38 	auto_tac (claset() addDs lemmata, simpset())]);
       
    39 
       
    40 
       
    41 (*#### patch for Isabelle98-1*)
       
    42 val major::prems = goal Trancl.thy
       
    43  "\\<lbrakk> (x,y) \\<in> r^+; \
       
    44 \    \\<And>x y. (x,y) \\<in> r \\<Longrightarrow> P x y; \
       
    45 \    \\<And>x y z. \\<lbrakk> (x,y) \\<in> r^+; P x y; (y,z) \\<in> r^+; P y z \\<rbrakk> \\<Longrightarrow> P x z \
       
    46 \ \\<rbrakk> \\<Longrightarrow> P x y";
       
    47 by(blast_tac (claset() addIs ([r_into_trancl,major RS trancl_induct]@prems))1);
       
    48 qed "trancl_trans_induct";
       
    49 
       
    50 Goalw [is_class_def] "G\\<turnstile>C\\<prec>C D \\<Longrightarrow> is_class G C";
       
    51 by(etac trancl_trans_induct 1);
       
    52 by (auto_tac (HOL_cs addSDs [subcls1D],simpset()));
       
    53 qed "subcls_is_class";
       
    54 
       
    55 
       
    56 (* A particular thm about wf;
       
    57    looks like it is an odd instance of something more general
       
    58 *)
       
    59 Goalw [wf_def] "wf{((A,x),(B,y)) . A=B \\<and> wf(R(A)) \\<and> (x,y)\\<in>R(A)}";
       
    60 by(full_simp_tac (simpset() delcongs [imp_cong] addsimps [split_paired_All]) 1);
       
    61 by(strip_tac 1);
       
    62 by(rename_tac "A x" 1);
       
    63 by(case_tac "wf(R A)" 1);
       
    64 by (eres_inst_tac [("a","x")] wf_induct 1);
       
    65 by (EVERY1[etac allE, etac allE, etac mp, rtac allI, rtac allI]);
       
    66 by (Fast_tac 1);
       
    67 by(rewrite_goals_tac [wf_def]);
       
    68 by(Blast_tac 1);
       
    69 val wf_rel_lemma = result();
       
    70 
       
    71 
       
    72 (* Proving the termination conditions *)
       
    73 
       
    74 goalw thy [subcls1_rel_def] "wf subcls1_rel";
       
    75 by(rtac (wf_rel_lemma RS wf_subset) 1);
       
    76 by(Force_tac 1);
       
    77 val wf_subcls1_rel = result();
       
    78 
       
    79 val cmethd_TC = prove_goalw_cterm [subcls1_rel_def]
       
    80  (cterm_of (sign_of thy) (HOLogic.mk_Trueprop (hd (tl (cmethd.tcs)))))
       
    81  (K [auto_tac (claset() addIs [subcls1I], simpset())]);
       
    82 
       
    83 val fields_TC = prove_goalw_cterm [subcls1_rel_def]
       
    84  (cterm_of (sign_of thy) (HOLogic.mk_Trueprop (hd (tl (fields.tcs)))))
       
    85  (K [auto_tac (claset() addIs [subcls1I], simpset())]);
       
    86 
       
    87 
       
    88 AddSIs   [widen.refl];
       
    89 Addsimps [widen.refl];
       
    90 
       
    91 val prove_widen_lemma = prove_typerel_lemma [] widen.elim;
       
    92 
       
    93 val widen_PrimT_RefT = prove_typerel "G\\<turnstile>PrimT x\\<preceq>RefT tname \\<Longrightarrow> R"
       
    94  [ prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S = PrimT x \\<longrightarrow> T = RefT tname \\<longrightarrow> R"];
       
    95 
       
    96 
       
    97 val widen_RefT = prove_typerel "G\\<turnstile>RefT R\\<preceq>T \\<Longrightarrow> \\<exists>t. T=RefT t" 
       
    98 	[prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S=RefT R \\<longrightarrow> (\\<exists>t. T=RefT t)"];
       
    99 
       
   100 val widen_RefT2 = prove_typerel "G\\<turnstile>S\\<preceq>RefT R \\<Longrightarrow> \\<exists>t. S=RefT t" 
       
   101 	[prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> T=RefT R \\<longrightarrow> (\\<exists>t. S=RefT t)"];
       
   102 
       
   103 val widen_Class = prove_typerel "G\\<turnstile>Class C\\<preceq>T \\<Longrightarrow> \\<exists>D. T=Class D"
       
   104  [ prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S = Class C \\<longrightarrow> (\\<exists>D. T=Class D)"];
       
   105 
       
   106 val widen_Class_RefT = prove_typerel 
       
   107 	"G\\<turnstile>Class C\\<preceq>RefT t \\<Longrightarrow> (\\<exists>tname. t=ClassT tname)" 
       
   108  [prove_widen_lemma 
       
   109  "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S=Class C \\<longrightarrow> T=RefT t \\<longrightarrow> (\\<exists>tname. t=ClassT tname)"];
       
   110 
       
   111 val widen_Class_NullT = prove_typerel "G\\<turnstile>Class C\\<preceq>RefT NullT \\<Longrightarrow> R" 
       
   112  [prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S=Class C \\<longrightarrow> T=RefT NullT \\<longrightarrow> R"];
       
   113 
       
   114 val widen_Class_Class = prove_typerel "G\\<turnstile>Class C\\<preceq>Class cm \\<Longrightarrow> C=cm |  G\\<turnstile>C\\<prec>C cm"
       
   115 [ prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S = Class C \\<longrightarrow> T = Class cm \\<longrightarrow> C=cm |  G\\<turnstile>C\\<prec>C cm"];
       
   116 
       
   117 Goal "\\<lbrakk>G\\<turnstile>S\\<preceq>U; \\<forall>C. is_class G C \\<longrightarrow> G\\<turnstile>Class C\\<preceq>Class Object;\
       
   118 \\\<forall>C. G\\<turnstile>Object\\<prec>C C \\<longrightarrow> False \\<rbrakk> \\<Longrightarrow> \\<forall>T. G\\<turnstile>U\\<preceq>T \\<longrightarrow> G\\<turnstile>S\\<preceq>T";
       
   119 by( etac widen.induct 1);
       
   120 by   Safe_tac;
       
   121 by(  ALLGOALS (forward_tac [widen_Class, widen_RefT]));
       
   122 by  Safe_tac;
       
   123 by(  rtac widen.null 2);
       
   124 by( forward_tac [widen_Class_Class] 1);
       
   125 by Safe_tac;
       
   126 by(  ALLGOALS(EVERY'[etac thin_rl,etac thin_rl,
       
   127 	fast_tac (claset() addIs [widen.subcls,trancl_trans])]));
       
   128 qed_spec_mp "widen_trans_lemma";
       
   129 
       
   130 
       
   131 val prove_cast_lemma = prove_typerel_lemma [] cast.elim;
       
   132 
       
   133 val cast_RefT = prove_typerel "G\\<turnstile>RefT R\\<Rightarrow>? T \\<Longrightarrow> \\<exists>t. T=RefT t" 
       
   134 	[prove_typerel_lemma [widen_RefT] cast.elim 
       
   135 	"G\\<turnstile>S\\<Rightarrow>? T \\<Longrightarrow> S=RefT R \\<longrightarrow> (\\<exists>t. T=RefT t)"];
       
   136 
       
   137 val cast_RefT2 = prove_typerel "G\\<turnstile>S\\<Rightarrow>? RefT R \\<Longrightarrow> \\<exists>t. S=RefT t" 
       
   138 	[prove_typerel_lemma [widen_RefT2] cast.elim 
       
   139 	"G\\<turnstile>S\\<Rightarrow>? T \\<Longrightarrow> T=RefT R \\<longrightarrow> (\\<exists>t. S=RefT t)"];
       
   140 
       
   141 val cast_PrimT2 = prove_typerel "G\\<turnstile>S\\<Rightarrow>? PrimT pt \\<Longrightarrow> G\\<turnstile>S\\<preceq>PrimT pt" 
       
   142  [prove_cast_lemma "G\\<turnstile>S\\<Rightarrow>? T \\<Longrightarrow> T=PrimT pt \\<longrightarrow> G\\<turnstile>S\\<preceq>PrimT pt"];
       
   143