src/HOL/HOL.thy
changeset 11432 8a203ae6efe3
parent 10489 a4684cf28edf
child 11438 3d9222b80989
--- a/src/HOL/HOL.thy	Fri Jul 20 17:49:21 2001 +0200
+++ b/src/HOL/HOL.thy	Fri Jul 20 21:52:54 2001 +0200
@@ -38,6 +38,7 @@
   (* Binders *)
 
   Eps           :: "('a => bool) => 'a"
+  The           :: "('a => bool) => 'a"
   All           :: "('a => bool) => bool"           (binder "ALL " 10)
   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
@@ -97,6 +98,7 @@
 syntax
   ~=            :: "['a, 'a] => bool"                    (infixl 50)
   "_Eps"        :: "[pttrn, bool] => 'a"                 ("(3SOME _./ _)" [0, 10] 10)
+  "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
 
   (* Let expressions *)
 
@@ -115,6 +117,7 @@
 translations
   "x ~= y"                == "~ (x = y)"
   "SOME x. P"             == "Eps (%x. P)"
+  "THE x. P"              == "The (%x. P)"
   "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
   "let x = a in e"        == "Let a (%x. e)"
 
@@ -171,6 +174,7 @@
   ext:          "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"
 
   someI:        "P (x::'a) ==> P (SOME x. P x)"
+  the_eq_trivial: "(THE x. x = a) = (a::'a)"
 
   impI:         "(P ==> Q) ==> P-->Q"
   mp:           "[| P-->Q;  P |] ==> Q"