diff -r 143f58bb34f9 -r 0909deb8059b src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Thu May 26 16:57:14 2016 +0200 +++ b/src/HOL/NanoJava/TypeRel.thy Thu May 26 17:51:22 2016 +0200 @@ -8,7 +8,7 @@ imports Decl begin -text{* Direct subclass relation *} +text\Direct subclass relation\ definition subcls1 :: "(cname \ cname) set" where @@ -24,7 +24,7 @@ subsection "Declarations and properties not used in the meta theory" -text{* Widening, viz. method invocation conversion *} +text\Widening, viz. method invocation conversion\ inductive widen :: "ty => ty => bool" ("_ \ _" [71,71] 70) where @@ -103,7 +103,7 @@ apply (subst cut_apply, auto intro: subcls1I) done ---{* Methods of a class, with inheritance and hiding *} +\\Methods of a class, with inheritance and hiding\ definition "method" :: "cname => (mname \ methd)" where "method C \ class_rec C methods" @@ -115,7 +115,7 @@ done ---{* Fields of a class, with inheritance and hiding *} +\\Fields of a class, with inheritance and hiding\ definition field :: "cname => (fname \ ty)" where "field C \ class_rec C flds"