(* Title: HOL/NanoJava/TypeRel.thy ID: $Id: TypeRel.thy,v 1.10 2005/06/17 14:13:09 haftmann Exp $ Author: David von Oheimb Copyright 2001 Technische Universitaet Muenchen *) header "Type relations" theory TypeRel imports Decl begin consts widen :: "(ty × ty ) set" --{* widening *} subcls1 :: "(cname × cname) set" --{* subclass *} syntax (xsymbols) widen :: "[ty , ty ] => bool" ("_ \<preceq> _" [71,71] 70) subcls1 :: "[cname, cname] => bool" ("_ \<prec>C1 _" [71,71] 70) subcls :: "[cname, cname] => bool" ("_ \<preceq>C _" [71,71] 70) syntax widen :: "[ty , ty ] => bool" ("_ <= _" [71,71] 70) subcls1 :: "[cname, cname] => bool" ("_ <=C1 _" [71,71] 70) subcls :: "[cname, cname] => bool" ("_ <=C _" [71,71] 70) translations "C \<prec>C1 D" == "(C,D) ∈ subcls1" "C \<preceq>C D" == "(C,D) ∈ subcls1^*" "S \<preceq> T" == "(S,T) ∈ widen" consts method :: "cname => (mname \<rightharpoonup> methd)" field :: "cname => (fname \<rightharpoonup> ty)" subsection "Declarations and properties not used in the meta theory" text{* Direct subclass relation *} defs subcls1_def: "subcls1 ≡ {(C,D). C≠Object ∧ (∃c. class C = Some c ∧ super c=D)}" text{* Widening, viz. method invocation conversion *} inductive widen intros refl [intro!, simp]: "T \<preceq> T" subcls : "C\<preceq>C D ==> Class C \<preceq> Class D" null [intro!]: "NT \<preceq> R" lemma subcls1D: "C\<prec>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\<prec>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 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 constdefs ws_prog :: "bool" "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^-1 ∩ subcls1^+ = {}" by (fast dest: subcls1D ws_progD) (* irrefl_tranclI in Transitive_Closure.thy is more general *) lemma irrefl_tranclI': "r^-1 Int 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, standard] lemma wf_subcls1: "ws_prog ==> wf (subcls1¯)" by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic) consts class_rec ::"cname => (class => ('a × 'b) list) => ('a \<rightharpoonup> 'b)" recdef (permissive) class_rec "subcls1¯" "class_rec C = (λf. case class C of None => arbitrary | Some m => if wf (subcls1¯) then (if C=Object then empty else class_rec (super m) f) ++ map_of (f m) else arbitrary)" (hints intro: subcls1I) lemma class_rec: "[|class C = Some m; ws_prog|] ==> class_rec C f = (if C = Object then empty else class_rec (super m) f) ++ map_of (f m)"; apply (drule wf_subcls1) apply (rule class_rec.simps [THEN trans [THEN fun_cong]]) apply assumption apply simp done --{* Methods of a class, with inheritance and hiding *} defs method_def: "method C ≡ class_rec C methods" lemma method_rec: "[|class C = Some m; ws_prog|] ==> method C = (if C=Object then empty else method (super m)) ++ map_of (methods m)" apply (unfold method_def) apply (erule (1) class_rec [THEN trans]); apply simp done --{* Fields of a class, with inheritance and hiding *} defs field_def: "field C ≡ class_rec C flds" lemma flds_rec: "[|class C = Some m; ws_prog|] ==> field C = (if C=Object then empty else field (super m)) ++ map_of (flds m)" apply (unfold field_def) apply (erule (1) class_rec [THEN trans]); apply simp done end
lemma subcls1D:
C <=C1 D ==> C ≠ Object ∧ (∃c. class C = Some c ∧ super c = D)
lemma subcls1I:
[| class C = Some m; super m = D; C ≠ Object |] ==> C <=C1 D
lemma subcls1_def2:
subcls1 = (SIGMA C:{C. is_class C}. {D. C ≠ Object ∧ super (the (class C)) = D})
lemma finite_subcls1:
finite subcls1
lemma ws_progD:
[| class C = Some c; C ≠ Object; ws_prog |] ==> is_class (super c) ∧ (super c, C) ∉ subcls1+
lemma subcls1_irrefl_lemma1:
ws_prog ==> subcls1^-1 ∩ subcls1+ = {}
lemma irrefl_tranclI':
r^-1 ∩ r+ = {} ==> ∀x. (x, x) ∉ r+
lemmas subcls1_irrefl_lemma2:
ws_prog ==> ∀x. (x, x) ∉ subcls1+
lemmas subcls1_irrefl_lemma2:
ws_prog ==> ∀x. (x, x) ∉ subcls1+
lemma subcls1_irrefl:
[| x <=C1 y; ws_prog |] ==> x ≠ y
lemmas subcls1_acyclic:
ws_prog ==> acyclic subcls1
lemmas subcls1_acyclic:
ws_prog ==> acyclic subcls1
lemma wf_subcls1:
ws_prog ==> wf (subcls1^-1)
lemma class_rec:
[| class C = Some m; ws_prog |] ==> class_rec C f = (if C = Object then empty else class_rec (super m) f) ++ map_of (f m)
lemma method_rec:
[| class C = Some m; ws_prog |] ==> method C = (if C = Object then empty else method (super m)) ++ map_of (methods m)
lemma flds_rec:
[| class C = Some m; ws_prog |] ==> field C = (if C = Object then empty else field (super m)) ++ map_of (flds m)