diff -r 0b217404522a -r dfc2e722fe47 src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Sun Aug 21 22:13:04 2011 +0200 +++ b/src/HOL/NanoJava/TypeRel.thy Sun Aug 21 22:13:04 2011 +0200 @@ -6,8 +6,11 @@ theory TypeRel imports Decl "~~/src/HOL/Library/Wfrec" begin -consts - subcls1 :: "(cname \ cname) set" --{* subclass *} +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) @@ -20,17 +23,9 @@ subcls1_syntax ("_ \C1 _" [71,71] 70) and subcls_syntax ("_ \C _" [71,71] 70) -consts - method :: "cname => (mname \ methd)" - field :: "cname => (fname \ 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 :: "ty => ty => bool" ("_ \ _" [71,71] 70) @@ -111,7 +106,8 @@ done --{* Methods of a class, with inheritance and hiding *} -defs method_def: "method C \ class_rec C methods" +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 empty else method (super m)) ++ map_of (methods m)" @@ -122,7 +118,8 @@ --{* Fields of a class, with inheritance and hiding *} -defs field_def: "field C \ class_rec C flds" +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 empty else field (super m)) ++ map_of (flds m)"