clasohm@923: (* Title: HOL/Set.thy clasohm@923: ID: $Id$ clasohm@923: Author: Tobias Nipkow clasohm@923: Copyright 1993 University of Cambridge clasohm@923: *) clasohm@923: clasohm@923: Set = Ord + clasohm@923: clasohm@923: types clasohm@923: 'a set clasohm@923: clasohm@923: arities clasohm@923: set :: (term) term clasohm@923: clasohm@923: instance clasohm@923: set :: (term) {ord, minus} clasohm@923: clasohm@923: consts clasohm@923: "{}" :: "'a set" ("{}") clasohm@923: insert :: "['a, 'a set] => 'a set" clasohm@923: Collect :: "('a => bool) => 'a set" (*comprehension*) clasohm@923: Compl :: "('a set) => 'a set" (*complement*) clasohm@923: Int :: "['a set, 'a set] => 'a set" (infixl 70) clasohm@923: Un :: "['a set, 'a set] => 'a set" (infixl 65) clasohm@923: UNION, INTER :: "['a set, 'a => 'b set] => 'b set" (*general*) clasohm@923: UNION1 :: "['a => 'b set] => 'b set" (binder "UN " 10) clasohm@923: INTER1 :: "['a => 'b set] => 'b set" (binder "INT " 10) clasohm@923: Union, Inter :: "(('a set)set) => 'a set" (*of a set*) clasohm@923: Pow :: "'a set => 'a set set" (*powerset*) clasohm@923: range :: "('a => 'b) => 'b set" (*of function*) clasohm@923: Ball, Bex :: "['a set, 'a => bool] => bool" (*bounded quantifiers*) clasohm@923: inj, surj :: "('a => 'b) => bool" (*inj/surjective*) clasohm@923: inj_onto :: "['a => 'b, 'a set] => bool" clasohm@923: "``" :: "['a => 'b, 'a set] => ('b set)" (infixl 90) clasohm@923: ":" :: "['a, 'a set] => bool" (infixl 50) (*membership*) clasohm@923: clasohm@923: clasohm@923: syntax clasohm@923: clasohm@923: "~:" :: "['a, 'a set] => bool" (infixl 50) clasohm@923: clasohm@923: "@Finset" :: "args => 'a set" ("{(_)}") clasohm@923: clasohm@923: "@Coll" :: "[idt, bool] => 'a set" ("(1{_./ _})") clasohm@923: "@SetCompr" :: "['a, idts, bool] => 'a set" ("(1{_ |/_./ _})") clasohm@923: clasohm@923: (* Big Intersection / Union *) clasohm@923: clasohm@923: "@INTER" :: "[idt, 'a set, 'b set] => 'b set" ("(3INT _:_./ _)" 10) clasohm@923: "@UNION" :: "[idt, 'a set, 'b set] => 'b set" ("(3UN _:_./ _)" 10) clasohm@923: clasohm@923: (* Bounded Quantifiers *) clasohm@923: clasohm@923: "@Ball" :: "[idt, 'a set, bool] => bool" ("(3! _:_./ _)" 10) clasohm@923: "@Bex" :: "[idt, 'a set, bool] => bool" ("(3? _:_./ _)" 10) clasohm@923: "*Ball" :: "[idt, 'a set, bool] => bool" ("(3ALL _:_./ _)" 10) clasohm@923: "*Bex" :: "[idt, 'a set, bool] => bool" ("(3EX _:_./ _)" 10) clasohm@923: clasohm@923: translations clasohm@923: "x ~: y" == "~ (x : y)" clasohm@923: "{x, xs}" == "insert x {xs}" clasohm@923: "{x}" == "insert x {}" clasohm@923: "{x. P}" == "Collect (%x. P)" clasohm@923: "INT x:A. B" == "INTER A (%x. B)" clasohm@923: "UN x:A. B" == "UNION A (%x. B)" clasohm@923: "! x:A. P" == "Ball A (%x. P)" clasohm@923: "? x:A. P" == "Bex A (%x. P)" clasohm@923: "ALL x:A. P" => "Ball A (%x. P)" clasohm@923: "EX x:A. P" => "Bex A (%x. P)" clasohm@923: clasohm@923: clasohm@923: rules clasohm@923: clasohm@923: (* Isomorphisms between Predicates and Sets *) clasohm@923: clasohm@923: mem_Collect_eq "(a : {x.P(x)}) = P(a)" clasohm@923: Collect_mem_eq "{x.x:A} = A" clasohm@923: clasohm@923: clasohm@923: defs clasohm@923: Ball_def "Ball A P == ! x. x:A --> P(x)" clasohm@923: Bex_def "Bex A P == ? x. x:A & P(x)" clasohm@923: subset_def "A <= B == ! x:A. x:B" clasohm@923: Compl_def "Compl(A) == {x. ~x:A}" clasohm@923: Un_def "A Un B == {x.x:A | x:B}" clasohm@923: Int_def "A Int B == {x.x:A & x:B}" clasohm@923: set_diff_def "A - B == {x. x:A & ~x:B}" clasohm@923: INTER_def "INTER A B == {y. ! x:A. y: B(x)}" clasohm@923: UNION_def "UNION A B == {y. ? x:A. y: B(x)}" clasohm@923: INTER1_def "INTER1(B) == INTER {x.True} B" clasohm@923: UNION1_def "UNION1(B) == UNION {x.True} B" clasohm@923: Inter_def "Inter(S) == (INT x:S. x)" clasohm@923: Union_def "Union(S) == (UN x:S. x)" clasohm@923: Pow_def "Pow(A) == {B. B <= A}" clasohm@923: empty_def "{} == {x. False}" clasohm@923: insert_def "insert a B == {x.x=a} Un B" clasohm@923: range_def "range(f) == {y. ? x. y=f(x)}" clasohm@923: image_def "f``A == {y. ? x:A. y=f(x)}" clasohm@923: inj_def "inj(f) == ! x y. f(x)=f(y) --> x=y" clasohm@923: inj_onto_def "inj_onto f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" clasohm@923: surj_def "surj(f) == ! y. ? x. y=f(x)" clasohm@923: clasohm@923: end clasohm@923: clasohm@923: ML clasohm@923: clasohm@923: local clasohm@923: clasohm@923: (* Translates between { e | x1..xn. P} and {u. ? x1..xn. u=e & P} *) clasohm@923: (* {y. ? x1..xn. y = e & P} is only translated if [0..n] subset bvs(e) *) clasohm@923: clasohm@923: val ex_tr = snd(mk_binder_tr("? ","Ex")); clasohm@923: clasohm@923: fun nvars(Const("_idts",_) $ _ $ idts) = nvars(idts)+1 clasohm@923: | nvars(_) = 1; clasohm@923: clasohm@923: fun setcompr_tr[e,idts,b] = clasohm@923: let val eq = Syntax.const("op =") $ Bound(nvars(idts)) $ e clasohm@923: val P = Syntax.const("op &") $ eq $ b clasohm@923: val exP = ex_tr [idts,P] clasohm@923: in Syntax.const("Collect") $ Abs("",dummyT,exP) end; clasohm@923: clasohm@923: val ex_tr' = snd(mk_binder_tr' ("Ex","DUMMY")); clasohm@923: clasohm@923: fun setcompr_tr'[Abs(_,_,P)] = clasohm@923: let fun ok(Const("Ex",_)$Abs(_,_,P),n) = ok(P,n+1) clasohm@923: | ok(Const("op &",_) $ (Const("op =",_) $ Bound(m) $ e) $ _, n) = clasohm@923: if n>0 andalso m=n andalso clasohm@923: ((0 upto (n-1)) subset add_loose_bnos(e,0,[])) clasohm@923: then () else raise Match clasohm@923: clasohm@923: fun tr'(_ $ abs) = clasohm@923: let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr'[abs] clasohm@923: in Syntax.const("@SetCompr") $ e $ idts $ Q end clasohm@923: in ok(P,0); tr'(P) end; clasohm@923: clasohm@923: in clasohm@923: clasohm@923: val parse_translation = [("@SetCompr", setcompr_tr)]; clasohm@923: val print_translation = [("Collect", setcompr_tr')]; clasohm@923: val print_ast_translation = clasohm@923: map HOL.alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")]; clasohm@923: clasohm@923: end; clasohm@923: