src/HOL/Vimage.thy
author wenzelm
Sun, 29 Nov 1998 13:16:47 +0100
changeset 5989 9670dae0143d
parent 4648 f04da668581c
child 7515 0c05469cad57
permissions -rw-r--r--
proof_general_trans (experimental);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4648
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     1
Vimage = Set +
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     2
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     3
consts
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     4
  "-``"          :: ['a => 'b, 'b set] => ('a set)   (infixr 90)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     5
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     6
defs
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     7
  vimage_def     "f-``B           == {x. f(x) : B}"
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     8
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     9
end