Theory TypeRel
section "Type relations"
theory TypeRel
imports Decl
begin
text‹Direct subclass relation›
definition subcls1 :: "(cname × cname) set"
where
"subcls1 ≡ {(C,D). C≠Object ∧ (∃c. class C = Some c ∧ super c=D)}"
abbreviation
subcls1_syntax :: "[cname, cname] => bool" ("_ ≺C1 _" [71,71] 70)
where "C ≺C1 D == (C,D) ∈ subcls1"
abbreviation
subcls_syntax :: "[cname, cname] => bool" ("_ ≼C _" [71,71] 70)
where "C ≼C D ≡ (C,D) ∈ subcls1⇧*"
subsection "Declarations and properties not used in the meta theory"
text‹Widening, viz. method invocation conversion›
inductive
widen :: "ty => ty => bool" ("_ ≼ _" [71,71] 70)
where
refl [intro!, simp]: "T ≼ T"
| subcls: "C≼C D ⟹ Class C ≼ Class D"
| null [intro!]: "NT ≼ R"
lemma subcls1D:
"C≺C1D ⟹ C ≠ Object ∧ (∃c. class C = Some c ∧ super c=D)"
apply (unfold subcls1_def)
apply auto
done
lemma subcls1I: "⟦class C = Some m; super m = D; C ≠ Object⟧ ⟹ C≺C1D"
apply (unfold subcls1_def)
apply auto
done
lemma subcls1_def2:
"subcls1 =
(SIGMA C: {C. is_class C} . {D. C≠Object ∧ super (the (class C)) = D})"
apply (unfold subcls1_def is_class_def)
apply (auto split:if_split_asm)
done
lemma finite_subcls1: "finite subcls1"
apply(subst subcls1_def2)
apply(rule finite_SigmaI [OF finite_is_class])
apply(rule_tac B = "{super (the (class C))}" in finite_subset)
apply auto
done
definition ws_prog :: "bool" where
"ws_prog ≡ ∀(C,c)∈set Prog. C≠Object ⟶
is_class (super c) ∧ (super c,C)∉subcls1⇧+"
lemma ws_progD: "⟦class C = Some c; C≠Object; ws_prog⟧ ⟹
is_class (super c) ∧ (super c,C)∉subcls1⇧+"
apply (unfold ws_prog_def class_def)
apply (drule_tac map_of_SomeD)
apply auto
done
lemma subcls1_irrefl_lemma1: "ws_prog ⟹ subcls1¯ ∩ subcls1⇧+ = {}"
by (fast dest: subcls1D ws_progD)
lemma irrefl_tranclI': "r¯ ∩ r⇧+ = {} ⟹ ∀x. (x, x) ∉ r⇧+"
by(blast elim: tranclE dest: trancl_into_rtrancl)
lemmas subcls1_irrefl_lemma2 = subcls1_irrefl_lemma1 [THEN irrefl_tranclI']
lemma subcls1_irrefl: "⟦(x, y) ∈ subcls1; ws_prog⟧ ⟹ x ≠ y"
apply (rule irrefl_trancl_rD)
apply (rule subcls1_irrefl_lemma2)
apply auto
done
lemmas subcls1_acyclic = subcls1_irrefl_lemma2 [THEN acyclicI]
lemma wf_subcls1: "ws_prog ⟹ wf (subcls1¯)"
by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic)
definition class_rec ::"cname ⇒ (class ⇒ ('a × 'b) list) ⇒ ('a ⇀ 'b)"
where
"class_rec ≡ wfrec (subcls1¯) (λrec C f.
case class C of None ⇒ undefined
| Some m ⇒ (if C = Object then Map.empty else rec (super m) f) ++ map_of (f m))"
lemma class_rec: "⟦class C = Some m; ws_prog⟧ ⟹
class_rec C f = (if C = Object then Map.empty else class_rec (super m) f) ++
map_of (f m)"
apply (drule wf_subcls1)
apply (subst def_wfrec[OF class_rec_def], auto)
apply (subst cut_apply, auto intro: subcls1I)
done
definition "method" :: "cname => (mname ⇀ methd)" where
"method C ≡ class_rec C methods"
lemma method_rec: "⟦class C = Some m; ws_prog⟧ ⟹
method C = (if C=Object then Map.empty else method (super m)) ++ map_of (methods m)"
apply (unfold method_def)
apply (erule (1) class_rec [THEN trans])
apply simp
done
definition field :: "cname => (fname ⇀ ty)" where
"field C ≡ class_rec C flds"
lemma flds_rec: "⟦class C = Some m; ws_prog⟧ ⟹
field C = (if C=Object then Map.empty else field (super m)) ++ map_of (flds m)"
apply (unfold field_def)
apply (erule (1) class_rec [THEN trans])
apply simp
done
end