Set.thy
author lcp
Wed, 22 Sep 1993 15:43:05 +0200
changeset 2 befa4e9f7c90
parent 0 7949f97df77a
child 4 d199410f1db1
permissions -rw-r--r--
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong, split_weak_cong, nat_case_weak_cong, nat_rec_weak_cong. Used in llist.ML to make simplifications faster. HOL/gfp: re-ordered premises to put mono(f) early (first or right after A==gfp(f) in the def_ rules). Renamed some variables in rules, A to X and h to A. Renamed coinduct to weak_coinduct, coinduct2 to coinduct. Strengthened coinduct as suggested by j. Frost, to have the premise X <= f(X Un gfp(f)). HOL/llist: used stronger coinduct rule to strengthen LList_coinduct, LList_equalityI, llist_equalityI, llist_fun_equalityI and to delete the "2" form of those rules. Proved List_Fun_LList_I, LListD_Fun_diag_I and llistD_Fun_range_I to help use the new coinduction rules; most proofs involving them required small changes. Proved prod_fun_range_eq_diag as lemma for llist_equalityI.

(*  Title:      HOL/set.thy
    ID:         $Id$
    Author:     Tobias Nipkow
    Copyright   1993  University of Cambridge
*)

Set = Ord +

types
  set 1

arities
  set :: (term) term
  set :: (term) ord
  set :: (term) minus


consts

  (* Constants *)

  Collect       :: "('a => bool) => 'a set"                 (*comprehension*)
  Compl         :: "('a set) => 'a set"                     (*complement*)
  Int           :: "['a set, 'a set] => 'a set"         (infixl 70)
  Un            :: "['a set, 'a set] => 'a set"         (infixl 65)
  UNION, INTER  :: "['a set, 'a => 'b set] => 'b set"       (*general*)
  UNION1        :: "['a => 'b set] => 'b set"           (binder "UN " 10)
  INTER1        :: "['a => 'b set] => 'b set"           (binder "INT " 10)
  Union, Inter  :: "(('a set)set) => 'a set"                (*of a set*)
  range         :: "('a => 'b) => 'b set"                   (*of function*)
  Ball, Bex     :: "['a set, 'a => bool] => bool"           (*bounded quantifiers*)
  inj, surj     :: "('a => 'b) => bool"                     (*injective/surjective*)
  inj_onto      :: "['a => 'b, 'a set] => bool"
  "``"          :: "['a => 'b, 'a set] => ('b set)"     (infixl 90)
  ":"           :: "['a, 'a set] => bool" (infixl 50)       (*membership*)

  (* Finite Sets *)

  insert        :: "['a, 'a set] => 'a set"
  "{}"          :: "'a set"                             ("{}")
  "@Finset"     :: "args => 'a set"                     ("{(_)}")


  (** Binding Constants **)

  "@Coll"       :: "[idt, bool] => 'a set"              ("(1{_./ _})")  (*collection*)

  (* Big Intersection / Union *)

  "@INTER"      :: "[idt, 'a set, 'b set] => 'b set"    ("(3INT _:_./ _)" 10)
  "@UNION"      :: "[idt, 'a set, 'b set] => 'b set"    ("(3UN _:_./ _)" 10)

  (* Bounded Quantifiers *)

  "@Ball"       :: "[idt, 'a set, bool] => bool"        ("(3! _:_./ _)" 10)
  "@Bex"        :: "[idt, 'a set, bool] => bool"        ("(3? _:_./ _)" 10)
  "*Ball"       :: "[idt, 'a set, bool] => bool"        ("(3ALL _:_./ _)" 10)
  "*Bex"        :: "[idt, 'a set, bool] => bool"        ("(3EX _:_./ _)" 10)


translations
  "{x. P}"      == "Collect(%x. P)"
  "INT x:A. B"  == "INTER(A, %x. B)"
  "UN x:A. B"   == "UNION(A, %x. B)"
  "! x:A. P"    == "Ball(A, %x. P)"
  "? x:A. P"    == "Bex(A, %x. P)"
  "ALL x:A. P"  => "Ball(A, %x. P)"
  "EX x:A. P"   => "Bex(A, %x. P)"

  "{x, xs}"     == "insert(x, {xs})"
  "{x}"         == "insert(x, {})"


rules

  (* Isomorphisms between Predicates and Sets *)

  mem_Collect_eq    "(a : {x.P(x)}) = P(a)"
  Collect_mem_eq    "{x.x:A} = A"

  (* Definitions *)

  Ball_def      "Ball(A, P)     == ! x. x:A --> P(x)"
  Bex_def       "Bex(A, P)      == ? x. x:A & P(x)"
  subset_def    "A <= B         == ! x:A. x:B"
  Compl_def     "Compl(A)       == {x. ~x:A}"
  Un_def        "A Un B         == {x.x:A | x:B}"
  Int_def       "A Int B        == {x.x:A & x:B}"
  set_diff_def  "A-B            == {x. x:A & ~x:B}"
  INTER_def     "INTER(A, B)    == {y. ! x:A. y: B(x)}"
  UNION_def     "UNION(A, B)    == {y. ? x:A. y: B(x)}"
  INTER1_def    "INTER1(B)      == INTER({x.True}, B)"
  UNION1_def    "UNION1(B)      == UNION({x.True}, B)"
  Inter_def     "Inter(S)       == (INT x:S. x)"
  Union_def     "Union(S)       == (UN x:S. x)"
  empty_def     "{}             == {x. False}"
  insert_def    "insert(a, B)   == {x.x=a} Un B"
  range_def     "range(f)       == {y. ? x. y=f(x)}"
  image_def     "f``A           == {y. ? x:A. y=f(x)}"
  inj_def       "inj(f)         == ! x y. f(x)=f(y) --> x=y"
  inj_onto_def  "inj_onto(f, A) == ! x:A. ! y:A. f(x)=f(y) --> x=y"
  surj_def      "surj(f)        == ! y. ? x. y=f(x)"

end


ML

val print_ast_translation =
  map HOL.mk_alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")];