src/HOL/MicroJava/BV/JType.thy
author oheimb
Thu, 18 Jan 2001 17:38:56 +0100
changeset 10925 5ffe7ed8899a
parent 10896 23386a5b63eb
child 11085 b830bf10bf71
permissions -rw-r--r--
is_class and class now as defs (rather than translations); corrected Digest.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     1
(*  Title:      HOL/BCV/JType.thy
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
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10797
diff changeset
     7
header "Java Type System as Semilattice"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     8
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
     9
theory JType = WellForm + Err:
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
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    12
  is_ref :: "ty => bool"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    13
  "is_ref T == case T of PrimT t => False | RefT r => True"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    14
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    15
  sup :: "'c prog => ty => ty => ty err"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    16
  "sup G T1 T2 ==
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    17
  case T1 of PrimT P1 => (case T2 of PrimT P2 => (if P1 = P2 then OK (PrimT P1) else Err) | RefT R => Err)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    18
           | RefT R1 => (case T2 of PrimT P => Err | RefT R2 => 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    19
  (case R1 of NullT => (case R2 of NullT => OK NT | ClassT C => OK (Class C))
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    20
            | ClassT C => (case R2 of NullT => OK (Class C) 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    21
                                    | ClassT D => OK (Class (some_lub ((subcls1 G)^* ) C D)))))"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    22
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    23
  subtype :: "'c prog => ty => ty => bool"
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    24
  "subtype G T1 T2 == G \\<turnstile> T1 \\<preceq> T2"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    25
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    26
  is_ty :: "'c prog => ty => bool"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    27
  "is_ty G T == case T of PrimT P => True | RefT R =>
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    28
               (case R of NullT => True | ClassT C => (C,Object):(subcls1 G)^*)"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    29
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    30
translations
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    31
  "types G" == "Collect (is_type G)"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    32
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    33
constdefs
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    34
  esl :: "'c prog => ty esl"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    35
  "esl G == (types G, subtype G, sup G)"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    36
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    37
lemma PrimT_PrimT: "(G \\<turnstile> xb \\<preceq> PrimT p) = (xb = PrimT p)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    38
  by (auto elim: widen.elims)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    39
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    40
lemma PrimT_PrimT2: "(G \\<turnstile> PrimT p \\<preceq> xb) = (xb = PrimT p)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    41
  by (auto elim: widen.elims)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    42
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    43
lemma is_tyI:
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    44
  "[| is_type G T; wf_prog wf_mb G |] ==> is_ty G T"
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    45
  by (auto simp add: is_ty_def intro: subcls_C_Object 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    46
           split: ty.splits ref_ty.splits)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    47
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    48
lemma is_type_conv: 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    49
  "wf_prog wf_mb G ==> is_type G T = is_ty G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    50
proof
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    51
  assume "is_type G T" "wf_prog wf_mb G" 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    52
  thus "is_ty G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    53
    by (rule is_tyI)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    54
next
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    55
  assume wf: "wf_prog wf_mb G" and
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    56
         ty: "is_ty G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    57
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    58
  show "is_type G T"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    59
  proof (cases T)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    60
    case PrimT
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    61
    thus ?thesis by simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    62
  next
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    63
    fix R assume R: "T = RefT R"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    64
    with wf
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    65
    have "R = ClassT Object \\<Longrightarrow> ?thesis" by simp
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    66
    moreover    
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    67
    from R wf ty
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    68
    have "R \\<noteq> ClassT Object \\<Longrightarrow> ?thesis"
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
    69
      by (auto simp add: is_ty_def subcls1_def is_class_def
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    70
               elim: converse_rtranclE
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    71
               split: ref_ty.splits)    
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    72
    ultimately    
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    73
    show ?thesis by blast
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    74
  qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    75
qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    76
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    77
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    78
lemma order_widen:
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    79
  "acyclic (subcls1 G) ==> order (subtype G)"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    80
  apply (unfold order_def lesub_def subtype_def)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    81
  apply (auto intro: widen_trans)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    82
  apply (case_tac x)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    83
   apply (case_tac y)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    84
    apply (auto simp add: PrimT_PrimT)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    85
   apply (case_tac y)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    86
    apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    87
  apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    88
  apply (case_tac ref_ty)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    89
   apply (case_tac ref_tya)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    90
    apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    91
   apply simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    92
  apply (case_tac ref_tya)
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 simp
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    95
  apply (auto dest: acyclic_impl_antisym_rtrancl antisymD)  
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    96
  done
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
    97
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
    98
lemma wf_converse_subcls1_impl_acc_subtype:
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
    99
  "wf ((subcls1 G)^-1) ==> acc (subtype G)"
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   100
apply (unfold acc_def lesssub_def)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   101
apply (drule_tac p = "(subcls1 G)^-1 - Id" in wf_subset)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   102
 apply blast
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   103
apply (drule wf_trancl)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   104
apply (simp add: wf_eq_minimal)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   105
apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   106
apply (unfold lesub_def subtype_def)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   107
apply (rename_tac M T) 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   108
apply (case_tac "EX C. Class C : M")
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   109
 prefer 2
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   110
 apply (case_tac T)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   111
  apply (fastsimp simp add: PrimT_PrimT2)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   112
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   113
 apply (subgoal_tac "ref_ty = NullT")
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   114
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   115
  apply (rule_tac x = NT in bexI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   116
   apply (rule allI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   117
   apply (rule impI, erule conjE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   118
   apply (drule widen_RefT)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   119
   apply clarsimp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   120
   apply (case_tac t)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   121
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   122
   apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   123
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   124
 apply (case_tac ref_ty)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   125
  apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   126
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   127
apply (erule_tac x = "{C. Class C : M}" in allE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   128
apply auto
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   129
apply (rename_tac D)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   130
apply (rule_tac x = "Class D" in bexI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   131
 prefer 2
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   132
 apply assumption
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   133
apply clarify 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   134
apply (frule widen_RefT)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   135
apply (erule exE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   136
apply (case_tac t)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   137
 apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   138
apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   139
apply (insert rtrancl_r_diff_Id [symmetric, standard, of "(subcls1 G)"])
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   140
apply simp
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   141
apply (erule rtranclE)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   142
 apply blast
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   143
apply (drule rtrancl_converseI)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   144
apply (subgoal_tac "((subcls1 G)-Id)^-1 = ((subcls1 G)^-1 - Id)")
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   145
 prefer 2
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   146
 apply blast
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   147
apply simp 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   148
apply (blast intro: rtrancl_into_trancl2)
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   149
done 
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   150
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   151
lemma closed_err_types:
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   152
  "[| wf_prog wf_mb G; single_valued (subcls1 G); acyclic (subcls1 G) |] 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   153
  ==> closed (err (types G)) (lift2 (sup G))"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   154
  apply (unfold closed_def plussub_def lift2_def sup_def)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   155
  apply (auto split: err.split)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   156
  apply (drule is_tyI, assumption)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   157
  apply (auto simp add: is_ty_def is_type_conv simp del: is_type.simps 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   158
              split: ty.split ref_ty.split)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   159
  apply (blast dest!: is_lub_some_lub is_lubD is_ubD intro!: is_ubI)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   160
  done
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   161
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   162
10896
kleing
parents: 10812
diff changeset
   163
lemma sup_subtype_greater:
kleing
parents: 10812
diff changeset
   164
  "[| wf_prog wf_mb G; single_valued (subcls1 G); acyclic (subcls1 G);
kleing
parents: 10812
diff changeset
   165
      is_type G t1; is_type G t2; sup G t1 t2 = OK s |] 
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   166
  ==> subtype G t1 s \\<and> subtype G t2 s"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   167
proof -
10896
kleing
parents: 10812
diff changeset
   168
  assume wf_prog:       "wf_prog wf_mb G"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   169
  assume single_valued: "single_valued (subcls1 G)" 
10896
kleing
parents: 10812
diff changeset
   170
  assume acyclic:       "acyclic (subcls1 G)"
kleing
parents: 10812
diff changeset
   171
 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   172
  { fix c1 c2
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   173
    assume is_class: "is_class G c1" "is_class G c2"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   174
    with wf_prog 
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   175
    obtain 
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   176
      "G \\<turnstile> c1 \\<preceq>C Object"
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   177
      "G \\<turnstile> c2 \\<preceq>C Object"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   178
      by (blast intro: subcls_C_Object)
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   179
    with wf_prog single_valued
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   180
    obtain u where
10896
kleing
parents: 10812
diff changeset
   181
      "is_lub ((subcls1 G)^* ) c1 c2 u"      
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   182
      by (blast dest: single_valued_has_lubs)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   183
    with acyclic
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   184
    have "G \\<turnstile> c1 \\<preceq>C some_lub ((subcls1 G)^* ) c1 c2 \\<and>
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   185
          G \\<turnstile> c2 \\<preceq>C some_lub ((subcls1 G)^* ) c1 c2"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   186
      by (simp add: some_lub_conv) (blast dest: is_lubD is_ubD)
10896
kleing
parents: 10812
diff changeset
   187
  } note this [simp]
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   188
      
10896
kleing
parents: 10812
diff changeset
   189
  assume "is_type G t1" "is_type G t2" "sup G t1 t2 = OK s"
kleing
parents: 10812
diff changeset
   190
  thus ?thesis
kleing
parents: 10812
diff changeset
   191
    apply (unfold sup_def subtype_def) 
kleing
parents: 10812
diff changeset
   192
    apply (cases s)
kleing
parents: 10812
diff changeset
   193
    apply (auto split: ty.split_asm ref_ty.split_asm split_if_asm)
kleing
parents: 10812
diff changeset
   194
    done
kleing
parents: 10812
diff changeset
   195
qed
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   196
10896
kleing
parents: 10812
diff changeset
   197
lemma sup_subtype_smallest:
kleing
parents: 10812
diff changeset
   198
  "[| wf_prog wf_mb G; single_valued (subcls1 G); acyclic (subcls1 G);
kleing
parents: 10812
diff changeset
   199
      is_type G a; is_type G b; is_type G c; 
kleing
parents: 10812
diff changeset
   200
      subtype G a c; subtype G b c; sup G a b = OK d |]
kleing
parents: 10812
diff changeset
   201
  ==> subtype G d c"
kleing
parents: 10812
diff changeset
   202
proof -
kleing
parents: 10812
diff changeset
   203
  assume wf_prog:       "wf_prog wf_mb G"
kleing
parents: 10812
diff changeset
   204
  assume single_valued: "single_valued (subcls1 G)" 
kleing
parents: 10812
diff changeset
   205
  assume acyclic:       "acyclic (subcls1 G)"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   206
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   207
  { fix c1 c2 D
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   208
    assume is_class: "is_class G c1" "is_class G c2"
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   209
    assume le: "G \\<turnstile> c1 \\<preceq>C D" "G \\<turnstile> c2 \\<preceq>C D"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   210
    from wf_prog is_class
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   211
    obtain 
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   212
      "G \\<turnstile> c1 \\<preceq>C Object"
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   213
      "G \\<turnstile> c2 \\<preceq>C Object"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   214
      by (blast intro: subcls_C_Object)
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   215
    with wf_prog single_valued
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   216
    obtain u where
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   217
      lub: "is_lub ((subcls1 G)^* ) c1 c2 u"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   218
      by (blast dest: single_valued_has_lubs)   
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   219
    with acyclic
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   220
    have "some_lub ((subcls1 G)^* ) c1 c2 = u"
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   221
      by (rule some_lub_conv)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   222
    moreover
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   223
    from lub le
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   224
    have "G \\<turnstile> u \\<preceq>C D" 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   225
      by (simp add: is_lub_def is_ub_def)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   226
    ultimately     
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   227
    have "G \\<turnstile> some_lub ((subcls1 G)\<^sup>*) c1 c2 \\<preceq>C D"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   228
      by blast
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   229
  } note this [intro]
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   230
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   231
  have [dest!]:
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   232
    "!!C T. G \\<turnstile> Class C \\<preceq> T ==> \\<exists>D. T=Class D \\<and> G \\<turnstile> C \\<preceq>C D"
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   233
    by (frule widen_Class, auto)
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   234
10896
kleing
parents: 10812
diff changeset
   235
  assume "is_type G a" "is_type G b" "is_type G c"
kleing
parents: 10812
diff changeset
   236
         "subtype G a c" "subtype G b c" "sup G a b = OK d"
kleing
parents: 10812
diff changeset
   237
  thus ?thesis
kleing
parents: 10812
diff changeset
   238
    by (auto simp add: subtype_def sup_def 
kleing
parents: 10812
diff changeset
   239
             split: ty.split_asm ref_ty.split_asm split_if_asm)
kleing
parents: 10812
diff changeset
   240
qed
kleing
parents: 10812
diff changeset
   241
kleing
parents: 10812
diff changeset
   242
lemma sup_exists:
kleing
parents: 10812
diff changeset
   243
  "[| subtype G a c; subtype G b c; sup G a b = Err |] ==> False"
kleing
parents: 10812
diff changeset
   244
  by (auto simp add: PrimT_PrimT PrimT_PrimT2 sup_def subtype_def
kleing
parents: 10812
diff changeset
   245
           split: ty.splits ref_ty.splits)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   246
10896
kleing
parents: 10812
diff changeset
   247
lemma err_semilat_JType_esl_lemma:
kleing
parents: 10812
diff changeset
   248
  "[| wf_prog wf_mb G; single_valued (subcls1 G); acyclic (subcls1 G) |] 
kleing
parents: 10812
diff changeset
   249
  ==> err_semilat (esl G)"
kleing
parents: 10812
diff changeset
   250
proof -
kleing
parents: 10812
diff changeset
   251
  assume wf_prog:   "wf_prog wf_mb G"
kleing
parents: 10812
diff changeset
   252
  assume single_valued: "single_valued (subcls1 G)" 
kleing
parents: 10812
diff changeset
   253
  assume acyclic:   "acyclic (subcls1 G)"
kleing
parents: 10812
diff changeset
   254
  
kleing
parents: 10812
diff changeset
   255
  hence "order (subtype G)"
kleing
parents: 10812
diff changeset
   256
    by (rule order_widen)
kleing
parents: 10812
diff changeset
   257
  moreover
kleing
parents: 10812
diff changeset
   258
  from wf_prog single_valued acyclic
kleing
parents: 10812
diff changeset
   259
  have "closed (err (types G)) (lift2 (sup G))"
kleing
parents: 10812
diff changeset
   260
    by (rule closed_err_types)
kleing
parents: 10812
diff changeset
   261
  moreover
kleing
parents: 10812
diff changeset
   262
kleing
parents: 10812
diff changeset
   263
  from wf_prog single_valued acyclic 
kleing
parents: 10812
diff changeset
   264
  have
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   265
    "(\\<forall>x\\<in>err (types G). \\<forall>y\\<in>err (types G). x <=_(le (subtype G)) x +_(lift2 (sup G)) y) \\<and> 
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   266
     (\\<forall>x\\<in>err (types G). \\<forall>y\\<in>err (types G). y <=_(le (subtype G)) x +_(lift2 (sup G)) y)"
10896
kleing
parents: 10812
diff changeset
   267
    by (auto simp add: lesub_def plussub_def le_def lift2_def sup_subtype_greater split: err.split)
kleing
parents: 10812
diff changeset
   268
kleing
parents: 10812
diff changeset
   269
  moreover    
kleing
parents: 10812
diff changeset
   270
kleing
parents: 10812
diff changeset
   271
  from wf_prog single_valued acyclic 
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   272
  have
10925
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   273
    "\\<forall>x\\<in>err (types G). \\<forall>y\\<in>err (types G). \\<forall>z\\<in>err (types G). 
5ffe7ed8899a is_class and class now as defs (rather than translations); corrected Digest.thy
oheimb
parents: 10896
diff changeset
   274
    x <=_(le (subtype G)) z \\<and> y <=_(le (subtype G)) z \\<longrightarrow> x +_(lift2 (sup G)) y <=_(le (subtype G)) z"
10896
kleing
parents: 10812
diff changeset
   275
    by (unfold lift2_def plussub_def lesub_def le_def)
kleing
parents: 10812
diff changeset
   276
       (auto intro: sup_subtype_smallest sup_exists split: err.split)
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   277
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   278
  ultimately
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   279
  
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   280
  show ?thesis
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   281
    by (unfold esl_def semilat_def sl_def) auto
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   282
qed
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   283
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   284
lemma single_valued_subcls1:
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   285
  "wf_prog wf_mb G ==> single_valued (subcls1 G)"
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   286
  by (unfold wf_prog_def unique_def single_valued_def subcls1_def) auto
10497
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   287
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   288
theorem err_semilat_JType_esl:
fc0b575a2cf7 BCV Integration
kleing
parents: 10497
diff changeset
   289
  "wf_prog wf_mb G ==> err_semilat (esl G)"
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10630
diff changeset
   290
  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
   291
7c6985b4de40 BCV integration (type system is semilattice)
kleing
parents:
diff changeset
   292
end