Set.thy
changeset 145 a9f7ff3a464c
parent 133 4a2bb4fbc168
child 191 ec175b039523
equal deleted inserted replaced
144:6254f50e5ec9 145:a9f7ff3a464c
     1 (*  Title:      HOL/set.thy
     1 (*  Title:      HOL/Set.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Tobias Nipkow
     3     Author:     Tobias Nipkow
     4     Copyright   1993  University of Cambridge
     4     Copyright   1993  University of Cambridge
     5 *)
     5 *)
     6 
     6 
     9 types
     9 types
    10   'a set
    10   'a set
    11 
    11 
    12 arities
    12 arities
    13   set :: (term) term
    13   set :: (term) term
    14   set :: (term) ord
       
    15   set :: (term) minus
       
    16 
    14 
       
    15 instance
       
    16   set :: (term) {ord, minus}
    17 
    17 
    18 consts
    18 consts
    19 
    19   "{}"          :: "'a set"                           ("{}")
    20   (* Constants *)
    20   insert        :: "['a, 'a set] => 'a set"
    21 
       
    22   Collect       :: "('a => bool) => 'a set"               (*comprehension*)
    21   Collect       :: "('a => bool) => 'a set"               (*comprehension*)
    23   Compl         :: "('a set) => 'a set"                   (*complement*)
    22   Compl         :: "('a set) => 'a set"                   (*complement*)
    24   Int           :: "['a set, 'a set] => 'a set"       (infixl 70)
    23   Int           :: "['a set, 'a set] => 'a set"       (infixl 70)
    25   Un            :: "['a set, 'a set] => 'a set"       (infixl 65)
    24   Un            :: "['a set, 'a set] => 'a set"       (infixl 65)
    26   UNION, INTER  :: "['a set, 'a => 'b set] => 'b set"     (*general*)
    25   UNION, INTER  :: "['a set, 'a => 'b set] => 'b set"     (*general*)
    27   UNION1        :: "['a => 'b set] => 'b set"         (binder "UN " 10)
    26   UNION1        :: "['a => 'b set] => 'b set"         (binder "UN " 10)
    28   INTER1        :: "['a => 'b set] => 'b set"         (binder "INT " 10)
    27   INTER1        :: "['a => 'b set] => 'b set"         (binder "INT " 10)
    29   Union, Inter  :: "(('a set)set) => 'a set"              (*of a set*)
    28   Union, Inter  :: "(('a set)set) => 'a set"              (*of a set*)
    30   Pow           :: "'a set => 'a set set"		  (*powerset*)
    29   Pow           :: "'a set => 'a set set"                 (*powerset*)
    31   range         :: "('a => 'b) => 'b set"                 (*of function*)
    30   range         :: "('a => 'b) => 'b set"                 (*of function*)
    32   Ball, Bex     :: "['a set, 'a => bool] => bool"         (*bounded quantifiers*)
    31   Ball, Bex     :: "['a set, 'a => bool] => bool"         (*bounded quantifiers*)
    33   inj, surj     :: "('a => 'b) => bool"                   (*inj/surjective*)
    32   inj, surj     :: "('a => 'b) => bool"                   (*inj/surjective*)
    34   inj_onto      :: "['a => 'b, 'a set] => bool"
    33   inj_onto      :: "['a => 'b, 'a set] => bool"
    35   "``"          :: "['a => 'b, 'a set] => ('b set)"   (infixl 90)
    34   "``"          :: "['a => 'b, 'a set] => ('b set)"   (infixl 90)
    36   ":"           :: "['a, 'a set] => bool"             (infixl 50) (*membership*)
    35   ":"           :: "['a, 'a set] => bool"             (infixl 50) (*membership*)
    37   "~:"          :: "['a, 'a set] => bool"             ("(_ ~:/ _)" [50, 51] 50)
       
    38 
       
    39   (* Finite Sets *)
       
    40 
       
    41   insert        :: "['a, 'a set] => 'a set"
       
    42   "{}"          :: "'a set"                           ("{}")
       
    43   "@Finset"     :: "args => 'a set"                   ("{(_)}")
       
    44 
    36 
    45 
    37 
    46   (** Binding Constants **)
    38 syntax
       
    39 
       
    40   "~:"          :: "['a, 'a set] => bool"             (infixl 50)
       
    41 
       
    42   "@Finset"     :: "args => 'a set"                   ("{(_)}")
    47 
    43 
    48   "@Coll"       :: "[idt, bool] => 'a set"            ("(1{_./ _})")
    44   "@Coll"       :: "[idt, bool] => 'a set"            ("(1{_./ _})")
    49   "@SetCompr"   :: "['a, idts, bool] => 'a set"       ("(1{_ |/_./ _})")
    45   "@SetCompr"   :: "['a, idts, bool] => 'a set"       ("(1{_ |/_./ _})")
    50 
    46 
    51   (* Big Intersection / Union *)
    47   (* Big Intersection / Union *)
    57 
    53 
    58   "@Ball"       :: "[idt, 'a set, bool] => bool"      ("(3! _:_./ _)" 10)
    54   "@Ball"       :: "[idt, 'a set, bool] => bool"      ("(3! _:_./ _)" 10)
    59   "@Bex"        :: "[idt, 'a set, bool] => bool"      ("(3? _:_./ _)" 10)
    55   "@Bex"        :: "[idt, 'a set, bool] => bool"      ("(3? _:_./ _)" 10)
    60   "*Ball"       :: "[idt, 'a set, bool] => bool"      ("(3ALL _:_./ _)" 10)
    56   "*Ball"       :: "[idt, 'a set, bool] => bool"      ("(3ALL _:_./ _)" 10)
    61   "*Bex"        :: "[idt, 'a set, bool] => bool"      ("(3EX _:_./ _)" 10)
    57   "*Bex"        :: "[idt, 'a set, bool] => bool"      ("(3EX _:_./ _)" 10)
    62 
       
    63 
    58 
    64 translations
    59 translations
    65   "x ~: y"      == "~ (x : y)"
    60   "x ~: y"      == "~ (x : y)"
    66   "{x, xs}"     == "insert(x, {xs})"
    61   "{x, xs}"     == "insert(x, {xs})"
    67   "{x}"         == "insert(x, {})"
    62   "{x}"         == "insert(x, {})"
    87   Bex_def       "Bex(A, P)      == ? x. x:A & P(x)"
    82   Bex_def       "Bex(A, P)      == ? x. x:A & P(x)"
    88   subset_def    "A <= B         == ! x:A. x:B"
    83   subset_def    "A <= B         == ! x:A. x:B"
    89   Compl_def     "Compl(A)       == {x. ~x:A}"
    84   Compl_def     "Compl(A)       == {x. ~x:A}"
    90   Un_def        "A Un B         == {x.x:A | x:B}"
    85   Un_def        "A Un B         == {x.x:A | x:B}"
    91   Int_def       "A Int B        == {x.x:A & x:B}"
    86   Int_def       "A Int B        == {x.x:A & x:B}"
    92   set_diff_def  "A-B            == {x. x:A & ~x:B}"
    87   set_diff_def  "A - B          == {x. x:A & ~x:B}"
    93   INTER_def     "INTER(A, B)    == {y. ! x:A. y: B(x)}"
    88   INTER_def     "INTER(A, B)    == {y. ! x:A. y: B(x)}"
    94   UNION_def     "UNION(A, B)    == {y. ? x:A. y: B(x)}"
    89   UNION_def     "UNION(A, B)    == {y. ? x:A. y: B(x)}"
    95   INTER1_def    "INTER1(B)      == INTER({x.True}, B)"
    90   INTER1_def    "INTER1(B)      == INTER({x.True}, B)"
    96   UNION1_def    "UNION1(B)      == UNION({x.True}, B)"
    91   UNION1_def    "UNION1(B)      == UNION({x.True}, B)"
    97   Inter_def     "Inter(S)       == (INT x:S. x)"
    92   Inter_def     "Inter(S)       == (INT x:S. x)"
    98   Union_def     "Union(S)       == (UN x:S. x)"
    93   Union_def     "Union(S)       == (UN x:S. x)"
    99   Pow_def	"Pow(A)         == {B. B <= A}"
    94   Pow_def       "Pow(A)         == {B. B <= A}"
   100   empty_def     "{}             == {x. False}"
    95   empty_def     "{}             == {x. False}"
   101   insert_def    "insert(a, B)   == {x.x=a} Un B"
    96   insert_def    "insert(a, B)   == {x.x=a} Un B"
   102   range_def     "range(f)       == {y. ? x. y=f(x)}"
    97   range_def     "range(f)       == {y. ? x. y=f(x)}"
   103   image_def     "f``A           == {y. ? x:A. y=f(x)}"
    98   image_def     "f``A           == {y. ? x:A. y=f(x)}"
   104   inj_def       "inj(f)         == ! x y. f(x)=f(y) --> x=y"
    99   inj_def       "inj(f)         == ! x y. f(x)=f(y) --> x=y"
   111 
   106 
   112 local
   107 local
   113 
   108 
   114 (* Translates between { e | x1..xn. P} and {u. ? x1..xn. u=e & P} *)
   109 (* Translates between { e | x1..xn. P} and {u. ? x1..xn. u=e & P} *)
   115 
   110 
   116 fun dummyC(s) = Const(s,dummyT);
       
   117 
       
   118 val ex_tr = snd(mk_binder_tr("? ","Ex"));
   111 val ex_tr = snd(mk_binder_tr("? ","Ex"));
   119 
   112 
   120 fun nvars(Const("_idts",_) $ _ $ idts) = nvars(idts)+1
   113 fun nvars(Const("_idts",_) $ _ $ idts) = nvars(idts)+1
   121   | nvars(_) = 1;
   114   | nvars(_) = 1;
   122 
   115 
   123 fun setcompr_tr[e,idts,b] =
   116 fun setcompr_tr[e,idts,b] =
   124   let val eq = dummyC("op =") $ Bound(nvars(idts)) $ e
   117   let val eq = Syntax.const("op =") $ Bound(nvars(idts)) $ e
   125       val P = dummyC("op &") $ eq $ b
   118       val P = Syntax.const("op &") $ eq $ b
   126       val exP = ex_tr [idts,P]
   119       val exP = ex_tr [idts,P]
   127   in dummyC("Collect") $ Abs("",dummyT,exP) end;
   120   in Syntax.const("Collect") $ Abs("",dummyT,exP) end;
   128 
   121 
   129 val ex_tr' = snd(mk_binder_tr' ("Ex","DUMMY"));
   122 val ex_tr' = snd(mk_binder_tr' ("Ex","DUMMY"));
   130 
   123 
   131 fun setcompr_tr'[Abs(_,_,P)] =
   124 fun setcompr_tr'[Abs(_,_,P)] =
   132   let fun ok(Const("Ex",_)$Abs(_,_,P),n) = ok(P,n+1)
   125   let fun ok(Const("Ex",_)$Abs(_,_,P),n) = ok(P,n+1)
   133         | ok(Const("op &",_) $ (Const("op =",_) $ Bound(m) $ _) $ _, n) =
   126         | ok(Const("op &",_) $ (Const("op =",_) $ Bound(m) $ _) $ _, n) =
   134             if n>0 andalso m=n then () else raise Match
   127             if n>0 andalso m=n then () else raise Match
   135 
   128 
   136       fun tr'(_ $ abs) =
   129       fun tr'(_ $ abs) =
   137         let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr'[abs]
   130         let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr'[abs]
   138         in dummyC("@SetCompr") $ e $ idts $ Q end
   131         in Syntax.const("@SetCompr") $ e $ idts $ Q end
   139   in ok(P,0); tr'(P) end;
   132   in ok(P,0); tr'(P) end;
   140 
   133 
   141 in
   134 in
   142 
   135 
   143 val parse_translation = [("@SetCompr",setcompr_tr)];
   136 val parse_translation = [("@SetCompr", setcompr_tr)];
   144 val print_translation = [("Collect",setcompr_tr')];
   137 val print_translation = [("Collect", setcompr_tr')];
   145 
       
   146 val print_ast_translation =
   138 val print_ast_translation =
   147   map HOL.alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")];
   139   map HOL.alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")];
   148 
   140 
   149 end;
   141 end;
   150 
   142