1475
|
1 |
(* Title: HOL/Fun.thy
|
923
|
2 |
ID: $Id$
|
1475
|
3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory
|
923
|
4 |
Copyright 1994 University of Cambridge
|
|
5 |
|
2912
|
6 |
Notions about functions.
|
923
|
7 |
*)
|
|
8 |
|
4648
|
9 |
Fun = Vimage +
|
2912
|
10 |
|
4059
|
11 |
instance set :: (term) order
|
|
12 |
(subset_refl,subset_trans,subset_antisym,psubset_eq)
|
2912
|
13 |
consts
|
|
14 |
|
5305
|
15 |
Id :: 'a => 'a
|
|
16 |
o :: ['b => 'c, 'a => 'b, 'a] => 'c (infixl 55)
|
4830
|
17 |
inj, surj :: ('a => 'b) => bool (*inj/surjective*)
|
|
18 |
inj_on :: ['a => 'b, 'a set] => bool
|
|
19 |
inv :: ('a => 'b) => ('b => 'a)
|
5305
|
20 |
fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)"
|
|
21 |
|
|
22 |
nonterminals
|
|
23 |
updbinds updbind
|
|
24 |
|
|
25 |
syntax
|
|
26 |
|
|
27 |
(* Let expressions *)
|
|
28 |
|
|
29 |
"_updbind" :: ['a, 'a] => updbind ("(2_ :=/ _)")
|
|
30 |
"" :: updbind => updbinds ("_")
|
|
31 |
"_updbinds" :: [updbind, updbinds] => updbinds ("_,/ _")
|
|
32 |
"_Update" :: ['a, updbinds] => 'a ("_/'((_)')" [900,0] 900)
|
|
33 |
|
|
34 |
translations
|
|
35 |
"_Update f (_updbinds b bs)" == "_Update (_Update f b) bs"
|
|
36 |
"f(x:=y)" == "fun_upd f x y"
|
2912
|
37 |
|
|
38 |
defs
|
|
39 |
|
5305
|
40 |
Id_def "Id == %x. x"
|
|
41 |
o_def "f o g == %x. f(g(x))"
|
|
42 |
inj_def "inj f == ! x y. f(x)=f(y) --> x=y"
|
|
43 |
inj_on_def "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y"
|
|
44 |
surj_def "surj f == ! y. ? x. y=f(x)"
|
|
45 |
inv_def "inv(f::'a=>'b) == % y. @x. f(x)=y"
|
|
46 |
fun_upd_def "f(a:=b) == % x. if x=a then b else f x"
|
2912
|
47 |
|
|
48 |
end
|