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