author | kleing |
Sun, 16 Dec 2001 00:18:17 +0100 | |
changeset 12517 | 360e3215f029 |
parent 12443 | e56ab6134b41 |
child 12911 | 704713ca07ea |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/J/TypeRel.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
11070 | 5 |
*) |
8011 | 6 |
|
11070 | 7 |
header "Relations between Java Types" |
8011 | 8 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
9 |
theory TypeRel = Decl: |
8011 | 10 |
|
11 |
consts |
|
12517 | 12 |
subcls1 :: "'c prog => (cname \<times> cname) set" -- "subclass" |
13 |
widen :: "'c prog => (ty \<times> ty ) set" -- "widening" |
|
14 |
cast :: "'c prog => (cname \<times> cname) set" -- "casting" |
|
8011 | 15 |
|
11372 | 16 |
syntax (xsymbols) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
17 |
subcls1 :: "'c prog => [cname, cname] => bool" ("_ \<turnstile> _ \<prec>C1 _" [71,71,71] 70) |
11372 | 18 |
subcls :: "'c prog => [cname, cname] => bool" ("_ \<turnstile> _ \<preceq>C _" [71,71,71] 70) |
19 |
widen :: "'c prog => [ty , ty ] => bool" ("_ \<turnstile> _ \<preceq> _" [71,71,71] 70) |
|
20 |
cast :: "'c prog => [cname, cname] => bool" ("_ \<turnstile> _ \<preceq>? _" [71,71,71] 70) |
|
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
21 |
|
11372 | 22 |
syntax |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
23 |
subcls1 :: "'c prog => [cname, cname] => bool" ("_ |- _ <=C1 _" [71,71,71] 70) |
11372 | 24 |
subcls :: "'c prog => [cname, cname] => bool" ("_ |- _ <=C _" [71,71,71] 70) |
25 |
widen :: "'c prog => [ty , ty ] => bool" ("_ |- _ <= _" [71,71,71] 70) |
|
26 |
cast :: "'c prog => [cname, cname] => bool" ("_ |- _ <=? _" [71,71,71] 70) |
|
8011 | 27 |
|
28 |
translations |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
29 |
"G\<turnstile>C \<prec>C1 D" == "(C,D) \<in> subcls1 G" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
30 |
"G\<turnstile>C \<preceq>C D" == "(C,D) \<in> (subcls1 G)^*" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
31 |
"G\<turnstile>S \<preceq> T" == "(S,T) \<in> widen G" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
32 |
"G\<turnstile>C \<preceq>? D" == "(C,D) \<in> cast G" |
8011 | 33 |
|
12517 | 34 |
-- "direct subclass, cf. 8.1.3" |
12443
e56ab6134b41
Turned subcls1 into an inductive relation to make it executable.
berghofe
parents:
11987
diff
changeset
|
35 |
inductive "subcls1 G" intros |
e56ab6134b41
Turned subcls1 into an inductive relation to make it executable.
berghofe
parents:
11987
diff
changeset
|
36 |
subcls1I: "\<lbrakk>class G C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>C1D" |
8011 | 37 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
38 |
lemma subcls1D: |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
39 |
"G\<turnstile>C\<prec>C1D \<Longrightarrow> C \<noteq> Object \<and> (\<exists>fs ms. class G C = Some (D,fs,ms))" |
12443
e56ab6134b41
Turned subcls1 into an inductive relation to make it executable.
berghofe
parents:
11987
diff
changeset
|
40 |
apply (erule subcls1.elims) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
41 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
42 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
43 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
44 |
lemma subcls1_def2: |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
45 |
"subcls1 G = (\<Sigma>C\<in>{C. is_class G C} . {D. C\<noteq>Object \<and> fst (the (class G C))=D})" |
12443
e56ab6134b41
Turned subcls1 into an inductive relation to make it executable.
berghofe
parents:
11987
diff
changeset
|
46 |
by (auto simp add: is_class_def dest: subcls1D intro: subcls1I) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
47 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
48 |
lemma finite_subcls1: "finite (subcls1 G)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
49 |
apply(subst subcls1_def2) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
50 |
apply(rule finite_SigmaI [OF finite_is_class]) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
51 |
apply(rule_tac B = "{fst (the (class G C))}" in finite_subset) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
52 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
53 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
54 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
55 |
lemma subcls_is_class: "(C,D) \<in> (subcls1 G)^+ ==> is_class G C" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
56 |
apply (unfold is_class_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
57 |
apply(erule trancl_trans_induct) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
58 |
apply (auto dest!: subcls1D) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
59 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
60 |
|
11266 | 61 |
lemma subcls_is_class2 [rule_format (no_asm)]: |
62 |
"G\<turnstile>C\<preceq>C D \<Longrightarrow> is_class G D \<longrightarrow> is_class G C" |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
63 |
apply (unfold is_class_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
64 |
apply (erule rtrancl_induct) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
65 |
apply (drule_tac [2] subcls1D) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
66 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
67 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
68 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
69 |
consts class_rec ::"'c prog \<times> cname \<Rightarrow> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
70 |
'a \<Rightarrow> (cname \<Rightarrow> fdecl list \<Rightarrow> 'c mdecl list \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a" |
11266 | 71 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
72 |
recdef class_rec "same_fst (\<lambda>G. wf ((subcls1 G)^-1)) (\<lambda>G. (subcls1 G)^-1)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
73 |
"class_rec (G,C) = (\<lambda>t f. case class G C of None \<Rightarrow> arbitrary |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
74 |
| Some (D,fs,ms) \<Rightarrow> if wf ((subcls1 G)^-1) then |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
75 |
f C fs ms (if C = Object then t else class_rec (G,D) t f) else arbitrary)" |
11266 | 76 |
(hints intro: subcls1I) |
11284 | 77 |
|
11266 | 78 |
declare class_rec.simps [simp del] |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
79 |
|
11284 | 80 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
81 |
lemma class_rec_lemma: "\<lbrakk> wf ((subcls1 G)^-1); class G C = Some (D,fs,ms)\<rbrakk> \<Longrightarrow> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
82 |
class_rec (G,C) t f = f C fs ms (if C=Object then t else class_rec (G,D) t f)"; |
11266 | 83 |
apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]]) |
84 |
apply simp |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
85 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
86 |
|
8011 | 87 |
consts |
88 |
||
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
89 |
method :: "'c prog \<times> cname => ( sig \<leadsto> cname \<times> ty \<times> 'c)" (* ###curry *) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
90 |
field :: "'c prog \<times> cname => ( vname \<leadsto> cname \<times> ty )" (* ###curry *) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
91 |
fields :: "'c prog \<times> cname => ((vname \<times> cname) \<times> ty) list" (* ###curry *) |
8011 | 92 |
|
12517 | 93 |
-- "methods of a class, with inheritance, overriding and hiding, cf. 8.4.6" |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
94 |
defs method_def: "method \<equiv> \<lambda>(G,C). class_rec (G,C) empty (\<lambda>C fs ms ts. |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
95 |
ts ++ map_of (map (\<lambda>(s,m). (s,(C,m))) ms))" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
96 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
97 |
lemma method_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
98 |
method (G,C) = (if C = Object then empty else method (G,D)) ++ |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
99 |
map_of (map (\<lambda>(s,m). (s,(C,m))) ms)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
100 |
apply (unfold method_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
101 |
apply (simp split del: split_if) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
102 |
apply (erule (1) class_rec_lemma [THEN trans]); |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
103 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
104 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
105 |
|
8011 | 106 |
|
12517 | 107 |
-- "list of fields of a class, including inherited and hidden ones" |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
108 |
defs fields_def: "fields \<equiv> \<lambda>(G,C). class_rec (G,C) [] (\<lambda>C fs ms ts. |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
109 |
map (\<lambda>(fn,ft). ((fn,C),ft)) fs @ ts)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
110 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
111 |
lemma fields_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
112 |
fields (G,C) = |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
113 |
map (\<lambda>(fn,ft). ((fn,C),ft)) fs @ (if C = Object then [] else fields (G,D))" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
114 |
apply (unfold fields_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
115 |
apply (simp split del: split_if) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
116 |
apply (erule (1) class_rec_lemma [THEN trans]); |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
117 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
118 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
119 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
120 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
121 |
defs field_def: "field == map_of o (map (\<lambda>((fn,fd),ft). (fn,(fd,ft)))) o fields" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
122 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
123 |
lemma field_fields: |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
124 |
"field (G,C) fn = Some (fd, fT) \<Longrightarrow> map_of (fields (G,C)) (fn, fd) = Some fT" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
125 |
apply (unfold field_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
126 |
apply (rule table_of_remap_SomeD) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
127 |
apply simp |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
128 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
129 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
130 |
|
12517 | 131 |
-- "widening, viz. method invocation conversion,cf. 5.3 i.e. sort of syntactic subtyping" |
132 |
inductive "widen G" intros |
|
133 |
refl [intro!, simp]: "G\<turnstile> T \<preceq> T" -- "identity conv., cf. 5.1.1" |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
134 |
subcls : "G\<turnstile>C\<preceq>C D ==> G\<turnstile>Class C \<preceq> Class D" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
135 |
null [intro!]: "G\<turnstile> NT \<preceq> RefT R" |
8011 | 136 |
|
12517 | 137 |
-- "casting conversion, cf. 5.5 / 5.1.5" |
138 |
-- "left out casts on primitve types" |
|
139 |
inductive "cast G" intros |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
140 |
widen: "G\<turnstile>C\<preceq>C D ==> G\<turnstile>C \<preceq>? D" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
141 |
subcls: "G\<turnstile>D\<preceq>C C ==> G\<turnstile>C \<preceq>? D" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
142 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
143 |
lemma widen_PrimT_RefT [iff]: "(G\<turnstile>PrimT pT\<preceq>RefT rT) = False" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
144 |
apply (rule iffI) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
145 |
apply (erule widen.elims) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
146 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
147 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
148 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
149 |
lemma widen_RefT: "G\<turnstile>RefT R\<preceq>T ==> \<exists>t. T=RefT t" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
150 |
apply (ind_cases "G\<turnstile>S\<preceq>T") |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
151 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
152 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
153 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
154 |
lemma widen_RefT2: "G\<turnstile>S\<preceq>RefT R ==> \<exists>t. S=RefT t" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
155 |
apply (ind_cases "G\<turnstile>S\<preceq>T") |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
156 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
157 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
158 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
159 |
lemma widen_Class: "G\<turnstile>Class C\<preceq>T ==> \<exists>D. T=Class D" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
160 |
apply (ind_cases "G\<turnstile>S\<preceq>T") |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
161 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
162 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
163 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
164 |
lemma widen_Class_NullT [iff]: "(G\<turnstile>Class C\<preceq>NT) = False" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
165 |
apply (rule iffI) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
166 |
apply (ind_cases "G\<turnstile>S\<preceq>T") |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
167 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
168 |
done |
8011 | 169 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
170 |
lemma widen_Class_Class [iff]: "(G\<turnstile>Class C\<preceq> Class D) = (G\<turnstile>C\<preceq>C D)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
171 |
apply (rule iffI) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
172 |
apply (ind_cases "G\<turnstile>S\<preceq>T") |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
173 |
apply (auto elim: widen.subcls) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
174 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
175 |
|
12517 | 176 |
theorem widen_trans[trans]: "\<lbrakk>G\<turnstile>S\<preceq>U; G\<turnstile>U\<preceq>T\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T" |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
177 |
proof - |
12517 | 178 |
assume "G\<turnstile>S\<preceq>U" thus "\<And>T. G\<turnstile>U\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>T" |
11987 | 179 |
proof induct |
12517 | 180 |
case (refl T T') thus "G\<turnstile>T\<preceq>T'" . |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
181 |
next |
11987 | 182 |
case (subcls C D T) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
183 |
then obtain E where "T = Class E" by (blast dest: widen_Class) |
11987 | 184 |
with subcls show "G\<turnstile>Class C\<preceq>T" by (auto elim: rtrancl_trans) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
185 |
next |
11987 | 186 |
case (null R RT) |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
187 |
then obtain rt where "RT = RefT rt" by (blast dest: widen_RefT) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
188 |
thus "G\<turnstile>NT\<preceq>RT" by auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
189 |
qed |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
190 |
qed |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10613
diff
changeset
|
191 |
|
8011 | 192 |
end |