src/HOL/Quot/HQUOT.thy
author wenzelm
Wed, 22 Sep 1999 21:04:34 +0200
changeset 7577 644f9b4ae764
parent 3842 b55686a7b22c
permissions -rw-r--r--
proper theory setup for Real/ex/BinEx;

(*  Title:      HOL/Quot/HQUOT.thy
    ID:         $Id$
    Author:     Oscar Slotosch
    Copyright   1997 Technische Universitaet Muenchen

quotient constructor for higher order quotients

*)

HQUOT = PER +      

typedef 'a quot = "{s::'a::per set. ? r.!y. y:s=y===r}" (quotNE)

(* constants for equivalence classes *)
consts
        peclass         :: "'a::per => 'a quot"
        any_in          :: "'a::per quot => 'a"

syntax          "@ecl"  :: "'a::per => 'a quot" ("<[ _ ]>")

translations    "<[x]>" == "peclass x"

defs
        peclass_def     "<[x]> == Abs_quot {y. y===x}"
        any_in_def      "any_in f == @x.<[x]>=f"
end