src/HOLCF/Ssum.thy
changeset 35427 ad039d29e01c
parent 33808 31169fdc5ae7
child 35547 991a6af75978
equal deleted inserted replaced
35426:c9b9d4fc270d 35427:ad039d29e01c
    22 by (rule typedef_finite_po [OF type_definition_Ssum])
    22 by (rule typedef_finite_po [OF type_definition_Ssum])
    23 
    23 
    24 instance "++" :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
    24 instance "++" :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
    25 by (rule typedef_chfin [OF type_definition_Ssum below_Ssum_def])
    25 by (rule typedef_chfin [OF type_definition_Ssum below_Ssum_def])
    26 
    26 
    27 syntax (xsymbols)
    27 type_notation (xsymbols)
    28   "++"          :: "[type, type] => type"       ("(_ \<oplus>/ _)" [21, 20] 20)
    28   "++"  ("(_ \<oplus>/ _)" [21, 20] 20)
    29 syntax (HTML output)
    29 type_notation (HTML output)
    30   "++"          :: "[type, type] => type"       ("(_ \<oplus>/ _)" [21, 20] 20)
    30   "++"  ("(_ \<oplus>/ _)" [21, 20] 20)
    31 
    31 
    32 subsection {* Definitions of constructors *}
    32 subsection {* Definitions of constructors *}
    33 
    33 
    34 definition
    34 definition
    35   sinl :: "'a \<rightarrow> ('a ++ 'b)" where
    35   sinl :: "'a \<rightarrow> ('a ++ 'b)" where