src/HOL/MicroJava/BV/JType.thy
author wenzelm
Sun, 20 Nov 2011 21:05:23 +0100
changeset 45605 a89b4bc311a5
parent 44890 22f665a2e91c
child 58263 6c907aad90ba
permissions -rw-r--r--
eliminated obsolete "standard";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12516
d09d0f160888 exceptions
kleing
parents: 12443
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/JType.thy
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10797
diff changeset
     2
    Author:     Tobias Nipkow, Gerwin Klein
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     3
    Copyright   2000 TUM
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     4
*)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     5
12911
704713ca07ea new document
kleing
parents: 12773
diff changeset
     6
header {* \isaheader{The Java Type System as Semilattice} *}
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     7
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
     8
theory JType
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
     9
imports "../DFA/Semilattices" "../J/WellForm"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
    10
begin
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    11
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    12
definition super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname" where
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
    13
  "super G C == fst (the (class G C))"
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
    14
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
    15
lemma superI:
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
    16
  "G \<turnstile> C \<prec>C1 D \<Longrightarrow> super G C = D"
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
    17
  by (unfold super_def) (auto dest: subcls1D)
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
    18
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    19
definition is_ref :: "ty \<Rightarrow> bool" where
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    20
  "is_ref T == case T of PrimT t \<Rightarrow> False | RefT r \<Rightarrow> True"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    21
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    22
definition sup :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty err" where
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    23
  "sup G T1 T2 ==
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    24
  case T1 of PrimT P1 \<Rightarrow> (case T2 of PrimT P2 \<Rightarrow> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    25
                         (if P1 = P2 then OK (PrimT P1) else Err) | RefT R \<Rightarrow> Err)
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    26
           | RefT R1 \<Rightarrow> (case T2 of PrimT P \<Rightarrow> Err | RefT R2 \<Rightarrow> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    27
  (case R1 of NullT \<Rightarrow> (case R2 of NullT \<Rightarrow> OK NT | ClassT C \<Rightarrow> OK (Class C))
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    28
            | ClassT C \<Rightarrow> (case R2 of NullT \<Rightarrow> OK (Class C) 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    29
                           | ClassT D \<Rightarrow> OK (Class (exec_lub (subcls1 G) (super G) C D)))))"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    30
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    31
definition subtype :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" where
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
    32
  "subtype G T1 T2 == G \<turnstile> T1 \<preceq> T2"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    33
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    34
definition is_ty :: "'c prog \<Rightarrow> ty \<Rightarrow> bool" where
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    35
  "is_ty G T == case T of PrimT P \<Rightarrow> True | RefT R \<Rightarrow>
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
    36
               (case R of NullT \<Rightarrow> True | ClassT C \<Rightarrow> (C, Object) \<in> (subcls1 G)^*)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    37
35102
cc7a0b9f938c modernized translations;
wenzelm
parents: 33954
diff changeset
    38
abbreviation "types G == Collect (is_type G)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    39
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    40
definition esl :: "'c prog \<Rightarrow> ty esl" where
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    41
  "esl G == (types G, subtype G, sup G)"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    42
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
    43
lemma PrimT_PrimT: "(G \<turnstile> xb \<preceq> PrimT p) = (xb = PrimT p)"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
    44
  by (auto elim: widen.cases)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    45
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
    46
lemma PrimT_PrimT2: "(G \<turnstile> PrimT p \<preceq> xb) = (xb = PrimT p)"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
    47
  by (auto elim: widen.cases)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    48
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    49
lemma is_tyI:
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
    50
  "\<lbrakk> is_type G T; ws_prog G \<rbrakk> \<Longrightarrow> is_ty G T"
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    51
  by (auto simp add: is_ty_def intro: subcls_C_Object 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    52
           split: ty.splits ref_ty.splits)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    53
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    54
lemma is_type_conv: 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
    55
  "ws_prog G \<Longrightarrow> is_type G T = is_ty G T"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    56
proof
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
    57
  assume "is_type G T" "ws_prog G" 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    58
  thus "is_ty G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    59
    by (rule is_tyI)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    60
next
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
    61
  assume wf: "ws_prog G" and
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    62
         ty: "is_ty G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    63
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    64
  show "is_type G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    65
  proof (cases T)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    66
    case PrimT
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    67
    thus ?thesis by simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    68
  next
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    69
    fix R assume R: "T = RefT R"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    70
    with wf
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
    71
    have "R = ClassT Object \<Longrightarrow> ?thesis" by simp
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    72
    moreover    
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    73
    from R wf ty
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
    74
    have "R \<noteq> ClassT Object \<Longrightarrow> ?thesis"
12443
e56ab6134b41 Turned subcls1 into an inductive relation to make it executable.
berghofe
parents: 11085
diff changeset
    75
     by (auto simp add: is_ty_def is_class_def split_tupled_all
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
    76
               elim!: subcls1.cases
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
    77
               elim: converse_rtranclE
12443
e56ab6134b41 Turned subcls1 into an inductive relation to make it executable.
berghofe
parents: 11085
diff changeset
    78
               split: ref_ty.splits)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    79
    ultimately    
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    80
    show ?thesis by blast
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    81
  qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    82
qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    83
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    84
lemma order_widen:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
    85
  "acyclic (subcls1 G) \<Longrightarrow> order (subtype G)"
22068
00bed5ac9884 renamed locale partial_order to order
haftmann
parents: 16417
diff changeset
    86
  apply (unfold Semilat.order_def lesub_def subtype_def)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    87
  apply (auto intro: widen_trans)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    88
  apply (case_tac x)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    89
   apply (case_tac y)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    90
    apply (auto simp add: PrimT_PrimT)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    91
   apply (case_tac y)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    92
    apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    93
  apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    94
  apply (case_tac ref_ty)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    95
   apply (case_tac ref_tya)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    96
    apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    97
   apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    98
  apply (case_tac ref_tya)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    99
   apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   100
  apply simp
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   101
  apply (auto dest: acyclic_impl_antisym_rtrancl antisymD)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   102
  done
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   103
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   104
lemma wf_converse_subcls1_impl_acc_subtype:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   105
  "wf ((subcls1 G)^-1) \<Longrightarrow> acc (subtype G)"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
   106
apply (unfold Semilat.acc_def lesssub_def)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   107
apply (drule_tac p = "((subcls1 G)^-1) - Id" in wf_subset)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
   108
 apply auto
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   109
apply (drule wf_trancl)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   110
apply (simp add: wf_eq_minimal)
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   111
apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   112
apply (unfold lesub_def subtype_def)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   113
apply (rename_tac M T) 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   114
apply (case_tac "EX C. Class C : M")
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   115
 prefer 2
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   116
 apply (case_tac T)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 35416
diff changeset
   117
  apply (fastforce simp add: PrimT_PrimT2)
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   118
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   119
 apply (subgoal_tac "ref_ty = NullT")
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   120
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   121
  apply (rule_tac x = NT in bexI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   122
   apply (rule allI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   123
   apply (rule impI, erule conjE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   124
   apply (drule widen_RefT)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   125
   apply clarsimp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   126
   apply (case_tac t)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   127
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   128
   apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   129
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   130
 apply (case_tac ref_ty)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   131
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   132
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   133
apply (erule_tac x = "{C. Class C : M}" in allE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   134
apply auto
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   135
apply (rename_tac D)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   136
apply (rule_tac x = "Class D" in bexI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   137
 prefer 2
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   138
 apply assumption
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   139
apply clarify 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   140
apply (frule widen_RefT)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   141
apply (erule exE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   142
apply (case_tac t)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   143
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   144
apply simp
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   145
apply (insert rtrancl_r_diff_Id [symmetric, of "subcls1 G"])
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   146
apply simp
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   147
apply (erule rtrancl.cases)
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   148
 apply blast
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   149
apply (drule rtrancl_converseI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   150
apply (subgoal_tac "(subcls1 G - Id)^-1 = (subcls1 G)^-1 - Id")
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   151
 prefer 2
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   152
 apply (simp add: converse_Int) apply safe[1]
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
   153
apply simp
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   154
apply (blast intro: rtrancl_into_trancl2)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   155
done
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   156
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   157
lemma closed_err_types:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   158
  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \<rbrakk> 
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   159
  \<Longrightarrow> closed (err (types G)) (lift2 (sup G))"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   160
  apply (unfold closed_def plussub_def lift2_def sup_def)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   161
  apply (auto split: err.split)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   162
  apply (drule is_tyI, assumption)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   163
  apply (auto simp add: is_ty_def is_type_conv simp del: is_type.simps 
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   164
              split: ty.split ref_ty.split)  
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   165
  apply (blast dest!: is_lub_exec_lub is_lubD is_ubD intro!: is_ubI superI)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   166
  done
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   167
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   168
10896
kleing
parents: 10812
diff changeset
   169
lemma sup_subtype_greater:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   170
  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G);
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   171
      is_type G t1; is_type G t2; sup G t1 t2 = OK s \<rbrakk> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   172
  \<Longrightarrow> subtype G t1 s \<and> subtype G t2 s"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   173
proof -
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   174
  assume ws_prog:       "ws_prog G"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   175
  assume single_valued: "single_valued (subcls1 G)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   176
  assume acyclic:       "acyclic (subcls1 G)"
10896
kleing
parents: 10812
diff changeset
   177
 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   178
  { fix c1 c2
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   179
    assume is_class: "is_class G c1" "is_class G c2"
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   180
    with ws_prog 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   181
    obtain 
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   182
      "G \<turnstile> c1 \<preceq>C Object"
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   183
      "G \<turnstile> c2 \<preceq>C Object"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   184
      by (blast intro: subcls_C_Object)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   185
    with ws_prog single_valued
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   186
    obtain u where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   187
      "is_lub ((subcls1 G)^* ) c1 c2 u"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   188
      by (blast dest: single_valued_has_lubs)
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   189
    moreover
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   190
    note acyclic
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   191
    moreover
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
   192
    have "\<forall>x y. G \<turnstile> x \<prec>C1 y \<longrightarrow> super G x = y"
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   193
      by (blast intro: superI)
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   194
    ultimately
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   195
    have "G \<turnstile> c1 \<preceq>C exec_lub (subcls1 G) (super G) c1 c2 \<and>
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   196
          G \<turnstile> c2 \<preceq>C exec_lub (subcls1 G) (super G) c1 c2"
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   197
      by (simp add: exec_lub_conv) (blast dest: is_lubD is_ubD)
10896
kleing
parents: 10812
diff changeset
   198
  } note this [simp]
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   199
      
10896
kleing
parents: 10812
diff changeset
   200
  assume "is_type G t1" "is_type G t2" "sup G t1 t2 = OK s"
kleing
parents: 10812
diff changeset
   201
  thus ?thesis
kleing
parents: 10812
diff changeset
   202
    apply (unfold sup_def subtype_def) 
kleing
parents: 10812
diff changeset
   203
    apply (cases s)
kleing
parents: 10812
diff changeset
   204
    apply (auto split: ty.split_asm ref_ty.split_asm split_if_asm)
kleing
parents: 10812
diff changeset
   205
    done
kleing
parents: 10812
diff changeset
   206
qed
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   207
10896
kleing
parents: 10812
diff changeset
   208
lemma sup_subtype_smallest:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   209
  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G);
10896
kleing
parents: 10812
diff changeset
   210
      is_type G a; is_type G b; is_type G c; 
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   211
      subtype G a c; subtype G b c; sup G a b = OK d \<rbrakk>
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   212
  \<Longrightarrow> subtype G d c"
10896
kleing
parents: 10812
diff changeset
   213
proof -
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   214
  assume ws_prog:       "ws_prog G"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   215
  assume single_valued: "single_valued (subcls1 G)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   216
  assume acyclic:       "acyclic (subcls1 G)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   217
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   218
  { fix c1 c2 D
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   219
    assume is_class: "is_class G c1" "is_class G c2"
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   220
    assume le: "G \<turnstile> c1 \<preceq>C D" "G \<turnstile> c2 \<preceq>C D"
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   221
    from ws_prog is_class
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   222
    obtain 
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   223
      "G \<turnstile> c1 \<preceq>C Object"
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   224
      "G \<turnstile> c2 \<preceq>C Object"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   225
      by (blast intro: subcls_C_Object)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   226
    with ws_prog single_valued
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   227
    obtain u where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   228
      lub: "is_lub ((subcls1 G)^*) c1 c2 u"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   229
      by (blast dest: single_valued_has_lubs)   
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   230
    with acyclic
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   231
    have "exec_lub (subcls1 G) (super G) c1 c2 = u"
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   232
      by (blast intro: superI exec_lub_conv)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   233
    moreover
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   234
    from lub le
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   235
    have "G \<turnstile> u \<preceq>C D" 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   236
      by (simp add: is_lub_def is_ub_def)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   237
    ultimately     
12773
a47f51daa6dc use exec_lub instead of some_lub
kleing
parents: 12516
diff changeset
   238
    have "G \<turnstile> exec_lub (subcls1 G) (super G) c1 c2 \<preceq>C D"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   239
      by blast
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   240
  } note this [intro]
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   241
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   242
  have [dest!]:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   243
    "\<And>C T. G \<turnstile> Class C \<preceq> T \<Longrightarrow> \<exists>D. T=Class D \<and> G \<turnstile> C \<preceq>C D"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   244
    by (frule widen_Class, auto)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   245
10896
kleing
parents: 10812
diff changeset
   246
  assume "is_type G a" "is_type G b" "is_type G c"
kleing
parents: 10812
diff changeset
   247
         "subtype G a c" "subtype G b c" "sup G a b = OK d"
kleing
parents: 10812
diff changeset
   248
  thus ?thesis
kleing
parents: 10812
diff changeset
   249
    by (auto simp add: subtype_def sup_def 
kleing
parents: 10812
diff changeset
   250
             split: ty.split_asm ref_ty.split_asm split_if_asm)
kleing
parents: 10812
diff changeset
   251
qed
kleing
parents: 10812
diff changeset
   252
kleing
parents: 10812
diff changeset
   253
lemma sup_exists:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   254
  "\<lbrakk> subtype G a c; subtype G b c; sup G a b = Err \<rbrakk> \<Longrightarrow> False"
10896
kleing
parents: 10812
diff changeset
   255
  by (auto simp add: PrimT_PrimT PrimT_PrimT2 sup_def subtype_def
kleing
parents: 10812
diff changeset
   256
           split: ty.splits ref_ty.splits)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   257
10896
kleing
parents: 10812
diff changeset
   258
lemma err_semilat_JType_esl_lemma:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   259
  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \<rbrakk> 
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   260
  \<Longrightarrow> err_semilat (esl G)"
10896
kleing
parents: 10812
diff changeset
   261
proof -
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   262
  assume ws_prog:   "ws_prog G"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   263
  assume single_valued: "single_valued (subcls1 G)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   264
  assume acyclic:   "acyclic (subcls1 G)"
10896
kleing
parents: 10812
diff changeset
   265
  
kleing
parents: 10812
diff changeset
   266
  hence "order (subtype G)"
kleing
parents: 10812
diff changeset
   267
    by (rule order_widen)
kleing
parents: 10812
diff changeset
   268
  moreover
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   269
  from ws_prog single_valued acyclic
10896
kleing
parents: 10812
diff changeset
   270
  have "closed (err (types G)) (lift2 (sup G))"
kleing
parents: 10812
diff changeset
   271
    by (rule closed_err_types)
kleing
parents: 10812
diff changeset
   272
  moreover
kleing
parents: 10812
diff changeset
   273
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   274
  from ws_prog single_valued acyclic
10896
kleing
parents: 10812
diff changeset
   275
  have
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   276
    "(\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). x <=_(Err.le (subtype G)) x +_(lift2 (sup G)) y) \<and> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   277
     (\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). y <=_(Err.le (subtype G)) x +_(lift2 (sup G)) y)"
13074
96bf406fd3e5 Started to convert to locales
nipkow
parents: 13006
diff changeset
   278
    by (auto simp add: lesub_def plussub_def Err.le_def lift2_def sup_subtype_greater split: err.split)
10896
kleing
parents: 10812
diff changeset
   279
13074
96bf406fd3e5 Started to convert to locales
nipkow
parents: 13006
diff changeset
   280
  moreover
10896
kleing
parents: 10812
diff changeset
   281
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   282
  from ws_prog single_valued acyclic 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   283
  have
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10925
diff changeset
   284
    "\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). \<forall>z\<in>err (types G). 
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   285
    x <=_(Err.le (subtype G)) z \<and> y <=_(Err.le (subtype G)) z \<longrightarrow> x +_(lift2 (sup G)) y <=_(Err.le (subtype G)) z"
13074
96bf406fd3e5 Started to convert to locales
nipkow
parents: 13006
diff changeset
   286
    by (unfold lift2_def plussub_def lesub_def Err.le_def)
10896
kleing
parents: 10812
diff changeset
   287
       (auto intro: sup_subtype_smallest sup_exists split: err.split)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   288
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   289
  ultimately
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   290
  
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   291
  show ?thesis
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   292
    by (unfold esl_def semilat_def Err.sl_def) auto
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   293
qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   294
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   295
lemma single_valued_subcls1:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 25363
diff changeset
   296
  "ws_prog G \<Longrightarrow> single_valued (subcls1 G)"
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   297
  by (auto simp add: ws_prog_def unique_def single_valued_def
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 22068
diff changeset
   298
    intro: subcls1I elim!: subcls1.cases)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   299
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   300
theorem err_semilat_JType_esl:
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13074
diff changeset
   301
  "ws_prog G \<Longrightarrow> err_semilat (esl G)"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   302
  by (frule acyclic_subcls1, frule single_valued_subcls1, rule err_semilat_JType_esl_lemma)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   303
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   304
end