| author | wenzelm | 
| Fri, 20 Aug 2010 11:47:33 +0200 | |
| changeset 38566 | 8176107637ce | 
| parent 37767 | a2b7a20d6ea3 | 
| child 38620 | b40524b74f77 | 
| permissions | -rw-r--r-- | 
| 1475 | 1 | (* Title: HOL/Fun.thy | 
| 2 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | |
| 923 | 3 | Copyright 1994 University of Cambridge | 
| 18154 | 4 | *) | 
| 923 | 5 | |
| 18154 | 6 | header {* Notions about functions *}
 | 
| 923 | 7 | |
| 15510 | 8 | theory Fun | 
| 32139 | 9 | imports Complete_Lattice | 
| 15131 | 10 | begin | 
| 2912 | 11 | |
| 26147 | 12 | text{*As a simplification rule, it replaces all function equalities by
 | 
| 13 | first-order equalities.*} | |
| 14 | lemma expand_fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)" | |
| 15 | apply (rule iffI) | |
| 16 | apply (simp (no_asm_simp)) | |
| 17 | apply (rule ext) | |
| 18 | apply (simp (no_asm_simp)) | |
| 19 | done | |
| 5305 | 20 | |
| 26147 | 21 | lemma apply_inverse: | 
| 26357 | 22 | "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u" | 
| 26147 | 23 | by auto | 
| 2912 | 24 | |
| 12258 | 25 | |
| 26147 | 26 | subsection {* The Identity Function @{text id} *}
 | 
| 6171 | 27 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 28 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 29 | id :: "'a \<Rightarrow> 'a" | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 30 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 31 | "id = (\<lambda>x. x)" | 
| 13910 | 32 | |
| 26147 | 33 | lemma id_apply [simp]: "id x = x" | 
| 34 | by (simp add: id_def) | |
| 35 | ||
| 36 | lemma image_ident [simp]: "(%x. x) ` Y = Y" | |
| 37 | by blast | |
| 38 | ||
| 39 | lemma image_id [simp]: "id ` Y = Y" | |
| 40 | by (simp add: id_def) | |
| 41 | ||
| 42 | lemma vimage_ident [simp]: "(%x. x) -` Y = Y" | |
| 43 | by blast | |
| 44 | ||
| 45 | lemma vimage_id [simp]: "id -` A = A" | |
| 46 | by (simp add: id_def) | |
| 47 | ||
| 48 | ||
| 49 | subsection {* The Composition Operator @{text "f \<circ> g"} *}
 | |
| 50 | ||
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 51 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 52 |   comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
 | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 53 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 54 | "f o g = (\<lambda>x. f (g x))" | 
| 11123 | 55 | |
| 21210 | 56 | notation (xsymbols) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 57 | comp (infixl "\<circ>" 55) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 58 | |
| 21210 | 59 | notation (HTML output) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 60 | comp (infixl "\<circ>" 55) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 61 | |
| 13585 | 62 | text{*compatibility*}
 | 
| 63 | lemmas o_def = comp_def | |
| 2912 | 64 | |
| 13585 | 65 | lemma o_apply [simp]: "(f o g) x = f (g x)" | 
| 66 | by (simp add: comp_def) | |
| 67 | ||
| 68 | lemma o_assoc: "f o (g o h) = f o g o h" | |
| 69 | by (simp add: comp_def) | |
| 70 | ||
| 71 | lemma id_o [simp]: "id o g = g" | |
| 72 | by (simp add: comp_def) | |
| 73 | ||
| 74 | lemma o_id [simp]: "f o id = f" | |
| 75 | by (simp add: comp_def) | |
| 76 | ||
| 34150 | 77 | lemma o_eq_dest: | 
| 78 | "a o b = c o d \<Longrightarrow> a (b v) = c (d v)" | |
| 79 | by (simp only: o_def) (fact fun_cong) | |
| 80 | ||
| 81 | lemma o_eq_elim: | |
| 82 | "a o b = c o d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R" | |
| 83 | by (erule meta_mp) (fact o_eq_dest) | |
| 84 | ||
| 13585 | 85 | lemma image_compose: "(f o g) ` r = f`(g`r)" | 
| 86 | by (simp add: comp_def, blast) | |
| 87 | ||
| 33044 | 88 | lemma vimage_compose: "(g \<circ> f) -` x = f -` (g -` x)" | 
| 89 | by auto | |
| 90 | ||
| 13585 | 91 | lemma UN_o: "UNION A (g o f) = UNION (f`A) g" | 
| 92 | by (unfold comp_def, blast) | |
| 93 | ||
| 94 | ||
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26357diff
changeset | 95 | subsection {* The Forward Composition Operator @{text fcomp} *}
 | 
| 26357 | 96 | |
| 97 | definition | |
| 37751 | 98 |   fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60)
 | 
| 26357 | 99 | where | 
| 37751 | 100 | "f \<circ>> g = (\<lambda>x. g (f x))" | 
| 26357 | 101 | |
| 37751 | 102 | lemma fcomp_apply [simp]: "(f \<circ>> g) x = g (f x)" | 
| 26357 | 103 | by (simp add: fcomp_def) | 
| 104 | ||
| 37751 | 105 | lemma fcomp_assoc: "(f \<circ>> g) \<circ>> h = f \<circ>> (g \<circ>> h)" | 
| 26357 | 106 | by (simp add: fcomp_def) | 
| 107 | ||
| 37751 | 108 | lemma id_fcomp [simp]: "id \<circ>> g = g" | 
| 26357 | 109 | by (simp add: fcomp_def) | 
| 110 | ||
| 37751 | 111 | lemma fcomp_id [simp]: "f \<circ>> id = f" | 
| 26357 | 112 | by (simp add: fcomp_def) | 
| 113 | ||
| 31202 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
31080diff
changeset | 114 | code_const fcomp | 
| 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
31080diff
changeset | 115 | (Eval infixl 1 "#>") | 
| 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
31080diff
changeset | 116 | |
| 37751 | 117 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26357diff
changeset | 118 | |
| 26357 | 119 | |
| 26147 | 120 | subsection {* Injectivity and Surjectivity *}
 | 
| 121 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 122 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 123 | inj_on :: "['a => 'b, 'a set] => bool" where | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 124 | -- "injective" | 
| 26147 | 125 | "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" | 
| 126 | ||
| 127 | text{*A common special case: functions injective over the entire domain type.*}
 | |
| 128 | ||
| 129 | abbreviation | |
| 130 | "inj f == inj_on f UNIV" | |
| 13585 | 131 | |
| 26147 | 132 | definition | 
| 133 |   bij_betw :: "('a => 'b) => 'a set => 'b set => bool" where -- "bijective"
 | |
| 37767 | 134 | "bij_betw f A B \<longleftrightarrow> inj_on f A & f ` A = B" | 
| 26147 | 135 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 136 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 137 |   surj :: "('a => 'b) => bool" where
 | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 138 | -- "surjective" | 
| 26147 | 139 | "surj f == ! y. ? x. y=f(x)" | 
| 13585 | 140 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 141 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 142 |   bij :: "('a => 'b) => bool" where
 | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 143 | -- "bijective" | 
| 26147 | 144 | "bij f == inj f & surj f" | 
| 145 | ||
| 146 | lemma injI: | |
| 147 | assumes "\<And>x y. f x = f y \<Longrightarrow> x = y" | |
| 148 | shows "inj f" | |
| 149 | using assms unfolding inj_on_def by auto | |
| 13585 | 150 | |
| 31775 | 151 | text{*For Proofs in @{text "Tools/Datatype/datatype_rep_proofs"}*}
 | 
| 13585 | 152 | lemma datatype_injI: | 
| 153 | "(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)" | |
| 154 | by (simp add: inj_on_def) | |
| 155 | ||
| 13637 | 156 | theorem range_ex1_eq: "inj f \<Longrightarrow> b : range f = (EX! x. b = f x)" | 
| 157 | by (unfold inj_on_def, blast) | |
| 158 | ||
| 13585 | 159 | lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y" | 
| 160 | by (simp add: inj_on_def) | |
| 161 | ||
| 32988 | 162 | lemma inj_on_eq_iff: "inj_on f A ==> x:A ==> y:A ==> (f(x) = f(y)) = (x=y)" | 
| 13585 | 163 | by (force simp add: inj_on_def) | 
| 164 | ||
| 32988 | 165 | lemma inj_eq: "inj f ==> (f(x) = f(y)) = (x=y)" | 
| 166 | by (simp add: inj_on_eq_iff) | |
| 167 | ||
| 26147 | 168 | lemma inj_on_id[simp]: "inj_on id A" | 
| 169 | by (simp add: inj_on_def) | |
| 13585 | 170 | |
| 26147 | 171 | lemma inj_on_id2[simp]: "inj_on (%x. x) A" | 
| 172 | by (simp add: inj_on_def) | |
| 173 | ||
| 174 | lemma surj_id[simp]: "surj id" | |
| 175 | by (simp add: surj_def) | |
| 176 | ||
| 177 | lemma bij_id[simp]: "bij id" | |
| 34209 | 178 | by (simp add: bij_def) | 
| 13585 | 179 | |
| 180 | lemma inj_onI: | |
| 181 | "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" | |
| 182 | by (simp add: inj_on_def) | |
| 183 | ||
| 184 | lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A" | |
| 185 | by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) | |
| 186 | ||
| 187 | lemma inj_onD: "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y" | |
| 188 | by (unfold inj_on_def, blast) | |
| 189 | ||
| 190 | lemma inj_on_iff: "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)" | |
| 191 | by (blast dest!: inj_onD) | |
| 192 | ||
| 193 | lemma comp_inj_on: | |
| 194 | "[| inj_on f A; inj_on g (f`A) |] ==> inj_on (g o f) A" | |
| 195 | by (simp add: comp_def inj_on_def) | |
| 196 | ||
| 15303 | 197 | lemma inj_on_imageI: "inj_on (g o f) A \<Longrightarrow> inj_on g (f ` A)" | 
| 198 | apply(simp add:inj_on_def image_def) | |
| 199 | apply blast | |
| 200 | done | |
| 201 | ||
| 15439 | 202 | lemma inj_on_image_iff: "\<lbrakk> ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y); | 
| 203 | inj_on f A \<rbrakk> \<Longrightarrow> inj_on g (f ` A) = inj_on g A" | |
| 204 | apply(unfold inj_on_def) | |
| 205 | apply blast | |
| 206 | done | |
| 207 | ||
| 13585 | 208 | lemma inj_on_contraD: "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)" | 
| 209 | by (unfold inj_on_def, blast) | |
| 12258 | 210 | |
| 13585 | 211 | lemma inj_singleton: "inj (%s. {s})"
 | 
| 212 | by (simp add: inj_on_def) | |
| 213 | ||
| 15111 | 214 | lemma inj_on_empty[iff]: "inj_on f {}"
 | 
| 215 | by(simp add: inj_on_def) | |
| 216 | ||
| 15303 | 217 | lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A" | 
| 13585 | 218 | by (unfold inj_on_def, blast) | 
| 219 | ||
| 15111 | 220 | lemma inj_on_Un: | 
| 221 | "inj_on f (A Un B) = | |
| 222 |   (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})"
 | |
| 223 | apply(unfold inj_on_def) | |
| 224 | apply (blast intro:sym) | |
| 225 | done | |
| 226 | ||
| 227 | lemma inj_on_insert[iff]: | |
| 228 |   "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))"
 | |
| 229 | apply(unfold inj_on_def) | |
| 230 | apply (blast intro:sym) | |
| 231 | done | |
| 232 | ||
| 233 | lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)" | |
| 234 | apply(unfold inj_on_def) | |
| 235 | apply (blast) | |
| 236 | done | |
| 237 | ||
| 13585 | 238 | lemma surjI: "(!! x. g(f x) = x) ==> surj g" | 
| 239 | apply (simp add: surj_def) | |
| 240 | apply (blast intro: sym) | |
| 241 | done | |
| 242 | ||
| 243 | lemma surj_range: "surj f ==> range f = UNIV" | |
| 244 | by (auto simp add: surj_def) | |
| 245 | ||
| 246 | lemma surjD: "surj f ==> EX x. y = f x" | |
| 247 | by (simp add: surj_def) | |
| 248 | ||
| 249 | lemma surjE: "surj f ==> (!!x. y = f x ==> C) ==> C" | |
| 250 | by (simp add: surj_def, blast) | |
| 251 | ||
| 252 | lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" | |
| 253 | apply (simp add: comp_def surj_def, clarify) | |
| 254 | apply (drule_tac x = y in spec, clarify) | |
| 255 | apply (drule_tac x = x in spec, blast) | |
| 256 | done | |
| 257 | ||
| 258 | lemma bijI: "[| inj f; surj f |] ==> bij f" | |
| 259 | by (simp add: bij_def) | |
| 260 | ||
| 261 | lemma bij_is_inj: "bij f ==> inj f" | |
| 262 | by (simp add: bij_def) | |
| 263 | ||
| 264 | lemma bij_is_surj: "bij f ==> surj f" | |
| 265 | by (simp add: bij_def) | |
| 266 | ||
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 267 | lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 268 | by (simp add: bij_betw_def) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 269 | |
| 32337 | 270 | lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g o f)" | 
| 271 | by(fastsimp intro: comp_inj_on comp_surj simp: bij_def surj_range) | |
| 272 | ||
| 31438 | 273 | lemma bij_betw_trans: | 
| 274 | "bij_betw f A B \<Longrightarrow> bij_betw g B C \<Longrightarrow> bij_betw (g o f) A C" | |
| 275 | by(auto simp add:bij_betw_def comp_inj_on) | |
| 276 | ||
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 277 | lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 278 | proof - | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 279 | have i: "inj_on f A" and s: "f ` A = B" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 280 | using assms by(auto simp:bij_betw_def) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 281 | let ?P = "%b a. a:A \<and> f a = b" let ?g = "%b. The (?P b)" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 282 |   { fix a b assume P: "?P b a"
 | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 283 | hence ex1: "\<exists>a. ?P b a" using s unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 284 | hence uex1: "\<exists>!a. ?P b a" by(blast dest:inj_onD[OF i]) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 285 | hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 286 | } note g = this | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 287 | have "inj_on ?g B" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 288 | proof(rule inj_onI) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 289 | fix x y assume "x:B" "y:B" "?g x = ?g y" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 290 | from s `x:B` obtain a1 where a1: "?P x a1" unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 291 | from s `y:B` obtain a2 where a2: "?P y a2" unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 292 | from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 293 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 294 | moreover have "?g ` B = A" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 295 | proof(auto simp:image_def) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 296 | fix b assume "b:B" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 297 | with s obtain a where P: "?P b a" unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 298 | thus "?g b \<in> A" using g[OF P] by auto | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 299 | next | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 300 | fix a assume "a:A" | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 301 | then obtain b where P: "?P b a" using s unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 302 | then have "b:B" using s unfolding image_def by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 303 | with g[OF P] show "\<exists>b\<in>B. a = ?g b" by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 304 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 305 | ultimately show ?thesis by(auto simp:bij_betw_def) | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 306 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 307 | |
| 13585 | 308 | lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" | 
| 309 | by (simp add: surj_range) | |
| 310 | ||
| 311 | lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" | |
| 312 | by (simp add: inj_on_def, blast) | |
| 313 | ||
| 314 | lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" | |
| 315 | apply (unfold surj_def) | |
| 316 | apply (blast intro: sym) | |
| 317 | done | |
| 318 | ||
| 319 | lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" | |
| 320 | by (unfold inj_on_def, blast) | |
| 321 | ||
| 322 | lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)" | |
| 323 | apply (unfold bij_def) | |
| 324 | apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD) | |
| 325 | done | |
| 326 | ||
| 31438 | 327 | lemma inj_on_Un_image_eq_iff: "inj_on f (A \<union> B) \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 328 | by(blast dest: inj_onD) | |
| 329 | ||
| 13585 | 330 | lemma inj_on_image_Int: | 
| 331 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A Int B) = f`A Int f`B" | |
| 332 | apply (simp add: inj_on_def, blast) | |
| 333 | done | |
| 334 | ||
| 335 | lemma inj_on_image_set_diff: | |
| 336 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A-B) = f`A - f`B" | |
| 337 | apply (simp add: inj_on_def, blast) | |
| 338 | done | |
| 339 | ||
| 340 | lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B" | |
| 341 | by (simp add: inj_on_def, blast) | |
| 342 | ||
| 343 | lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B" | |
| 344 | by (simp add: inj_on_def, blast) | |
| 345 | ||
| 346 | lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)" | |
| 347 | by (blast dest: injD) | |
| 348 | ||
| 349 | lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)" | |
| 350 | by (simp add: inj_on_def, blast) | |
| 351 | ||
| 352 | lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)" | |
| 353 | by (blast dest: injD) | |
| 354 | ||
| 355 | (*injectivity's required. Left-to-right inclusion holds even if A is empty*) | |
| 356 | lemma image_INT: | |
| 357 | "[| inj_on f C; ALL x:A. B x <= C; j:A |] | |
| 358 | ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 359 | apply (simp add: inj_on_def, blast) | |
| 360 | done | |
| 361 | ||
| 362 | (*Compare with image_INT: no use of inj_on, and if f is surjective then | |
| 363 | it doesn't matter whether A is empty*) | |
| 364 | lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 365 | apply (simp add: bij_def) | |
| 366 | apply (simp add: inj_on_def surj_def, blast) | |
| 367 | done | |
| 368 | ||
| 369 | lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" | |
| 370 | by (auto simp add: surj_def) | |
| 371 | ||
| 372 | lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" | |
| 373 | by (auto simp add: inj_on_def) | |
| 5852 | 374 | |
| 13585 | 375 | lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)" | 
| 376 | apply (simp add: bij_def) | |
| 377 | apply (rule equalityI) | |
| 378 | apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) | |
| 379 | done | |
| 380 | ||
| 35584 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 381 | lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A" | 
| 35580 | 382 | by (auto intro!: inj_onI) | 
| 13585 | 383 | |
| 35584 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 384 | lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \<Longrightarrow> inj_on f A" | 
| 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 385 | by (auto intro!: inj_onI dest: strict_mono_eq) | 
| 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 386 | |
| 13585 | 387 | subsection{*Function Updating*}
 | 
| 388 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 389 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 390 |   fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
 | 
| 26147 | 391 | "fun_upd f a b == % x. if x=a then b else f x" | 
| 392 | ||
| 393 | nonterminals | |
| 394 | updbinds updbind | |
| 395 | syntax | |
| 396 |   "_updbind" :: "['a, 'a] => updbind"             ("(2_ :=/ _)")
 | |
| 397 |   ""         :: "updbind => updbinds"             ("_")
 | |
| 398 |   "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _")
 | |
| 35115 | 399 |   "_Update"  :: "['a, updbinds] => 'a"            ("_/'((_)')" [1000, 0] 900)
 | 
| 26147 | 400 | |
| 401 | translations | |
| 35115 | 402 | "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs" | 
| 403 | "f(x:=y)" == "CONST fun_upd f x y" | |
| 26147 | 404 | |
| 405 | (* Hint: to define the sum of two functions (or maps), use sum_case. | |
| 406 | A nice infix syntax could be defined (in Datatype.thy or below) by | |
| 35115 | 407 | notation | 
| 408 | sum_case (infixr "'(+')"80) | |
| 26147 | 409 | *) | 
| 410 | ||
| 13585 | 411 | lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)" | 
| 412 | apply (simp add: fun_upd_def, safe) | |
| 413 | apply (erule subst) | |
| 414 | apply (rule_tac [2] ext, auto) | |
| 415 | done | |
| 416 | ||
| 417 | (* f x = y ==> f(x:=y) = f *) | |
| 418 | lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard] | |
| 419 | ||
| 420 | (* f(x := f x) = f *) | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 421 | lemmas fun_upd_triv = refl [THEN fun_upd_idem] | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 422 | declare fun_upd_triv [iff] | 
| 13585 | 423 | |
| 424 | lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)" | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 425 | by (simp add: fun_upd_def) | 
| 13585 | 426 | |
| 427 | (* fun_upd_apply supersedes these two, but they are useful | |
| 428 | if fun_upd_apply is intentionally removed from the simpset *) | |
| 429 | lemma fun_upd_same: "(f(x:=y)) x = y" | |
| 430 | by simp | |
| 431 | ||
| 432 | lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z" | |
| 433 | by simp | |
| 434 | ||
| 435 | lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)" | |
| 436 | by (simp add: expand_fun_eq) | |
| 437 | ||
| 438 | lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)" | |
| 439 | by (rule ext, auto) | |
| 440 | ||
| 15303 | 441 | lemma inj_on_fun_updI: "\<lbrakk> inj_on f A; y \<notin> f`A \<rbrakk> \<Longrightarrow> inj_on (f(x:=y)) A" | 
| 34209 | 442 | by (fastsimp simp:inj_on_def image_def) | 
| 15303 | 443 | |
| 15510 | 444 | lemma fun_upd_image: | 
| 445 |      "f(x:=y) ` A = (if x \<in> A then insert y (f ` (A-{x})) else f ` A)"
 | |
| 446 | by auto | |
| 447 | ||
| 31080 | 448 | lemma fun_upd_comp: "f \<circ> (g(x := y)) = (f \<circ> g)(x := f y)" | 
| 34209 | 449 | by (auto intro: ext) | 
| 31080 | 450 | |
| 26147 | 451 | |
| 452 | subsection {* @{text override_on} *}
 | |
| 453 | ||
| 454 | definition | |
| 455 |   override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
 | |
| 456 | where | |
| 457 | "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)" | |
| 13910 | 458 | |
| 15691 | 459 | lemma override_on_emptyset[simp]: "override_on f g {} = f"
 | 
| 460 | by(simp add:override_on_def) | |
| 13910 | 461 | |
| 15691 | 462 | lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a" | 
| 463 | by(simp add:override_on_def) | |
| 13910 | 464 | |
| 15691 | 465 | lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a" | 
| 466 | by(simp add:override_on_def) | |
| 13910 | 467 | |
| 26147 | 468 | |
| 469 | subsection {* @{text swap} *}
 | |
| 15510 | 470 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 471 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 472 |   swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 473 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 474 | "swap a b f = f (a := f b, b:= f a)" | 
| 15510 | 475 | |
| 34101 | 476 | lemma swap_self [simp]: "swap a a f = f" | 
| 15691 | 477 | by (simp add: swap_def) | 
| 15510 | 478 | |
| 479 | lemma swap_commute: "swap a b f = swap b a f" | |
| 480 | by (rule ext, simp add: fun_upd_def swap_def) | |
| 481 | ||
| 482 | lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" | |
| 483 | by (rule ext, simp add: fun_upd_def swap_def) | |
| 484 | ||
| 34145 | 485 | lemma swap_triple: | 
| 486 | assumes "a \<noteq> c" and "b \<noteq> c" | |
| 487 | shows "swap a b (swap b c (swap a b f)) = swap a c f" | |
| 488 | using assms by (simp add: expand_fun_eq swap_def) | |
| 489 | ||
| 34101 | 490 | lemma comp_swap: "f \<circ> swap a b g = swap a b (f \<circ> g)" | 
| 491 | by (rule ext, simp add: fun_upd_def swap_def) | |
| 492 | ||
| 15510 | 493 | lemma inj_on_imp_inj_on_swap: | 
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 494 | "[|inj_on f A; a \<in> A; b \<in> A|] ==> inj_on (swap a b f) A" | 
| 15510 | 495 | by (simp add: inj_on_def swap_def, blast) | 
| 496 | ||
| 497 | lemma inj_on_swap_iff [simp]: | |
| 498 | assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A = inj_on f A" | |
| 499 | proof | |
| 500 | assume "inj_on (swap a b f) A" | |
| 501 | with A have "inj_on (swap a b (swap a b f)) A" | |
| 17589 | 502 | by (iprover intro: inj_on_imp_inj_on_swap) | 
| 15510 | 503 | thus "inj_on f A" by simp | 
| 504 | next | |
| 505 | assume "inj_on f A" | |
| 34209 | 506 | with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) | 
| 15510 | 507 | qed | 
| 508 | ||
| 509 | lemma surj_imp_surj_swap: "surj f ==> surj (swap a b f)" | |
| 510 | apply (simp add: surj_def swap_def, clarify) | |
| 27125 | 511 | apply (case_tac "y = f b", blast) | 
| 512 | apply (case_tac "y = f a", auto) | |
| 15510 | 513 | done | 
| 514 | ||
| 515 | lemma surj_swap_iff [simp]: "surj (swap a b f) = surj f" | |
| 516 | proof | |
| 517 | assume "surj (swap a b f)" | |
| 518 | hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) | |
| 519 | thus "surj f" by simp | |
| 520 | next | |
| 521 | assume "surj f" | |
| 522 | thus "surj (swap a b f)" by (rule surj_imp_surj_swap) | |
| 523 | qed | |
| 524 | ||
| 525 | lemma bij_swap_iff: "bij (swap a b f) = bij f" | |
| 526 | by (simp add: bij_def) | |
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 527 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35584diff
changeset | 528 | hide_const (open) swap | 
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 529 | |
| 31949 | 530 | |
| 531 | subsection {* Inversion of injective functions *}
 | |
| 532 | ||
| 33057 | 533 | definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
 | 
| 534 | "the_inv_into A f == %x. THE y. y : A & f y = x" | |
| 32961 | 535 | |
| 33057 | 536 | lemma the_inv_into_f_f: | 
| 537 | "[| inj_on f A; x : A |] ==> the_inv_into A f (f x) = x" | |
| 538 | apply (simp add: the_inv_into_def inj_on_def) | |
| 34209 | 539 | apply blast | 
| 32961 | 540 | done | 
| 541 | ||
| 33057 | 542 | lemma f_the_inv_into_f: | 
| 543 | "inj_on f A ==> y : f`A ==> f (the_inv_into A f y) = y" | |
| 544 | apply (simp add: the_inv_into_def) | |
| 32961 | 545 | apply (rule the1I2) | 
| 546 | apply(blast dest: inj_onD) | |
| 547 | apply blast | |
| 548 | done | |
| 549 | ||
| 33057 | 550 | lemma the_inv_into_into: | 
| 551 | "[| inj_on f A; x : f ` A; A <= B |] ==> the_inv_into A f x : B" | |
| 552 | apply (simp add: the_inv_into_def) | |
| 32961 | 553 | apply (rule the1I2) | 
| 554 | apply(blast dest: inj_onD) | |
| 555 | apply blast | |
| 556 | done | |
| 557 | ||
| 33057 | 558 | lemma the_inv_into_onto[simp]: | 
| 559 | "inj_on f A ==> the_inv_into A f ` (f ` A) = A" | |
| 560 | by (fast intro:the_inv_into_into the_inv_into_f_f[symmetric]) | |
| 32961 | 561 | |
| 33057 | 562 | lemma the_inv_into_f_eq: | 
| 563 | "[| inj_on f A; f x = y; x : A |] ==> the_inv_into A f y = x" | |
| 32961 | 564 | apply (erule subst) | 
| 33057 | 565 | apply (erule the_inv_into_f_f, assumption) | 
| 32961 | 566 | done | 
| 567 | ||
| 33057 | 568 | lemma the_inv_into_comp: | 
| 32961 | 569 | "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==> | 
| 33057 | 570 | the_inv_into A (f o g) x = (the_inv_into A g o the_inv_into (g ` A) f) x" | 
| 571 | apply (rule the_inv_into_f_eq) | |
| 32961 | 572 | apply (fast intro: comp_inj_on) | 
| 33057 | 573 | apply (simp add: f_the_inv_into_f the_inv_into_into) | 
| 574 | apply (simp add: the_inv_into_into) | |
| 32961 | 575 | done | 
| 576 | ||
| 33057 | 577 | lemma inj_on_the_inv_into: | 
| 578 | "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)" | |
| 579 | by (auto intro: inj_onI simp: image_def the_inv_into_f_f) | |
| 32961 | 580 | |
| 33057 | 581 | lemma bij_betw_the_inv_into: | 
| 582 | "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A" | |
| 583 | by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) | |
| 32961 | 584 | |
| 32998 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 585 | abbreviation the_inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)" where
 | 
| 33057 | 586 | "the_inv f \<equiv> the_inv_into UNIV f" | 
| 32998 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 587 | |
| 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 588 | lemma the_inv_f_f: | 
| 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 589 | assumes "inj f" | 
| 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 590 | shows "the_inv f (f x) = x" using assms UNIV_I | 
| 33057 | 591 | by (rule the_inv_into_f_f) | 
| 32998 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 592 | |
| 31949 | 593 | |
| 22845 | 594 | subsection {* Proof tool setup *} 
 | 
| 595 | ||
| 596 | text {* simplifies terms of the form
 | |
| 597 | f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *} | |
| 598 | ||
| 24017 | 599 | simproc_setup fun_upd2 ("f(v := w, x := y)") = {* fn _ =>
 | 
| 22845 | 600 | let | 
| 601 | fun gen_fun_upd NONE T _ _ = NONE | |
| 24017 | 602 |     | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y)
 | 
| 22845 | 603 | fun dest_fun_T1 (Type (_, T :: Ts)) = T | 
| 604 |   fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) =
 | |
| 605 | let | |
| 606 |       fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) =
 | |
| 607 | if v aconv x then SOME g else gen_fun_upd (find g) T v w | |
| 608 | | find t = NONE | |
| 609 | in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end | |
| 24017 | 610 | |
| 611 | fun proc ss ct = | |
| 612 | let | |
| 613 | val ctxt = Simplifier.the_context ss | |
| 614 | val t = Thm.term_of ct | |
| 615 | in | |
| 616 | case find_double t of | |
| 617 | (T, NONE) => NONE | |
| 618 | | (T, SOME rhs) => | |
| 27330 | 619 | SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) | 
| 24017 | 620 | (fn _ => | 
| 621 | rtac eq_reflection 1 THEN | |
| 622 | rtac ext 1 THEN | |
| 623 |               simp_tac (Simplifier.inherit_context ss @{simpset}) 1))
 | |
| 624 | end | |
| 625 | in proc end | |
| 22845 | 626 | *} | 
| 627 | ||
| 628 | ||
| 21870 | 629 | subsection {* Code generator setup *}
 | 
| 630 | ||
| 25886 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 631 | types_code | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 632 |   "fun"  ("(_ ->/ _)")
 | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 633 | attach (term_of) {*
 | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 634 | fun term_of_fun_type _ aT _ bT _ = Free ("<function>", aT --> bT);
 | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 635 | *} | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 636 | attach (test) {*
 | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 637 | fun gen_fun_type aF aT bG bT i = | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 638 | let | 
| 32740 | 639 | val tab = Unsynchronized.ref []; | 
| 25886 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 640 |     fun mk_upd (x, (_, y)) t = Const ("Fun.fun_upd",
 | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 641 | (aT --> bT) --> aT --> bT --> aT --> bT) $ t $ aF x $ y () | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 642 | in | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 643 | (fn x => | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 644 | case AList.lookup op = (!tab) x of | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 645 | NONE => | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 646 | let val p as (y, _) = bG i | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 647 | in (tab := (x, p) :: !tab; y) end | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 648 | | SOME (y, _) => y, | 
| 28711 | 649 |      fn () => Basics.fold mk_upd (!tab) (Const ("HOL.undefined", aT --> bT)))
 | 
| 25886 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 650 | end; | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 651 | *} | 
| 
7753e0d81b7a
Added test data generator for function type (from Pure/codegen.ML).
 berghofe parents: 
24286diff
changeset | 652 | |
| 21870 | 653 | code_const "op \<circ>" | 
| 654 | (SML infixl 5 "o") | |
| 655 | (Haskell infixr 9 ".") | |
| 656 | ||
| 21906 | 657 | code_const "id" | 
| 658 | (Haskell "id") | |
| 659 | ||
| 2912 | 660 | end |