equal
deleted
inserted
replaced
8 theory FuncSet |
8 theory FuncSet |
9 imports Main |
9 imports Main |
10 begin |
10 begin |
11 |
11 |
12 definition |
12 definition |
13 Pi :: "['a set, 'a => 'b set] => ('a => 'b) set" |
13 Pi :: "['a set, 'a => 'b set] => ('a => 'b) set" where |
14 "Pi A B = {f. \<forall>x. x \<in> A --> f x \<in> B x}" |
14 "Pi A B = {f. \<forall>x. x \<in> A --> f x \<in> B x}" |
15 |
15 |
16 extensional :: "'a set => ('a => 'b) set" |
16 definition |
|
17 extensional :: "'a set => ('a => 'b) set" where |
17 "extensional A = {f. \<forall>x. x~:A --> f x = arbitrary}" |
18 "extensional A = {f. \<forall>x. x~:A --> f x = arbitrary}" |
18 |
19 |
19 "restrict" :: "['a => 'b, 'a set] => ('a => 'b)" |
20 definition |
|
21 "restrict" :: "['a => 'b, 'a set] => ('a => 'b)" where |
20 "restrict f A = (%x. if x \<in> A then f x else arbitrary)" |
22 "restrict f A = (%x. if x \<in> A then f x else arbitrary)" |
21 |
23 |
22 abbreviation |
24 abbreviation |
23 funcset :: "['a set, 'b set] => ('a => 'b) set" (infixr "->" 60) |
25 funcset :: "['a set, 'b set] => ('a => 'b) set" |
|
26 (infixr "->" 60) where |
24 "A -> B == Pi A (%_. B)" |
27 "A -> B == Pi A (%_. B)" |
25 |
28 |
26 notation (xsymbols) |
29 notation (xsymbols) |
27 funcset (infixr "\<rightarrow>" 60) |
30 funcset (infixr "\<rightarrow>" 60) |
28 |
31 |
41 translations |
44 translations |
42 "PI x:A. B" == "CONST Pi A (%x. B)" |
45 "PI x:A. B" == "CONST Pi A (%x. B)" |
43 "%x:A. f" == "CONST restrict (%x. f) A" |
46 "%x:A. f" == "CONST restrict (%x. f) A" |
44 |
47 |
45 definition |
48 definition |
46 "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)" |
49 "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)" where |
47 "compose A g f = (\<lambda>x\<in>A. g (f x))" |
50 "compose A g f = (\<lambda>x\<in>A. g (f x))" |
48 |
51 |
49 |
52 |
50 subsection{*Basic Properties of @{term Pi}*} |
53 subsection{*Basic Properties of @{term Pi}*} |
51 |
54 |
140 |
143 |
141 text{*The basic definition could be moved to @{text "Fun.thy"}, but most of |
144 text{*The basic definition could be moved to @{text "Fun.thy"}, but most of |
142 the theorems belong here, or need at least @{term Hilbert_Choice}.*} |
145 the theorems belong here, or need at least @{term Hilbert_Choice}.*} |
143 |
146 |
144 definition |
147 definition |
145 bij_betw :: "['a => 'b, 'a set, 'b set] => bool" -- {* bijective *} |
148 bij_betw :: "['a => 'b, 'a set, 'b set] => bool" where -- {* bijective *} |
146 "bij_betw f A B = (inj_on f A & f ` A = B)" |
149 "bij_betw f A B = (inj_on f A & f ` A = B)" |
147 |
150 |
148 lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" |
151 lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" |
149 by (simp add: bij_betw_def) |
152 by (simp add: bij_betw_def) |
150 |
153 |