src/ZF/Order.thy
author oheimb
Wed, 12 Nov 1997 18:58:50 +0100
changeset 4223 f60e3d2c81d3
parent 2469 b50b8c0eec01
child 9683 f87c8c449018
permissions -rw-r--r--
added thin_refl to hyp_subst_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/Order.thy
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     5
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     6
Orders in Zermelo-Fraenkel Set Theory 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     7
*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     8
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     9
Order = WF + Perm + 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    10
consts
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    11
  part_ord        :: [i,i]=>o           (*Strict partial ordering*)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    12
  linear, tot_ord :: [i,i]=>o           (*Strict total ordering*)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    13
  well_ord        :: [i,i]=>o           (*Well-ordering*)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    14
  mono_map        :: [i,i,i,i]=>i       (*Order-preserving maps*)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    15
  ord_iso         :: [i,i,i,i]=>i       (*Order isomorphisms*)
1851
7b1e1c298e50 Corrected indentation
paulson
parents: 1478
diff changeset
    16
  pred            :: [i,i,i]=>i		(*Set of predecessors*)
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    17
  ord_iso_map     :: [i,i,i,i]=>i       (*Construction for linearity theorem*)
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    18
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents: 435
diff changeset
    19
defs
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    20
  part_ord_def "part_ord(A,r) == irrefl(A,r) & trans[A](r)"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    21
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    22
  linear_def   "linear(A,r) == (ALL x:A. ALL y:A. <x,y>:r | x=y | <y,x>:r)"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    23
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    24
  tot_ord_def  "tot_ord(A,r) == part_ord(A,r) & linear(A,r)"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    25
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    26
  well_ord_def "well_ord(A,r) == tot_ord(A,r) & wf[A](r)"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    27
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    28
  mono_map_def "mono_map(A,r,B,s) == 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    29
                   {f: A->B. ALL x:A. ALL y:A. <x,y>:r --> <f`x,f`y>:s}"
786
2a871417e7fc added constants mono_map, ord_iso_map
lcp
parents: 578
diff changeset
    30
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    31
  ord_iso_def  "ord_iso(A,r,B,s) == 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    32
                   {f: bij(A,B). ALL x:A. ALL y:A. <x,y>:r <-> <f`x,f`y>:s}"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    33
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    34
  pred_def     "pred(A,x,r) == {y:A. <y,x>:r}"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    35
786
2a871417e7fc added constants mono_map, ord_iso_map
lcp
parents: 578
diff changeset
    36
  ord_iso_map_def
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    37
     "ord_iso_map(A,r,B,s) == 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    38
       UN x:A. UN y:B. UN f: ord_iso(pred(A,x,r), r, pred(B,y,s), s).   
928a16e02f9f removed \...\ inside strings
clasohm
parents: 786
diff changeset
    39
            {<x,y>}"
786
2a871417e7fc added constants mono_map, ord_iso_map
lcp
parents: 578
diff changeset
    40
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1851
diff changeset
    41
constdefs
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1851
diff changeset
    42
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1851
diff changeset
    43
  first :: [i, i, i] => o
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1851
diff changeset
    44
    "first(u, X, R) == u:X & (ALL v:X. v~=u --> <u,v> : R)"
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1851
diff changeset
    45
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    46
end