diff -r ff7e6751a1a7 -r 9bbd5497befd src/HOL/Main.thy --- a/src/HOL/Main.thy Sat Nov 10 07:57:18 2018 +0000 +++ b/src/HOL/Main.thy Sat Nov 10 07:57:19 2018 +0000 @@ -11,12 +11,50 @@ Quickcheck_Narrowing Extraction Nunchaku - BNF_Greatest_Fixpoint Filter + BNF_Greatest_Fixpoint + Filter Conditionally_Complete_Lattices Binomial GCD begin +text \Legacy\ + +context Inf +begin + +abbreviation (input) INFIMUM :: "'b set \ ('b \ 'a) \ 'a" + where "INFIMUM A f \ \(f ` A)" + +end + +context Sup +begin + +abbreviation (input) SUPREMUM :: "'b set \ ('b \ 'a) \ 'a" + where "SUPREMUM A f \ \(f ` A)" + +end + +abbreviation (input) INTER :: "'a set \ ('a \ 'b set) \ 'b set" + where "INTER \ INFIMUM" + +abbreviation (input) UNION :: "'a set \ ('a \ 'b set) \ 'b set" + where "UNION \ SUPREMUM" + + +text \Namespace cleanup\ + +hide_const (open) + czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect + fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift + shift proj id_bnf + +hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV + + +text \Syntax cleanup\ + no_notation bot ("\") and top ("\") and @@ -29,17 +67,10 @@ ordLess2 (infix "(_,/ _)\") - -hide_const (open) - czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect - fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift - shift proj id_bnf - -hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV + BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and + BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and + BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90) and + BNF_Def.convol ("\(_,/ _)\") no_syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10)