| author | chaieb | 
| Thu, 23 Jul 2009 22:25:09 +0200 | |
| changeset 32162 | c6a045559ce6 | 
| parent 23757 | 087b0a241557 | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/J/WellForm.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: David von Oheimb | |
| 4 | Copyright 1999 Technische Universitaet Muenchen | |
| 11070 | 5 | *) | 
| 8011 | 6 | |
| 12911 | 7 | header {* \isaheader{Well-formedness of Java programs} *}
 | 
| 8011 | 8 | |
| 16417 | 9 | theory WellForm imports TypeRel SystemClasses begin | 
| 8011 | 10 | |
| 11070 | 11 | text {*
 | 
| 12 | for static checks on expressions and statements, see WellType. | |
| 13 | ||
| 14 | \begin{description}
 | |
| 15 | \item[improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):]\ \\ | |
| 16 | \begin{itemize}
 | |
| 17 | \item a method implementing or overwriting another method may have a result type | |
| 18 | that widens to the result type of the other method (instead of identical type) | |
| 19 | \end{itemize}
 | |
| 20 | ||
| 21 | \item[simplifications:]\ \\ | |
| 22 | \begin{itemize}
 | |
| 23 | \item for uniformity, Object is assumed to be declared like any other class | |
| 24 | \end{itemize}
 | |
| 25 | \end{description}
 | |
| 26 | *} | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 27 | types 'c wf_mb = "'c prog => cname => 'c mdecl => bool" | 
| 8011 | 28 | |
| 29 | constdefs | |
| 14045 | 30 | wf_syscls :: "'c prog => bool" | 
| 31 | "wf_syscls G == let cs = set G in Object \<in> fst ` cs \<and> (\<forall>x. Xcpt x \<in> fst ` cs)" | |
| 32 | ||
| 10069 | 33 | wf_fdecl :: "'c prog => fdecl => bool" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 34 | "wf_fdecl G == \<lambda>(fn,ft). is_type G ft" | 
| 8011 | 35 | |
| 10069 | 36 | wf_mhead :: "'c prog => sig => ty => bool" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 37 | "wf_mhead G == \<lambda>(mn,pTs) rT. (\<forall>T\<in>set pTs. is_type G T) \<and> is_type G rT" | 
| 8011 | 38 | |
| 14045 | 39 | ws_cdecl :: "'c prog => 'c cdecl => bool" | 
| 40 | "ws_cdecl G == | |
| 41 | \<lambda>(C,(D,fs,ms)). | |
| 42 | (\<forall>f\<in>set fs. wf_fdecl G f) \<and> unique fs \<and> | |
| 43 | (\<forall>(sig,rT,mb)\<in>set ms. wf_mhead G sig rT) \<and> unique ms \<and> | |
| 44 | (C \<noteq> Object \<longrightarrow> is_class G D \<and> \<not>G\<turnstile>D\<preceq>C C)" | |
| 45 | ||
| 46 | ws_prog :: "'c prog => bool" | |
| 47 | "ws_prog G == | |
| 48 | wf_syscls G \<and> (\<forall>c\<in>set G. ws_cdecl G c) \<and> unique G" | |
| 49 | ||
| 50 | wf_mrT :: "'c prog => 'c cdecl => bool" | |
| 51 | "wf_mrT G == | |
| 52 | \<lambda>(C,(D,fs,ms)). | |
| 53 | (C \<noteq> Object \<longrightarrow> (\<forall>(sig,rT,b)\<in>set ms. \<forall>D' rT' b'. | |
| 54 | method(G,D) sig = Some(D',rT',b') --> G\<turnstile>rT\<preceq>rT'))" | |
| 55 | ||
| 56 | wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" | |
| 57 | "wf_cdecl_mdecl wf_mb G == | |
| 58 | \<lambda>(C,(D,fs,ms)). (\<forall>m\<in>set ms. wf_mb G C m)" | |
| 59 | ||
| 60 | wf_prog :: "'c wf_mb => 'c prog => bool" | |
| 61 | "wf_prog wf_mb G == | |
| 62 | ws_prog G \<and> (\<forall>c\<in> set G. wf_mrT G c \<and> wf_cdecl_mdecl wf_mb G c)" | |
| 63 | ||
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 64 | wf_mdecl :: "'c wf_mb => 'c wf_mb" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 65 | "wf_mdecl wf_mb G C == \<lambda>(sig,rT,mb). wf_mhead G sig rT \<and> wf_mb G C (sig,rT,mb)" | 
| 8011 | 66 | |
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 67 | wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" | 
| 10042 | 68 | "wf_cdecl wf_mb G == | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 69 | \<lambda>(C,(D,fs,ms)). | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 70 | (\<forall>f\<in>set fs. wf_fdecl G f) \<and> unique fs \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 71 | (\<forall>m\<in>set ms. wf_mdecl wf_mb G C m) \<and> unique ms \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 72 | (C \<noteq> Object \<longrightarrow> is_class G D \<and> \<not>G\<turnstile>D\<preceq>C C \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 73 | (\<forall>(sig,rT,b)\<in>set ms. \<forall>D' rT' b'. | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 74 | method(G,D) sig = Some(D',rT',b') --> G\<turnstile>rT\<preceq>rT'))" | 
| 8011 | 75 | |
| 14045 | 76 | lemma wf_cdecl_mrT_cdecl_mdecl: | 
| 77 | "(wf_cdecl wf_mb G c) = (ws_cdecl G c \<and> wf_mrT G c \<and> wf_cdecl_mdecl wf_mb G c)" | |
| 78 | apply (rule iffI) | |
| 79 | apply (simp add: wf_cdecl_def ws_cdecl_def wf_mrT_def wf_cdecl_mdecl_def | |
| 80 | wf_mdecl_def wf_mhead_def split_beta)+ | |
| 81 | done | |
| 12951 | 82 | |
| 14045 | 83 | lemma wf_cdecl_ws_cdecl [intro]: "wf_cdecl wf_mb G cd \<Longrightarrow> ws_cdecl G cd" | 
| 84 | by (simp add: wf_cdecl_mrT_cdecl_mdecl) | |
| 85 | ||
| 86 | lemma wf_prog_ws_prog [intro]: "wf_prog wf_mb G \<Longrightarrow> ws_prog G" | |
| 87 | by (simp add: wf_prog_def ws_prog_def) | |
| 88 | ||
| 89 | lemma wf_prog_wf_mdecl: | |
| 90 | "\<lbrakk> wf_prog wf_mb G; (C,S,fs,mdecls) \<in> set G; ((mn,pTs),rT,code) \<in> set mdecls\<rbrakk> | |
| 91 | \<Longrightarrow> wf_mdecl wf_mb G C ((mn,pTs),rT,code)" | |
| 92 | by (auto simp add: wf_prog_def ws_prog_def wf_mdecl_def | |
| 93 | wf_cdecl_mdecl_def ws_cdecl_def) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 94 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 95 | lemma class_wf: | 
| 14045 | 96 | "[|class G C = Some c; wf_prog wf_mb G|] | 
| 97 | ==> wf_cdecl wf_mb G (C,c) \<and> wf_mrT G (C,c)" | |
| 98 | apply (unfold wf_prog_def ws_prog_def wf_cdecl_def class_def) | |
| 99 | apply clarify | |
| 100 | apply (drule_tac x="(C,c)" in bspec, fast dest: map_of_SomeD) | |
| 101 | apply (drule_tac x="(C,c)" in bspec, fast dest: map_of_SomeD) | |
| 102 | apply (simp add: wf_cdecl_def ws_cdecl_def wf_mdecl_def | |
| 103 | wf_cdecl_mdecl_def wf_mrT_def split_beta) | |
| 104 | done | |
| 105 | ||
| 106 | lemma class_wf_struct: | |
| 107 | "[|class G C = Some c; ws_prog G|] | |
| 108 | ==> ws_cdecl G (C,c)" | |
| 109 | apply (unfold ws_prog_def class_def) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 110 | apply (fast dest: map_of_SomeD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 111 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 112 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 113 | lemma class_Object [simp]: | 
| 14045 | 114 | "ws_prog G ==> \<exists>X fs ms. class G Object = Some (X,fs,ms)" | 
| 115 | apply (unfold ws_prog_def wf_syscls_def class_def) | |
| 12951 | 116 | apply (auto simp: map_of_SomeI) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 117 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 118 | |
| 14045 | 119 | lemma class_Object_syscls [simp]: | 
| 120 | "wf_syscls G ==> unique G \<Longrightarrow> \<exists>X fs ms. class G Object = Some (X,fs,ms)" | |
| 121 | apply (unfold wf_syscls_def class_def) | |
| 122 | apply (auto simp: map_of_SomeI) | |
| 123 | done | |
| 124 | ||
| 125 | lemma is_class_Object [simp]: "ws_prog G ==> is_class G Object" | |
| 18576 | 126 | by (simp add: is_class_def) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 127 | |
| 14045 | 128 | lemma is_class_xcpt [simp]: "ws_prog G \<Longrightarrow> is_class G (Xcpt x)" | 
| 129 | apply (simp add: ws_prog_def wf_syscls_def) | |
| 13051 | 130 | apply (simp add: is_class_def class_def) | 
| 131 | apply clarify | |
| 132 | apply (erule_tac x = x in allE) | |
| 133 | apply clarify | |
| 134 | apply (auto intro!: map_of_SomeI) | |
| 135 | done | |
| 136 | ||
| 22271 | 137 | lemma subcls1_wfD: "[|G\<turnstile>C\<prec>C1D; ws_prog G|] ==> D \<noteq> C \<and> \<not> (subcls1 G)^++ D C" | 
| 23757 | 138 | apply( frule tranclp.r_into_trancl [where r="subcls1 G"]) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 139 | apply( drule subcls1D) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 140 | apply(clarify) | 
| 14045 | 141 | apply( drule (1) class_wf_struct) | 
| 142 | apply( unfold ws_cdecl_def) | |
| 23757 | 143 | apply(force simp add: reflcl_tranclp [THEN sym] simp del: reflcl_tranclp) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 144 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 145 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 146 | lemma wf_cdecl_supD: | 
| 14045 | 147 | "!!r. \<lbrakk>ws_cdecl G (C,D,r); C \<noteq> Object\<rbrakk> \<Longrightarrow> is_class G D" | 
| 148 | apply (unfold ws_cdecl_def) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 149 | apply (auto split add: option.split_asm) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 150 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 151 | |
| 22271 | 152 | lemma subcls_asym: "[|ws_prog G; (subcls1 G)^++ C D|] ==> \<not> (subcls1 G)^++ D C" | 
| 23757 | 153 | apply(erule tranclp.cases) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 154 | apply(fast dest!: subcls1_wfD ) | 
| 23757 | 155 | apply(fast dest!: subcls1_wfD intro: tranclp_trans) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 156 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 157 | |
| 22271 | 158 | lemma subcls_irrefl: "[|ws_prog G; (subcls1 G)^++ C D|] ==> C \<noteq> D" | 
| 23757 | 159 | apply (erule tranclp_trans_induct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 160 | apply (auto dest: subcls1_wfD subcls_asym) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 161 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 162 | |
| 22271 | 163 | lemma acyclic_subcls1: "ws_prog G ==> acyclicP (subcls1 G)" | 
| 164 | apply (simp add: acyclic_def [to_pred]) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 165 | apply (fast dest: subcls_irrefl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 166 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 167 | |
| 22271 | 168 | lemma wf_subcls1: "ws_prog G ==> wfP ((subcls1 G)^--1)" | 
| 169 | apply (rule finite_acyclic_wf [to_pred]) | |
| 170 | apply ( subst finite_converse [to_pred]) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 171 | apply ( rule finite_subcls1) | 
| 22271 | 172 | apply (subst acyclic_converse [to_pred]) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 173 | apply (erule acyclic_subcls1) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 174 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 175 | |
| 14045 | 176 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 177 | lemma subcls_induct: | 
| 22271 | 178 | "[|wf_prog wf_mb G; !!C. \<forall>D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 179 | (is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 180 | proof - | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 181 | assume p: "PROP ?P" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 182 | assume ?A thus ?thesis apply - | 
| 14045 | 183 | apply (drule wf_prog_ws_prog) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 184 | apply(drule wf_subcls1) | 
| 22271 | 185 | apply(drule wfP_trancl) | 
| 23757 | 186 | apply(simp only: tranclp_converse) | 
| 22271 | 187 | apply(erule_tac a = C in wfP_induct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 188 | apply(rule p) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 189 | apply(auto) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 190 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 191 | qed | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 192 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 193 | lemma subcls1_induct: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 194 | "[|is_class G C; wf_prog wf_mb G; P Object; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 195 | !!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 196 | wf_cdecl wf_mb G (C,D,fs,ms) \<and> G\<turnstile>C\<prec>C1D \<and> is_class G D \<and> P D|] ==> P C | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 197 | |] ==> P C" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 198 | (is "?A \<Longrightarrow> ?B \<Longrightarrow> ?C \<Longrightarrow> PROP ?P \<Longrightarrow> _") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 199 | proof - | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 200 | assume p: "PROP ?P" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 201 | assume ?A ?B ?C thus ?thesis apply - | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 202 | apply(unfold is_class_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 203 | apply( rule impE) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 204 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 205 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 206 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 207 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 208 | apply( erule thin_rl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 209 | apply( rule subcls_induct) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 210 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 211 | apply( rule impI) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 212 | apply( case_tac "C = Object") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 213 | apply( fast) | 
| 18447 | 214 | apply auto | 
| 14045 | 215 | apply( frule (1) class_wf) apply (erule conjE)+ | 
| 216 | apply (frule wf_cdecl_ws_cdecl) | |
| 217 | apply( frule (1) wf_cdecl_supD) | |
| 218 | ||
| 219 | apply( subgoal_tac "G\<turnstile>C\<prec>C1a") | |
| 220 | apply( erule_tac [2] subcls1I) | |
| 221 | apply( rule p) | |
| 222 | apply (unfold is_class_def) | |
| 223 | apply auto | |
| 224 | done | |
| 225 | qed | |
| 226 | ||
| 227 | lemma subcls_induct_struct: | |
| 22271 | 228 | "[|ws_prog G; !!C. \<forall>D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C" | 
| 14045 | 229 | (is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _") | 
| 230 | proof - | |
| 231 | assume p: "PROP ?P" | |
| 232 | assume ?A thus ?thesis apply - | |
| 233 | apply(drule wf_subcls1) | |
| 22271 | 234 | apply(drule wfP_trancl) | 
| 23757 | 235 | apply(simp only: tranclp_converse) | 
| 22271 | 236 | apply(erule_tac a = C in wfP_induct) | 
| 14045 | 237 | apply(rule p) | 
| 238 | apply(auto) | |
| 239 | done | |
| 240 | qed | |
| 241 | ||
| 242 | ||
| 243 | lemma subcls1_induct_struct: | |
| 244 | "[|is_class G C; ws_prog G; P Object; | |
| 245 | !!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and> | |
| 246 | ws_cdecl G (C,D,fs,ms) \<and> G\<turnstile>C\<prec>C1D \<and> is_class G D \<and> P D|] ==> P C | |
| 247 | |] ==> P C" | |
| 248 | (is "?A \<Longrightarrow> ?B \<Longrightarrow> ?C \<Longrightarrow> PROP ?P \<Longrightarrow> _") | |
| 249 | proof - | |
| 250 | assume p: "PROP ?P" | |
| 251 | assume ?A ?B ?C thus ?thesis apply - | |
| 252 | apply(unfold is_class_def) | |
| 253 | apply( rule impE) | |
| 254 | prefer 2 | |
| 255 | apply( assumption) | |
| 256 | prefer 2 | |
| 257 | apply( assumption) | |
| 258 | apply( erule thin_rl) | |
| 259 | apply( rule subcls_induct_struct) | |
| 260 | apply( assumption) | |
| 261 | apply( rule impI) | |
| 262 | apply( case_tac "C = Object") | |
| 263 | apply( fast) | |
| 18447 | 264 | apply auto | 
| 14045 | 265 | apply( frule (1) class_wf_struct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 266 | apply( frule (1) wf_cdecl_supD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 267 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 268 | apply( subgoal_tac "G\<turnstile>C\<prec>C1a") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 269 | apply( erule_tac [2] subcls1I) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 270 | apply( rule p) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 271 | apply (unfold is_class_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 272 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 273 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 274 | qed | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 275 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 276 | lemmas method_rec = wf_subcls1 [THEN [2] method_rec_lemma]; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 277 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 278 | lemmas fields_rec = wf_subcls1 [THEN [2] fields_rec_lemma]; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 279 | |
| 14045 | 280 | lemma field_rec: "\<lbrakk>class G C = Some (D, fs, ms); ws_prog G\<rbrakk> | 
| 13672 | 281 | \<Longrightarrow> field (G, C) = | 
| 282 | (if C = Object then empty else field (G, D)) ++ | |
| 283 | map_of (map (\<lambda>(s, f). (s, C, f)) fs)" | |
| 284 | apply (simp only: field_def) | |
| 285 | apply (frule fields_rec, assumption) | |
| 286 | apply (rule HOL.trans) | |
| 287 | apply (simp add: o_def) | |
| 288 | apply (simp (no_asm_use) | |
| 289 | add: split_beta split_def o_def map_compose [THEN sym] del: map_compose) | |
| 290 | done | |
| 291 | ||
| 12951 | 292 | lemma method_Object [simp]: | 
| 14045 | 293 | "method (G, Object) sig = Some (D, mh, code) \<Longrightarrow> ws_prog G \<Longrightarrow> D = Object" | 
| 12951 | 294 | apply (frule class_Object, clarify) | 
| 295 | apply (drule method_rec, assumption) | |
| 296 | apply (auto dest: map_of_SomeD) | |
| 297 | done | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 298 | |
| 13672 | 299 | |
| 14045 | 300 | lemma fields_Object [simp]: "\<lbrakk> ((vn, C), T) \<in> set (fields (G, Object)); ws_prog G \<rbrakk> | 
| 13672 | 301 | \<Longrightarrow> C = Object" | 
| 302 | apply (frule class_Object) | |
| 303 | apply clarify | |
| 304 | apply (subgoal_tac "fields (G, Object) = map (\<lambda>(fn,ft). ((fn,Object),ft)) fs") | |
| 305 | apply (simp add: image_iff split_beta) | |
| 306 | apply auto | |
| 307 | apply (rule trans) | |
| 308 | apply (rule fields_rec, assumption+) | |
| 309 | apply simp | |
| 310 | done | |
| 311 | ||
| 14045 | 312 | lemma subcls_C_Object: "[|is_class G C; ws_prog G|] ==> G\<turnstile>C\<preceq>C Object" | 
| 313 | apply(erule subcls1_induct_struct) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 314 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 315 | apply( fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 316 | apply(auto dest!: wf_cdecl_supD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 317 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 318 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 319 | lemma is_type_rTI: "wf_mhead G sig rT ==> is_type G rT" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 320 | apply (unfold wf_mhead_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 321 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 322 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 323 | |
| 14045 | 324 | lemma widen_fields_defpl': "[|is_class G C; ws_prog G|] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 325 | \<forall>((fn,fd),fT)\<in>set (fields (G,C)). G\<turnstile>C\<preceq>C fd" | 
| 14045 | 326 | apply( erule subcls1_induct_struct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 327 | apply( assumption) | 
| 12951 | 328 | apply( frule class_Object) | 
| 329 | apply( clarify) | |
| 330 | apply( frule fields_rec, assumption) | |
| 331 | apply( fastsimp) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 332 | apply( tactic "safe_tac HOL_cs") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 333 | apply( subst fields_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 334 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 335 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 336 | apply( simp (no_asm) split del: split_if) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 337 | apply( rule ballI) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 338 | apply( simp (no_asm_simp) only: split_tupled_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 339 | apply( simp (no_asm)) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 340 | apply( erule UnE) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 341 | apply( force) | 
| 23757 | 342 | apply( erule r_into_rtranclp [THEN rtranclp_trans]) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 343 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 344 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 345 | |
| 12517 | 346 | lemma widen_fields_defpl: | 
| 14045 | 347 | "[|((fn,fd),fT) \<in> set (fields (G,C)); ws_prog G; is_class G C|] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 348 | G\<turnstile>C\<preceq>C fd" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 349 | apply( drule (1) widen_fields_defpl') | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 350 | apply (fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 351 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 352 | |
| 12517 | 353 | lemma unique_fields: | 
| 14045 | 354 | "[|is_class G C; ws_prog G|] ==> unique (fields (G,C))" | 
| 355 | apply( erule subcls1_induct_struct) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 356 | apply( assumption) | 
| 12951 | 357 | apply( frule class_Object) | 
| 358 | apply( clarify) | |
| 359 | apply( frule fields_rec, assumption) | |
| 14045 | 360 | apply( drule class_wf_struct, assumption) | 
| 361 | apply( simp add: ws_cdecl_def) | |
| 12951 | 362 | apply( rule unique_map_inj) | 
| 363 | apply( simp) | |
| 13585 | 364 | apply( rule inj_onI) | 
| 12951 | 365 | apply( simp) | 
| 366 | apply( safe dest!: wf_cdecl_supD) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 367 | apply( drule subcls1_wfD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 368 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 369 | apply( subst fields_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 370 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 371 | apply( rotate_tac -1) | 
| 14045 | 372 | apply( frule class_wf_struct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 373 | apply auto | 
| 14045 | 374 | apply( simp add: ws_cdecl_def) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 375 | apply( erule unique_append) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 376 | apply( rule unique_map_inj) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 377 | apply( clarsimp) | 
| 13585 | 378 | apply (rule inj_onI) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 379 | apply( simp) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 380 | apply(auto dest!: widen_fields_defpl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 381 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 382 | |
| 12517 | 383 | lemma fields_mono_lemma [rule_format (no_asm)]: | 
| 22271 | 384 | "[|ws_prog G; (subcls1 G)^** C' C|] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 385 | x \<in> set (fields (G,C)) --> x \<in> set (fields (G,C'))" | 
| 23757 | 386 | apply(erule converse_rtranclp_induct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 387 | apply( safe dest!: subcls1D) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 388 | apply(subst fields_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 389 | apply( auto) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 390 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 391 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 392 | lemma fields_mono: | 
| 14045 | 393 | "\<lbrakk>map_of (fields (G,C)) fn = Some f; G\<turnstile>D\<preceq>C C; is_class G D; ws_prog G\<rbrakk> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 394 | \<Longrightarrow> map_of (fields (G,D)) fn = Some f" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 395 | apply (rule map_of_SomeI) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 396 | apply (erule (1) unique_fields) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 397 | apply (erule (1) fields_mono_lemma) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 398 | apply (erule map_of_SomeD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 399 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 400 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 401 | lemma widen_cfs_fields: | 
| 14045 | 402 | "[|field (G,C) fn = Some (fd, fT); G\<turnstile>D\<preceq>C C; ws_prog G|]==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 403 | map_of (fields (G,D)) (fn, fd) = Some fT" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 404 | apply (drule field_fields) | 
| 23757 | 405 | apply (drule rtranclpD) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 406 | apply safe | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 407 | apply (frule subcls_is_class) | 
| 23757 | 408 | apply (drule tranclp_into_rtranclp) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 409 | apply (fast dest: fields_mono) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 410 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 411 | |
| 12517 | 412 | lemma method_wf_mdecl [rule_format (no_asm)]: | 
| 413 | "wf_prog wf_mb G ==> is_class G C \<Longrightarrow> | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 414 | method (G,C) sig = Some (md,mh,m) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 415 | --> G\<turnstile>C\<preceq>C md \<and> wf_mdecl wf_mb G md (sig,(mh,m))" | 
| 14045 | 416 | apply (frule wf_prog_ws_prog) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 417 | apply( erule subcls1_induct) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 418 | apply( assumption) | 
| 12951 | 419 | apply( clarify) | 
| 420 | apply( frule class_Object) | |
| 421 | apply( clarify) | |
| 422 | apply( frule method_rec, assumption) | |
| 423 | apply( drule class_wf, assumption) | |
| 424 | apply( simp add: wf_cdecl_def) | |
| 425 | apply( drule map_of_SomeD) | |
| 426 | apply( subgoal_tac "md = Object") | |
| 14045 | 427 | apply( fastsimp) | 
| 12951 | 428 | apply( fastsimp) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 429 | apply( clarify) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 430 | apply( frule_tac C = C in method_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 431 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 432 | apply( rotate_tac -1) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 433 | apply( simp) | 
| 14025 | 434 | apply( drule map_add_SomeD) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 435 | apply( erule disjE) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 436 | apply( erule_tac V = "?P --> ?Q" in thin_rl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 437 | apply (frule map_of_SomeD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 438 | apply (clarsimp simp add: wf_cdecl_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 439 | apply( clarify) | 
| 23757 | 440 | apply( rule rtranclp_trans) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 441 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 442 | apply( assumption) | 
| 23757 | 443 | apply( rule r_into_rtranclp) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 444 | apply( fast intro: subcls1I) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 445 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 446 | |
| 13672 | 447 | |
| 14045 | 448 | lemma method_wf_mhead [rule_format (no_asm)]: | 
| 449 | "ws_prog G ==> is_class G C \<Longrightarrow> | |
| 450 | method (G,C) sig = Some (md,rT,mb) | |
| 451 | --> G\<turnstile>C\<preceq>C md \<and> wf_mhead G sig rT" | |
| 452 | apply( erule subcls1_induct_struct) | |
| 453 | apply( assumption) | |
| 454 | apply( clarify) | |
| 455 | apply( frule class_Object) | |
| 456 | apply( clarify) | |
| 457 | apply( frule method_rec, assumption) | |
| 458 | apply( drule class_wf_struct, assumption) | |
| 459 | apply( simp add: ws_cdecl_def) | |
| 460 | apply( drule map_of_SomeD) | |
| 461 | apply( subgoal_tac "md = Object") | |
| 462 | apply( fastsimp) | |
| 463 | apply( fastsimp) | |
| 464 | apply( clarify) | |
| 465 | apply( frule_tac C = C in method_rec) | |
| 466 | apply( assumption) | |
| 467 | apply( rotate_tac -1) | |
| 468 | apply( simp) | |
| 469 | apply( drule map_add_SomeD) | |
| 470 | apply( erule disjE) | |
| 471 | apply( erule_tac V = "?P --> ?Q" in thin_rl) | |
| 472 | apply (frule map_of_SomeD) | |
| 473 | apply (clarsimp simp add: ws_cdecl_def) | |
| 474 | apply blast | |
| 475 | apply clarify | |
| 23757 | 476 | apply( rule rtranclp_trans) | 
| 14045 | 477 | prefer 2 | 
| 478 | apply( assumption) | |
| 23757 | 479 | apply( rule r_into_rtranclp) | 
| 14045 | 480 | apply( fast intro: subcls1I) | 
| 13672 | 481 | done | 
| 482 | ||
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 483 | lemma subcls_widen_methd [rule_format (no_asm)]: | 
| 14045 | 484 | "[|G\<turnstile>T'\<preceq>C T; wf_prog wf_mb G|] ==> | 
| 485 | \<forall>D rT b. method (G,T) sig = Some (D,rT ,b) --> | |
| 486 | (\<exists>D' rT' b'. method (G,T') sig = Some (D',rT',b') \<and> G\<turnstile>D'\<preceq>C D \<and> G\<turnstile>rT'\<preceq>rT)" | |
| 23757 | 487 | apply( drule rtranclpD) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 488 | apply( erule disjE) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 489 | apply( fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 490 | apply( erule conjE) | 
| 23757 | 491 | apply( erule tranclp_trans_induct) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 492 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 493 | apply( clarify) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 494 | apply( drule spec, drule spec, drule spec, erule (1) impE) | 
| 23757 | 495 | apply( fast elim: widen_trans rtranclp_trans) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 496 | apply( clarify) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 497 | apply( drule subcls1D) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 498 | apply( clarify) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 499 | apply( subst method_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 500 | apply( assumption) | 
| 14025 | 501 | apply( unfold map_add_def) | 
| 14045 | 502 | apply( simp (no_asm_simp) add: wf_prog_ws_prog del: split_paired_Ex) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 503 | apply( case_tac "\<exists>z. map_of(map (\<lambda>(s,m). (s, ?C, m)) ms) sig = Some z") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 504 | apply( erule exE) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 505 | apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 506 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 507 | apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl) | 
| 22633 | 508 | apply(  tactic "asm_full_simp_tac (HOL_ss addsimps [@{thm not_None_eq} RS sym]) 1")
 | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 509 | apply( simp_all (no_asm_simp) del: split_paired_Ex) | 
| 14045 | 510 | apply( frule (1) class_wf) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 511 | apply( simp (no_asm_simp) only: split_tupled_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 512 | apply( unfold wf_cdecl_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 513 | apply( drule map_of_SomeD) | 
| 14045 | 514 | apply (auto simp add: wf_mrT_def) | 
| 23757 | 515 | apply (rule rtranclp_trans) | 
| 14045 | 516 | defer | 
| 517 | apply (rule method_wf_mhead [THEN conjunct1]) | |
| 518 | apply (simp only: wf_prog_def) | |
| 519 | apply (simp add: is_class_def) | |
| 520 | apply assumption | |
| 521 | apply (auto intro: subcls1I) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 522 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 523 | |
| 14045 | 524 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 525 | lemma subtype_widen_methd: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 526 | "[| G\<turnstile> C\<preceq>C D; wf_prog wf_mb G; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 527 | method (G,D) sig = Some (md, rT, b) |] | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 528 | ==> \<exists>mD' rT' b'. method (G,C) sig= Some(mD',rT',b') \<and> G\<turnstile>rT'\<preceq>rT" | 
| 14045 | 529 | apply(auto dest: subcls_widen_methd | 
| 12517 | 530 | simp add: wf_mdecl_def wf_mhead_def split_def) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 531 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 532 | |
| 14045 | 533 | |
| 12517 | 534 | lemma method_in_md [rule_format (no_asm)]: | 
| 14045 | 535 | "ws_prog G ==> is_class G C \<Longrightarrow> \<forall>D. method (G,C) sig = Some(D,mh,code) | 
| 12517 | 536 | --> is_class G D \<and> method (G,D) sig = Some(D,mh,code)" | 
| 14045 | 537 | apply (erule (1) subcls1_induct_struct) | 
| 12951 | 538 | apply clarify | 
| 539 | apply (frule method_Object, assumption) | |
| 540 | apply hypsubst | |
| 541 | apply simp | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 542 | apply (erule conjE) | 
| 15481 | 543 | apply (simplesubst method_rec, assumption+) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 544 | apply (clarify) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14045diff
changeset | 545 | apply (erule_tac x = "Da" in allE) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 546 | apply (clarsimp) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 547 | apply (simp add: map_of_map) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 548 | apply (clarify) | 
| 15481 | 549 | apply (subst method_rec, assumption+) | 
| 14025 | 550 | apply (simp add: map_add_def map_of_map split add: option.split) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 551 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 552 | |
| 13672 | 553 | |
| 14045 | 554 | lemma method_in_md_struct [rule_format (no_asm)]: | 
| 555 | "ws_prog G ==> is_class G C \<Longrightarrow> \<forall>D. method (G,C) sig = Some(D,mh,code) | |
| 556 | --> is_class G D \<and> method (G,D) sig = Some(D,mh,code)" | |
| 557 | apply (erule (1) subcls1_induct_struct) | |
| 558 | apply clarify | |
| 559 | apply (frule method_Object, assumption) | |
| 560 | apply hypsubst | |
| 561 | apply simp | |
| 562 | apply (erule conjE) | |
| 15481 | 563 | apply (simplesubst method_rec, assumption+) | 
| 14045 | 564 | apply (clarify) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14045diff
changeset | 565 | apply (erule_tac x = "Da" in allE) | 
| 14045 | 566 | apply (clarsimp) | 
| 567 | apply (simp add: map_of_map) | |
| 568 | apply (clarify) | |
| 15481 | 569 | apply (subst method_rec, assumption+) | 
| 14045 | 570 | apply (simp add: map_add_def map_of_map split add: option.split) | 
| 571 | done | |
| 572 | ||
| 13672 | 573 | lemma fields_in_fd [rule_format (no_asm)]: "\<lbrakk> wf_prog wf_mb G; is_class G C\<rbrakk> | 
| 574 | \<Longrightarrow> \<forall> vn D T. (((vn,D),T) \<in> set (fields (G,C)) | |
| 575 | \<longrightarrow> (is_class G D \<and> ((vn,D),T) \<in> set (fields (G,D))))" | |
| 576 | apply (erule (1) subcls1_induct) | |
| 577 | ||
| 578 | apply clarify | |
| 14045 | 579 | apply (frule wf_prog_ws_prog) | 
| 13672 | 580 | apply (frule fields_Object, assumption+) | 
| 581 | apply (simp only: is_class_Object) apply simp | |
| 582 | ||
| 583 | apply clarify | |
| 584 | apply (frule fields_rec) | |
| 14045 | 585 | apply (simp (no_asm_simp) add: wf_prog_ws_prog) | 
| 13672 | 586 | |
| 587 | apply (case_tac "Da=C") | |
| 588 | apply blast (* case Da=C *) | |
| 589 | ||
| 590 | apply (subgoal_tac "((vn, Da), T) \<in> set (fields (G, D))") apply blast | |
| 591 | apply (erule thin_rl) | |
| 592 | apply (rotate_tac 1) | |
| 593 | apply (erule thin_rl, erule thin_rl, erule thin_rl, | |
| 594 | erule thin_rl, erule thin_rl, erule thin_rl) | |
| 595 | apply auto | |
| 596 | done | |
| 597 | ||
| 598 | lemma field_in_fd [rule_format (no_asm)]: "\<lbrakk> wf_prog wf_mb G; is_class G C\<rbrakk> | |
| 599 | \<Longrightarrow> \<forall> vn D T. (field (G,C) vn = Some(D,T) | |
| 600 | \<longrightarrow> is_class G D \<and> field (G,D) vn = Some(D,T))" | |
| 601 | apply (erule (1) subcls1_induct) | |
| 602 | ||
| 603 | apply clarify | |
| 604 | apply (frule field_fields) | |
| 605 | apply (drule map_of_SomeD) | |
| 14045 | 606 | apply (frule wf_prog_ws_prog) | 
| 13672 | 607 | apply (drule fields_Object, assumption+) | 
| 608 | apply simp | |
| 609 | ||
| 610 | apply clarify | |
| 611 | apply (subgoal_tac "((field (G, D)) ++ map_of (map (\<lambda>(s, f). (s, C, f)) fs)) vn = Some (Da, T)") | |
| 14025 | 612 | apply (simp (no_asm_use) only: map_add_Some_iff) | 
| 13672 | 613 | apply (erule disjE) | 
| 614 | apply (simp (no_asm_use) add: map_of_map) apply blast | |
| 615 | apply blast | |
| 616 | apply (rule trans [THEN sym], rule sym, assumption) | |
| 617 | apply (rule_tac x=vn in fun_cong) | |
| 618 | apply (rule trans, rule field_rec, assumption+) | |
| 14045 | 619 | apply (simp (no_asm_simp) add: wf_prog_ws_prog) | 
| 13672 | 620 | apply (simp (no_asm_use)) apply blast | 
| 621 | done | |
| 622 | ||
| 623 | ||
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 624 | lemma widen_methd: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 625 | "[| method (G,C) sig = Some (md,rT,b); wf_prog wf_mb G; G\<turnstile>T''\<preceq>C C|] | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 626 | ==> \<exists>md' rT' b'. method (G,T'') sig = Some (md',rT',b') \<and> G\<turnstile>rT'\<preceq>rT" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 627 | apply( drule subcls_widen_methd) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 628 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 629 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 630 | |
| 13672 | 631 | lemma widen_field: "\<lbrakk> (field (G,C) fn) = Some (fd, fT); wf_prog wf_mb G; is_class G C \<rbrakk> | 
| 632 | \<Longrightarrow> G\<turnstile>C\<preceq>C fd" | |
| 633 | apply (rule widen_fields_defpl) | |
| 634 | apply (simp add: field_def) | |
| 635 | apply (rule map_of_SomeD) | |
| 636 | apply (rule table_of_remap_SomeD) | |
| 637 | apply assumption+ | |
| 14045 | 638 | apply (simp (no_asm_simp) add: wf_prog_ws_prog)+ | 
| 13672 | 639 | done | 
| 640 | ||
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 641 | lemma Call_lemma: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 642 | "[|method (G,C) sig = Some (md,rT,b); G\<turnstile>T''\<preceq>C C; wf_prog wf_mb G; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 643 | class G C = Some y|] ==> \<exists>T' rT' b. method (G,T'') sig = Some (T',rT',b) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 644 | G\<turnstile>rT'\<preceq>rT \<and> G\<turnstile>T''\<preceq>C T' \<and> wf_mhead G sig rT' \<and> wf_mb G T' (sig,rT',b)" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 645 | apply( drule (2) widen_methd) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 646 | apply( clarify) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 647 | apply( frule subcls_is_class2) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 648 | apply (unfold is_class_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 649 | apply (simp (no_asm_simp)) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 650 | apply( drule method_wf_mdecl) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 651 | apply( unfold wf_mdecl_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 652 | apply( unfold is_class_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 653 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 654 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 655 | |
| 12517 | 656 | lemma fields_is_type_lemma [rule_format (no_asm)]: | 
| 14045 | 657 | "[|is_class G C; ws_prog G|] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 658 | \<forall>f\<in>set (fields (G,C)). is_type G (snd f)" | 
| 14045 | 659 | apply( erule (1) subcls1_induct_struct) | 
| 12951 | 660 | apply( frule class_Object) | 
| 661 | apply( clarify) | |
| 662 | apply( frule fields_rec, assumption) | |
| 14045 | 663 | apply( drule class_wf_struct, assumption) | 
| 664 | apply( simp add: ws_cdecl_def wf_fdecl_def) | |
| 12951 | 665 | apply( fastsimp) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 666 | apply( subst fields_rec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 667 | apply( fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 668 | apply( assumption) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 669 | apply( clarsimp) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 670 | apply( safe) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 671 | prefer 2 | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 672 | apply( force) | 
| 14045 | 673 | apply( drule (1) class_wf_struct) | 
| 674 | apply( unfold ws_cdecl_def) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 675 | apply( clarsimp) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 676 | apply( drule (1) bspec) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 677 | apply( unfold wf_fdecl_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 678 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 679 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 680 | |
| 14045 | 681 | |
| 12517 | 682 | lemma fields_is_type: | 
| 14045 | 683 | "[|map_of (fields (G,C)) fn = Some f; ws_prog G; is_class G C|] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 684 | is_type G f" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 685 | apply(drule map_of_SomeD) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 686 | apply(drule (2) fields_is_type_lemma) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 687 | apply(auto) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 688 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 689 | |
| 14045 | 690 | |
| 691 | lemma field_is_type: "\<lbrakk> ws_prog G; is_class G C; field (G, C) fn = Some (fd, fT) \<rbrakk> | |
| 692 | \<Longrightarrow> is_type G fT" | |
| 693 | apply (frule_tac f="((fn, fd), fT)" in fields_is_type_lemma) | |
| 694 | apply (auto simp add: field_def dest: map_of_SomeD) | |
| 695 | done | |
| 696 | ||
| 697 | ||
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 698 | lemma methd: | 
| 14045 | 699 | "[| ws_prog G; (C,S,fs,mdecls) \<in> set G; (sig,rT,code) \<in> set mdecls |] | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 700 | ==> method (G,C) sig = Some(C,rT,code) \<and> is_class G C" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 701 | proof - | 
| 14045 | 702 | assume wf: "ws_prog G" and C: "(C,S,fs,mdecls) \<in> set G" and | 
| 12951 | 703 | m: "(sig,rT,code) \<in> set mdecls" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 704 | moreover | 
| 12951 | 705 | from wf C have "class G C = Some (S,fs,mdecls)" | 
| 14045 | 706 | by (auto simp add: ws_prog_def class_def is_class_def intro: map_of_SomeI) | 
| 12951 | 707 | moreover | 
| 708 | from wf C | |
| 14045 | 709 | have "unique mdecls" by (unfold ws_prog_def ws_cdecl_def) auto | 
| 12951 | 710 | hence "unique (map (\<lambda>(s,m). (s,C,m)) mdecls)" by (induct mdecls, auto) | 
| 711 | with m | |
| 712 | have "map_of (map (\<lambda>(s,m). (s,C,m)) mdecls) sig = Some (C,rT,code)" | |
| 713 | by (force intro: map_of_SomeI) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 714 | ultimately | 
| 12517 | 715 | show ?thesis by (auto simp add: is_class_def dest: method_rec) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10613diff
changeset | 716 | qed | 
| 8011 | 717 | |
| 12951 | 718 | |
| 719 | lemma wf_mb'E: | |
| 720 | "\<lbrakk> wf_prog wf_mb G; \<And>C S fs ms m.\<lbrakk>(C,S,fs,ms) \<in> set G; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_mb' G C m \<rbrakk> | |
| 721 | \<Longrightarrow> wf_prog wf_mb' G" | |
| 14045 | 722 | apply (simp only: wf_prog_def) | 
| 12951 | 723 | apply auto | 
| 14045 | 724 | apply (simp add: wf_cdecl_mdecl_def) | 
| 12951 | 725 | apply safe | 
| 726 | apply (drule bspec, assumption) apply simp | |
| 727 | done | |
| 728 | ||
| 729 | ||
| 730 | lemma fst_mono: "A \<subseteq> B \<Longrightarrow> fst ` A \<subseteq> fst ` B" by fast | |
| 731 | ||
| 732 | lemma wf_syscls: | |
| 733 | "set SystemClasses \<subseteq> set G \<Longrightarrow> wf_syscls G" | |
| 734 | apply (drule fst_mono) | |
| 735 | apply (simp add: SystemClasses_def wf_syscls_def) | |
| 736 | apply (simp add: ObjectC_def) | |
| 737 | apply (rule allI, case_tac x) | |
| 738 | apply (auto simp add: NullPointerC_def ClassCastC_def OutOfMemoryC_def) | |
| 739 | done | |
| 740 | ||
| 8011 | 741 | end |