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 |
|
10212
|
9 |
Fun = Inverse_Image + equalities +
|
2912
|
10 |
|
4059
|
11 |
instance set :: (term) order
|
|
12 |
(subset_refl,subset_trans,subset_antisym,psubset_eq)
|
6171
|
13 |
consts
|
|
14 |
fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)"
|
|
15 |
|
9141
|
16 |
nonterminals
|
|
17 |
updbinds updbind
|
5305
|
18 |
syntax
|
|
19 |
"_updbind" :: ['a, 'a] => updbind ("(2_ :=/ _)")
|
|
20 |
"" :: updbind => updbinds ("_")
|
|
21 |
"_updbinds" :: [updbind, updbinds] => updbinds ("_,/ _")
|
8258
|
22 |
"_Update" :: ['a, updbinds] => 'a ("_/'((_)')" [1000,0] 900)
|
5305
|
23 |
|
|
24 |
translations
|
|
25 |
"_Update f (_updbinds b bs)" == "_Update (_Update f b) bs"
|
|
26 |
"f(x:=y)" == "fun_upd f x y"
|
2912
|
27 |
|
|
28 |
defs
|
6171
|
29 |
fun_upd_def "f(a:=b) == % x. if x=a then b else f x"
|
2912
|
30 |
|
9340
|
31 |
(* Hint: to define the sum of two functions (or maps), use sum_case.
|
|
32 |
A nice infix syntax could be defined (in Datatype.thy or below) by
|
|
33 |
consts
|
|
34 |
fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
|
|
35 |
translations
|
|
36 |
"fun_sum" == "sum_case"
|
|
37 |
*)
|
6171
|
38 |
|
|
39 |
constdefs
|
|
40 |
id :: 'a => 'a
|
|
41 |
"id == %x. x"
|
|
42 |
|
|
43 |
o :: ['b => 'c, 'a => 'b, 'a] => 'c (infixl 55)
|
|
44 |
"f o g == %x. f(g(x))"
|
7374
|
45 |
|
|
46 |
inv :: ('a => 'b) => ('b => 'a)
|
|
47 |
"inv(f::'a=>'b) == % y. @x. f(x)=y"
|
6171
|
48 |
|
|
49 |
inj_on :: ['a => 'b, 'a set] => bool
|
|
50 |
"inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y"
|
2912
|
51 |
|
9352
|
52 |
syntax (symbols)
|
|
53 |
"op o" :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixl "\\<circ>" 55)
|
|
54 |
|
6171
|
55 |
syntax
|
|
56 |
inj :: ('a => 'b) => bool (*injective*)
|
|
57 |
|
|
58 |
translations
|
|
59 |
"inj f" == "inj_on f UNIV"
|
5852
|
60 |
|
7374
|
61 |
constdefs
|
|
62 |
surj :: ('a => 'b) => bool (*surjective*)
|
|
63 |
"surj f == ! y. ? x. y=f(x)"
|
|
64 |
|
|
65 |
bij :: ('a => 'b) => bool (*bijective*)
|
|
66 |
"bij f == inj f & surj f"
|
|
67 |
|
|
68 |
|
5852
|
69 |
(*The Pi-operator, by Florian Kammueller*)
|
|
70 |
|
|
71 |
constdefs
|
|
72 |
Pi :: "['a set, 'a => 'b set] => ('a => 'b) set"
|
|
73 |
"Pi A B == {f. ! x. if x:A then f(x) : B(x) else f(x) = (@ y. True)}"
|
|
74 |
|
|
75 |
restrict :: "['a => 'b, 'a set] => ('a => 'b)"
|
|
76 |
"restrict f A == (%x. if x : A then f x else (@ y. True))"
|
|
77 |
|
|
78 |
syntax
|
|
79 |
"@Pi" :: "[idt, 'a set, 'b set] => ('a => 'b) set" ("(3PI _:_./ _)" 10)
|
|
80 |
funcset :: "['a set, 'b set] => ('a => 'b) set" (infixr 60)
|
|
81 |
"@lam" :: "[pttrn, 'a set, 'a => 'b] => ('a => 'b)" ("(3lam _:_./ _)" 10)
|
|
82 |
|
|
83 |
(*Giving funcset the nice arrow syntax -> clashes with existing theories*)
|
|
84 |
|
|
85 |
translations
|
|
86 |
"PI x:A. B" => "Pi A (%x. B)"
|
|
87 |
"A funcset B" => "Pi A (_K B)"
|
|
88 |
"lam x:A. f" == "restrict (%x. f) A"
|
|
89 |
|
|
90 |
constdefs
|
|
91 |
Applyall :: "[('a => 'b) set, 'a]=> 'b set"
|
|
92 |
"Applyall F a == (%f. f a) `` F"
|
|
93 |
|
9309
|
94 |
compose :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)"
|
5852
|
95 |
"compose A g f == lam x : A. g(f x)"
|
|
96 |
|
|
97 |
Inv :: "['a set, 'a => 'b] => ('b => 'a)"
|
|
98 |
"Inv A f == (% x. (@ y. y : A & f y = x))"
|
|
99 |
|
|
100 |
|
2912
|
101 |
end
|
5852
|
102 |
|
|
103 |
ML
|
|
104 |
val print_translation = [("Pi", dependent_tr' ("@Pi", "op funcset"))];
|