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