Integ/Equiv.thy
author clasohm
Tue, 24 Oct 1995 14:59:17 +0100
changeset 251 f04b33ce250f
parent 249 492493334e0f
permissions -rw-r--r--
added calls of init_html and make_chart
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
216
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     1
(*  Title: 	Equiv.thy
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     2
    ID:         $Id$
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     3
    Authors: 	Riccardo Mattolini, Dip. Sistemi e Informatica
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     4
        	Lawrence C Paulson, Cambridge University Computer Laboratory
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     7
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     8
Equivalence relations in Higher-Order Set Theory 
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     9
*)
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    10
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    11
Equiv = Relation +
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    12
consts
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    13
    refl,equiv 	::      "['a set,('a*'a) set]=>bool"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    14
    sym         ::      "('a*'a) set=>bool"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    15
    "'/"        ::      "['a set,('a*'a) set]=>'a set set"  (infixl 90) 
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    16
                        (*set of equiv classes*)
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    17
    congruent	::	"[('a*'a) set,'a=>'b]=>bool"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    18
    congruent2  ::      "[('a*'a) set,['a,'a]=>'b]=>bool"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    19
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    20
defs
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    21
    refl_def      "refl(A,r) == r <= Sigma(A,%x.A) & (ALL x: A. <x,x> : r)"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    22
    sym_def       "sym(r)    == ALL x y. <x,y>: r --> <y,x>: r"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    23
    equiv_def     "equiv(A,r) == refl(A,r) & sym(r) & trans(r)"
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    24
    quotient_def  "A/r == UN x:A. {r^^{x}}"  
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    25
    congruent_def   "congruent(r,b)  == ALL y z. <y,z>:r --> b(y)=b(z)"
249
492493334e0f removed \...\ inside strings
clasohm
parents: 216
diff changeset
    26
    congruent2_def  "congruent2(r,b) == ALL y1 z1 y2 z2. 
492493334e0f removed \...\ inside strings
clasohm
parents: 216
diff changeset
    27
           <y1,z1>:r --> <y2,z2>:r --> b(y1,y2) = b(z1,z2)"
216
12943ab62cc5 Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    28
end