src/HOL/NanoJava/TypeRel.thy
changeset 61990 39e4a93ad36e
parent 59682 d662d096f72b
child 62390 842917225d56
equal deleted inserted replaced
61989:ba8c284d4725 61990:39e4a93ad36e
    13 definition subcls1 :: "(cname \<times> cname) set"
    13 definition subcls1 :: "(cname \<times> cname) set"
    14 where
    14 where
    15   "subcls1 \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c. class C = Some c \<and> super c=D)}"
    15   "subcls1 \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c. class C = Some c \<and> super c=D)}"
    16 
    16 
    17 abbreviation
    17 abbreviation
    18   subcls1_syntax :: "[cname, cname] => bool"  ("_ <=C1 _" [71,71] 70)
    18   subcls1_syntax :: "[cname, cname] => bool"  ("_ \<prec>C1 _" [71,71] 70)
    19   where "C <=C1 D == (C,D) \<in> subcls1"
    19   where "C \<prec>C1 D == (C,D) \<in> subcls1"
    20 abbreviation
    20 abbreviation
    21   subcls_syntax  :: "[cname, cname] => bool" ("_ <=C _"  [71,71] 70)
    21   subcls_syntax  :: "[cname, cname] => bool" ("_ \<preceq>C _"  [71,71] 70)
    22   where "C <=C D == (C,D) \<in> subcls1^*"
    22   where "C \<preceq>C D == (C,D) \<in> subcls1^*"
    23 
       
    24 notation (xsymbols)
       
    25   subcls1_syntax  ("_ \<prec>C1 _"  [71,71] 70) and
       
    26   subcls_syntax  ("_ \<preceq>C _"   [71,71] 70)
       
    27 
    23 
    28 
    24 
    29 subsection "Declarations and properties not used in the meta theory"
    25 subsection "Declarations and properties not used in the meta theory"
    30 
    26 
    31 text{* Widening, viz. method invocation conversion *}
    27 text{* Widening, viz. method invocation conversion *}