clasohm@1475: (* Title: HOL/Fun.thy clasohm@1475: Author: Tobias Nipkow, Cambridge University Computer Laboratory clasohm@923: Copyright 1994 University of Cambridge huffman@18154: *) clasohm@923: huffman@18154: header {* Notions about functions *} clasohm@923: paulson@15510: theory Fun haftmann@32139: imports Complete_Lattice haftmann@32554: uses ("Tools/transfer.ML") nipkow@15131: begin nipkow@2912: haftmann@26147: text{*As a simplification rule, it replaces all function equalities by haftmann@26147: first-order equalities.*} haftmann@26147: lemma expand_fun_eq: "f = g \ (\x. f x = g x)" haftmann@26147: apply (rule iffI) haftmann@26147: apply (simp (no_asm_simp)) haftmann@26147: apply (rule ext) haftmann@26147: apply (simp (no_asm_simp)) haftmann@26147: done paulson@6171: haftmann@26147: lemma apply_inverse: haftmann@26357: "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" haftmann@26147: by auto oheimb@5305: nipkow@2912: haftmann@26147: subsection {* The Identity Function @{text id} *} nipkow@2912: haftmann@22744: definition haftmann@22744: id :: "'a \ 'a" haftmann@22744: where haftmann@22744: "id = (\x. x)" nipkow@13910: haftmann@26147: lemma id_apply [simp]: "id x = x" haftmann@26147: by (simp add: id_def) haftmann@26147: haftmann@26147: lemma image_ident [simp]: "(%x. x) ` Y = Y" haftmann@26147: by blast haftmann@26147: haftmann@26147: lemma image_id [simp]: "id ` Y = Y" haftmann@26147: by (simp add: id_def) haftmann@26147: haftmann@26147: lemma vimage_ident [simp]: "(%x. x) -` Y = Y" haftmann@26147: by blast haftmann@26147: haftmann@26147: lemma vimage_id [simp]: "id -` A = A" haftmann@26147: by (simp add: id_def) haftmann@26147: haftmann@26147: haftmann@26147: subsection {* The Composition Operator @{text "f \ g"} *} haftmann@26147: haftmann@22744: definition haftmann@22744: comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "o" 55) haftmann@22744: where haftmann@22744: "f o g = (\x. f (g x))" oheimb@11123: wenzelm@21210: notation (xsymbols) wenzelm@19656: comp (infixl "\" 55) wenzelm@19656: wenzelm@21210: notation (HTML output) wenzelm@19656: comp (infixl "\" 55) wenzelm@19656: paulson@13585: text{*compatibility*} paulson@13585: lemmas o_def = comp_def paulson@13585: paulson@13585: lemma o_apply [simp]: "(f o g) x = f (g x)" paulson@13585: by (simp add: comp_def) paulson@13585: paulson@13585: lemma o_assoc: "f o (g o h) = f o g o h" paulson@13585: by (simp add: comp_def) paulson@13585: paulson@13585: lemma id_o [simp]: "id o g = g" paulson@13585: by (simp add: comp_def) paulson@13585: paulson@13585: lemma o_id [simp]: "f o id = f" paulson@13585: by (simp add: comp_def) paulson@13585: paulson@13585: lemma image_compose: "(f o g) ` r = f`(g`r)" paulson@13585: by (simp add: comp_def, blast) paulson@13585: paulson@33044: lemma vimage_compose: "(g \ f) -` x = f -` (g -` x)" paulson@33044: by auto paulson@33044: paulson@13585: lemma UN_o: "UNION A (g o f) = UNION (f`A) g" paulson@13585: by (unfold comp_def, blast) paulson@13585: paulson@13585: haftmann@26588: subsection {* The Forward Composition Operator @{text fcomp} *} haftmann@26357: haftmann@26357: definition haftmann@26357: fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "o>" 60) haftmann@26357: where haftmann@26357: "f o> g = (\x. g (f x))" haftmann@26357: haftmann@26357: lemma fcomp_apply: "(f o> g) x = g (f x)" haftmann@26357: by (simp add: fcomp_def) haftmann@26357: haftmann@26357: lemma fcomp_assoc: "(f o> g) o> h = f o> (g o> h)" haftmann@26357: by (simp add: fcomp_def) haftmann@26357: haftmann@26357: lemma id_fcomp [simp]: "id o> g = g" haftmann@26357: by (simp add: fcomp_def) haftmann@26357: haftmann@26357: lemma fcomp_id [simp]: "f o> id = f" haftmann@26357: by (simp add: fcomp_def) haftmann@26357: haftmann@31202: code_const fcomp haftmann@31202: (Eval infixl 1 "#>") haftmann@31202: haftmann@26588: no_notation fcomp (infixl "o>" 60) haftmann@26588: haftmann@26357: haftmann@26147: subsection {* Injectivity and Surjectivity *} paulson@13585: haftmann@26147: constdefs haftmann@26147: inj_on :: "['a => 'b, 'a set] => bool" -- "injective" haftmann@26147: "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" paulson@13585: haftmann@26147: text{*A common special case: functions injective over the entire domain type.*} haftmann@26147: haftmann@26147: abbreviation haftmann@26147: "inj f == inj_on f UNIV" haftmann@26147: haftmann@26147: definition haftmann@26147: bij_betw :: "('a => 'b) => 'a set => 'b set => bool" where -- "bijective" haftmann@28562: [code del]: "bij_betw f A B \ inj_on f A & f ` A = B" haftmann@26147: haftmann@26147: constdefs haftmann@26147: surj :: "('a => 'b) => bool" (*surjective*) haftmann@26147: "surj f == ! y. ? x. y=f(x)" haftmann@26147: haftmann@26147: bij :: "('a => 'b) => bool" (*bijective*) haftmann@26147: "bij f == inj f & surj f" haftmann@26147: haftmann@26147: lemma injI: haftmann@26147: assumes "\x y. f x = f y \ x = y" haftmann@26147: shows "inj f" haftmann@26147: using assms unfolding inj_on_def by auto paulson@13585: haftmann@31771: text{*For Proofs in @{text "Tools/Datatype/datatype_rep_proofs"}*} paulson@13585: lemma datatype_injI: paulson@13585: "(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)" paulson@13585: by (simp add: inj_on_def) paulson@13585: berghofe@13637: theorem range_ex1_eq: "inj f \ b : range f = (EX! x. b = f x)" berghofe@13637: by (unfold inj_on_def, blast) berghofe@13637: paulson@13585: lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y" paulson@13585: by (simp add: inj_on_def) paulson@13585: nipkow@32988: lemma inj_on_eq_iff: "inj_on f A ==> x:A ==> y:A ==> (f(x) = f(y)) = (x=y)" paulson@13585: by (force simp add: inj_on_def) paulson@13585: nipkow@32988: lemma inj_eq: "inj f ==> (f(x) = f(y)) = (x=y)" nipkow@32988: by (simp add: inj_on_eq_iff) nipkow@32988: haftmann@26147: lemma inj_on_id[simp]: "inj_on id A" haftmann@26147: by (simp add: inj_on_def) paulson@13585: haftmann@26147: lemma inj_on_id2[simp]: "inj_on (%x. x) A" haftmann@26147: by (simp add: inj_on_def) haftmann@26147: haftmann@26147: lemma surj_id[simp]: "surj id" haftmann@26147: by (simp add: surj_def) haftmann@26147: haftmann@26147: lemma bij_id[simp]: "bij id" haftmann@26147: by (simp add: bij_def inj_on_id surj_id) paulson@13585: paulson@13585: lemma inj_onI: paulson@13585: "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" paulson@13585: by (simp add: inj_on_def) paulson@13585: paulson@13585: lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A" paulson@13585: by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) paulson@13585: paulson@13585: lemma inj_onD: "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y" paulson@13585: by (unfold inj_on_def, blast) paulson@13585: paulson@13585: lemma inj_on_iff: "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)" paulson@13585: by (blast dest!: inj_onD) paulson@13585: paulson@13585: lemma comp_inj_on: paulson@13585: "[| inj_on f A; inj_on g (f`A) |] ==> inj_on (g o f) A" paulson@13585: by (simp add: comp_def inj_on_def) paulson@13585: nipkow@15303: lemma inj_on_imageI: "inj_on (g o f) A \ inj_on g (f ` A)" nipkow@15303: apply(simp add:inj_on_def image_def) nipkow@15303: apply blast nipkow@15303: done nipkow@15303: nipkow@15439: lemma inj_on_image_iff: "\ ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y); nipkow@15439: inj_on f A \ \ inj_on g (f ` A) = inj_on g A" nipkow@15439: apply(unfold inj_on_def) nipkow@15439: apply blast nipkow@15439: done nipkow@15439: paulson@13585: lemma inj_on_contraD: "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)" paulson@13585: by (unfold inj_on_def, blast) paulson@13585: paulson@13585: lemma inj_singleton: "inj (%s. {s})" paulson@13585: by (simp add: inj_on_def) paulson@13585: nipkow@15111: lemma inj_on_empty[iff]: "inj_on f {}" nipkow@15111: by(simp add: inj_on_def) nipkow@15111: nipkow@15303: lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A" paulson@13585: by (unfold inj_on_def, blast) paulson@13585: nipkow@15111: lemma inj_on_Un: nipkow@15111: "inj_on f (A Un B) = nipkow@15111: (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})" nipkow@15111: apply(unfold inj_on_def) nipkow@15111: apply (blast intro:sym) nipkow@15111: done nipkow@15111: nipkow@15111: lemma inj_on_insert[iff]: nipkow@15111: "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))" nipkow@15111: apply(unfold inj_on_def) nipkow@15111: apply (blast intro:sym) nipkow@15111: done nipkow@15111: nipkow@15111: lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)" nipkow@15111: apply(unfold inj_on_def) nipkow@15111: apply (blast) nipkow@15111: done nipkow@15111: paulson@13585: lemma surjI: "(!! x. g(f x) = x) ==> surj g" paulson@13585: apply (simp add: surj_def) paulson@13585: apply (blast intro: sym) paulson@13585: done paulson@13585: paulson@13585: lemma surj_range: "surj f ==> range f = UNIV" paulson@13585: by (auto simp add: surj_def) paulson@13585: paulson@13585: lemma surjD: "surj f ==> EX x. y = f x" paulson@13585: by (simp add: surj_def) paulson@13585: paulson@13585: lemma surjE: "surj f ==> (!!x. y = f x ==> C) ==> C" paulson@13585: by (simp add: surj_def, blast) paulson@13585: paulson@13585: lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" paulson@13585: apply (simp add: comp_def surj_def, clarify) paulson@13585: apply (drule_tac x = y in spec, clarify) paulson@13585: apply (drule_tac x = x in spec, blast) paulson@13585: done paulson@13585: paulson@13585: lemma bijI: "[| inj f; surj f |] ==> bij f" paulson@13585: by (simp add: bij_def) paulson@13585: paulson@13585: lemma bij_is_inj: "bij f ==> inj f" paulson@13585: by (simp add: bij_def) paulson@13585: paulson@13585: lemma bij_is_surj: "bij f ==> surj f" paulson@13585: by (simp add: bij_def) paulson@13585: nipkow@26105: lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" nipkow@26105: by (simp add: bij_betw_def) nipkow@26105: nipkow@32337: lemma bij_comp: "bij f \ bij g \ bij (g o f)" nipkow@32337: by(fastsimp intro: comp_inj_on comp_surj simp: bij_def surj_range) nipkow@32337: nipkow@31424: lemma bij_betw_trans: nipkow@31424: "bij_betw f A B \ bij_betw g B C \ bij_betw (g o f) A C" nipkow@31424: by(auto simp add:bij_betw_def comp_inj_on) nipkow@31424: nipkow@26105: lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A" nipkow@26105: proof - nipkow@26105: have i: "inj_on f A" and s: "f ` A = B" nipkow@26105: using assms by(auto simp:bij_betw_def) nipkow@26105: let ?P = "%b a. a:A \ f a = b" let ?g = "%b. The (?P b)" nipkow@26105: { fix a b assume P: "?P b a" nipkow@26105: hence ex1: "\a. ?P b a" using s unfolding image_def by blast nipkow@26105: hence uex1: "\!a. ?P b a" by(blast dest:inj_onD[OF i]) nipkow@26105: hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp nipkow@26105: } note g = this nipkow@26105: have "inj_on ?g B" nipkow@26105: proof(rule inj_onI) nipkow@26105: fix x y assume "x:B" "y:B" "?g x = ?g y" nipkow@26105: from s `x:B` obtain a1 where a1: "?P x a1" unfolding image_def by blast nipkow@26105: from s `y:B` obtain a2 where a2: "?P y a2" unfolding image_def by blast nipkow@26105: from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp nipkow@26105: qed nipkow@26105: moreover have "?g ` B = A" nipkow@26105: proof(auto simp:image_def) nipkow@26105: fix b assume "b:B" nipkow@26105: with s obtain a where P: "?P b a" unfolding image_def by blast nipkow@26105: thus "?g b \ A" using g[OF P] by auto nipkow@26105: next nipkow@26105: fix a assume "a:A" nipkow@26105: then obtain b where P: "?P b a" using s unfolding image_def by blast nipkow@26105: then have "b:B" using s unfolding image_def by blast nipkow@26105: with g[OF P] show "\b\B. a = ?g b" by blast nipkow@26105: qed nipkow@26105: ultimately show ?thesis by(auto simp:bij_betw_def) nipkow@26105: qed nipkow@26105: paulson@13585: lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" paulson@13585: by (simp add: surj_range) paulson@13585: paulson@13585: lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" paulson@13585: by (simp add: inj_on_def, blast) paulson@13585: paulson@13585: lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" paulson@13585: apply (unfold surj_def) paulson@13585: apply (blast intro: sym) paulson@13585: done paulson@13585: paulson@13585: lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" paulson@13585: by (unfold inj_on_def, blast) paulson@13585: paulson@13585: lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)" paulson@13585: apply (unfold bij_def) paulson@13585: apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD) paulson@13585: done paulson@13585: nipkow@31424: lemma inj_on_Un_image_eq_iff: "inj_on f (A \ B) \ f ` A = f ` B \ A = B" nipkow@31424: by(blast dest: inj_onD) nipkow@31424: paulson@13585: lemma inj_on_image_Int: paulson@13585: "[| inj_on f C; A<=C; B<=C |] ==> f`(A Int B) = f`A Int f`B" paulson@13585: apply (simp add: inj_on_def, blast) paulson@13585: done paulson@13585: paulson@13585: lemma inj_on_image_set_diff: paulson@13585: "[| inj_on f C; A<=C; B<=C |] ==> f`(A-B) = f`A - f`B" paulson@13585: apply (simp add: inj_on_def, blast) paulson@13585: done paulson@13585: paulson@13585: lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B" paulson@13585: by (simp add: inj_on_def, blast) paulson@13585: paulson@13585: lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B" paulson@13585: by (simp add: inj_on_def, blast) paulson@13585: paulson@13585: lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)" paulson@13585: by (blast dest: injD) paulson@13585: paulson@13585: lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)" paulson@13585: by (simp add: inj_on_def, blast) paulson@13585: paulson@13585: lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)" paulson@13585: by (blast dest: injD) paulson@13585: paulson@13585: (*injectivity's required. Left-to-right inclusion holds even if A is empty*) paulson@13585: lemma image_INT: paulson@13585: "[| inj_on f C; ALL x:A. B x <= C; j:A |] paulson@13585: ==> f ` (INTER A B) = (INT x:A. f ` B x)" paulson@13585: apply (simp add: inj_on_def, blast) paulson@13585: done paulson@13585: paulson@13585: (*Compare with image_INT: no use of inj_on, and if f is surjective then paulson@13585: it doesn't matter whether A is empty*) paulson@13585: lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)" paulson@13585: apply (simp add: bij_def) paulson@13585: apply (simp add: inj_on_def surj_def, blast) paulson@13585: done paulson@13585: paulson@13585: lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" paulson@13585: by (auto simp add: surj_def) paulson@13585: paulson@13585: lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" paulson@13585: by (auto simp add: inj_on_def) paulson@13585: paulson@13585: lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)" paulson@13585: apply (simp add: bij_def) paulson@13585: apply (rule equalityI) paulson@13585: apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) paulson@13585: done paulson@13585: paulson@13585: paulson@13585: subsection{*Function Updating*} paulson@13585: haftmann@26147: constdefs haftmann@26147: fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" haftmann@26147: "fun_upd f a b == % x. if x=a then b else f x" haftmann@26147: haftmann@26147: nonterminals haftmann@26147: updbinds updbind haftmann@26147: syntax haftmann@26147: "_updbind" :: "['a, 'a] => updbind" ("(2_ :=/ _)") haftmann@26147: "" :: "updbind => updbinds" ("_") haftmann@26147: "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _") haftmann@26147: "_Update" :: "['a, updbinds] => 'a" ("_/'((_)')" [1000,0] 900) haftmann@26147: haftmann@26147: translations haftmann@26147: "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs" haftmann@26147: "f(x:=y)" == "fun_upd f x y" haftmann@26147: haftmann@26147: (* Hint: to define the sum of two functions (or maps), use sum_case. haftmann@26147: A nice infix syntax could be defined (in Datatype.thy or below) by haftmann@26147: consts haftmann@26147: fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80) haftmann@26147: translations haftmann@26147: "fun_sum" == sum_case haftmann@26147: *) haftmann@26147: paulson@13585: lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)" paulson@13585: apply (simp add: fun_upd_def, safe) paulson@13585: apply (erule subst) paulson@13585: apply (rule_tac [2] ext, auto) paulson@13585: done paulson@13585: paulson@13585: (* f x = y ==> f(x:=y) = f *) paulson@13585: lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard] paulson@13585: paulson@13585: (* f(x := f x) = f *) paulson@17084: lemmas fun_upd_triv = refl [THEN fun_upd_idem] paulson@17084: declare fun_upd_triv [iff] paulson@13585: paulson@13585: lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)" paulson@17084: by (simp add: fun_upd_def) paulson@13585: paulson@13585: (* fun_upd_apply supersedes these two, but they are useful paulson@13585: if fun_upd_apply is intentionally removed from the simpset *) paulson@13585: lemma fun_upd_same: "(f(x:=y)) x = y" paulson@13585: by simp paulson@13585: paulson@13585: lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z" paulson@13585: by simp paulson@13585: paulson@13585: lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)" paulson@13585: by (simp add: expand_fun_eq) paulson@13585: paulson@13585: lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)" paulson@13585: by (rule ext, auto) paulson@13585: nipkow@15303: lemma inj_on_fun_updI: "\ inj_on f A; y \ f`A \ \ inj_on (f(x:=y)) A" nipkow@15303: by(fastsimp simp:inj_on_def image_def) nipkow@15303: paulson@15510: lemma fun_upd_image: paulson@15510: "f(x:=y) ` A = (if x \ A then insert y (f ` (A-{x})) else f ` A)" paulson@15510: by auto paulson@15510: nipkow@31080: lemma fun_upd_comp: "f \ (g(x := y)) = (f \ g)(x := f y)" nipkow@31080: by(auto intro: ext) nipkow@31080: haftmann@26147: haftmann@26147: subsection {* @{text override_on} *} haftmann@26147: haftmann@26147: definition haftmann@26147: override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" haftmann@26147: where haftmann@26147: "override_on f g A = (\a. if a \ A then g a else f a)" nipkow@13910: nipkow@15691: lemma override_on_emptyset[simp]: "override_on f g {} = f" nipkow@15691: by(simp add:override_on_def) nipkow@13910: nipkow@15691: lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a" nipkow@15691: by(simp add:override_on_def) nipkow@13910: nipkow@15691: lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a" nipkow@15691: by(simp add:override_on_def) nipkow@13910: haftmann@26147: haftmann@26147: subsection {* @{text swap} *} paulson@15510: haftmann@22744: definition haftmann@22744: swap :: "'a \ 'a \ ('a \ 'b) \ ('a \ 'b)" haftmann@22744: where haftmann@22744: "swap a b f = f (a := f b, b:= f a)" paulson@15510: paulson@15510: lemma swap_self: "swap a a f = f" nipkow@15691: by (simp add: swap_def) paulson@15510: paulson@15510: lemma swap_commute: "swap a b f = swap b a f" paulson@15510: by (rule ext, simp add: fun_upd_def swap_def) paulson@15510: paulson@15510: lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" paulson@15510: by (rule ext, simp add: fun_upd_def swap_def) paulson@15510: paulson@15510: lemma inj_on_imp_inj_on_swap: haftmann@22744: "[|inj_on f A; a \ A; b \ A|] ==> inj_on (swap a b f) A" paulson@15510: by (simp add: inj_on_def swap_def, blast) paulson@15510: paulson@15510: lemma inj_on_swap_iff [simp]: paulson@15510: assumes A: "a \ A" "b \ A" shows "inj_on (swap a b f) A = inj_on f A" paulson@15510: proof paulson@15510: assume "inj_on (swap a b f) A" paulson@15510: with A have "inj_on (swap a b (swap a b f)) A" nipkow@17589: by (iprover intro: inj_on_imp_inj_on_swap) paulson@15510: thus "inj_on f A" by simp paulson@15510: next paulson@15510: assume "inj_on f A" nipkow@27165: with A show "inj_on (swap a b f) A" by(iprover intro: inj_on_imp_inj_on_swap) paulson@15510: qed paulson@15510: paulson@15510: lemma surj_imp_surj_swap: "surj f ==> surj (swap a b f)" paulson@15510: apply (simp add: surj_def swap_def, clarify) wenzelm@27125: apply (case_tac "y = f b", blast) wenzelm@27125: apply (case_tac "y = f a", auto) paulson@15510: done paulson@15510: paulson@15510: lemma surj_swap_iff [simp]: "surj (swap a b f) = surj f" paulson@15510: proof paulson@15510: assume "surj (swap a b f)" paulson@15510: hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) paulson@15510: thus "surj f" by simp paulson@15510: next paulson@15510: assume "surj f" paulson@15510: thus "surj (swap a b f)" by (rule surj_imp_surj_swap) paulson@15510: qed paulson@15510: paulson@15510: lemma bij_swap_iff: "bij (swap a b f) = bij f" paulson@15510: by (simp add: bij_def) haftmann@21547: nipkow@27188: hide (open) const swap haftmann@21547: haftmann@31949: haftmann@31949: subsection {* Inversion of injective functions *} haftmann@31949: nipkow@33057: definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where nipkow@33057: "the_inv_into A f == %x. THE y. y : A & f y = x" nipkow@32957: nipkow@33057: lemma the_inv_into_f_f: nipkow@33057: "[| inj_on f A; x : A |] ==> the_inv_into A f (f x) = x" nipkow@33057: apply (simp add: the_inv_into_def inj_on_def) nipkow@32957: apply (blast intro: the_equality) nipkow@32957: done nipkow@32957: nipkow@33057: lemma f_the_inv_into_f: nipkow@33057: "inj_on f A ==> y : f`A ==> f (the_inv_into A f y) = y" nipkow@33057: apply (simp add: the_inv_into_def) nipkow@32957: apply (rule the1I2) nipkow@32957: apply(blast dest: inj_onD) nipkow@32957: apply blast nipkow@32957: done nipkow@32957: nipkow@33057: lemma the_inv_into_into: nipkow@33057: "[| inj_on f A; x : f ` A; A <= B |] ==> the_inv_into A f x : B" nipkow@33057: apply (simp add: the_inv_into_def) nipkow@32957: apply (rule the1I2) nipkow@32957: apply(blast dest: inj_onD) nipkow@32957: apply blast nipkow@32957: done nipkow@32957: nipkow@33057: lemma the_inv_into_onto[simp]: nipkow@33057: "inj_on f A ==> the_inv_into A f ` (f ` A) = A" nipkow@33057: by (fast intro:the_inv_into_into the_inv_into_f_f[symmetric]) nipkow@32957: nipkow@33057: lemma the_inv_into_f_eq: nipkow@33057: "[| inj_on f A; f x = y; x : A |] ==> the_inv_into A f y = x" nipkow@32957: apply (erule subst) nipkow@33057: apply (erule the_inv_into_f_f, assumption) nipkow@32957: done nipkow@32957: nipkow@33057: lemma the_inv_into_comp: nipkow@32957: "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==> nipkow@33057: the_inv_into A (f o g) x = (the_inv_into A g o the_inv_into (g ` A) f) x" nipkow@33057: apply (rule the_inv_into_f_eq) nipkow@32957: apply (fast intro: comp_inj_on) nipkow@33057: apply (simp add: f_the_inv_into_f the_inv_into_into) nipkow@33057: apply (simp add: the_inv_into_into) nipkow@32957: done nipkow@32957: nipkow@33057: lemma inj_on_the_inv_into: nipkow@33057: "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" nipkow@33057: by (auto intro: inj_onI simp: image_def the_inv_into_f_f) nipkow@32957: nipkow@33057: lemma bij_betw_the_inv_into: nipkow@33057: "bij_betw f A B \ bij_betw (the_inv_into A f) B A" nipkow@33057: by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) nipkow@32957: berghofe@32998: abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" where nipkow@33057: "the_inv f \ the_inv_into UNIV f" berghofe@32998: berghofe@32998: lemma the_inv_f_f: berghofe@32998: assumes "inj f" berghofe@32998: shows "the_inv f (f x) = x" using assms UNIV_I nipkow@33057: by (rule the_inv_into_f_f) berghofe@32998: haftmann@31949: haftmann@22845: subsection {* Proof tool setup *} haftmann@22845: haftmann@22845: text {* simplifies terms of the form haftmann@22845: f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *} haftmann@22845: wenzelm@24017: simproc_setup fun_upd2 ("f(v := w, x := y)") = {* fn _ => haftmann@22845: let haftmann@22845: fun gen_fun_upd NONE T _ _ = NONE wenzelm@24017: | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y) haftmann@22845: fun dest_fun_T1 (Type (_, T :: Ts)) = T haftmann@22845: fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) = haftmann@22845: let haftmann@22845: fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) = haftmann@22845: if v aconv x then SOME g else gen_fun_upd (find g) T v w haftmann@22845: | find t = NONE haftmann@22845: in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end wenzelm@24017: wenzelm@24017: fun proc ss ct = wenzelm@24017: let wenzelm@24017: val ctxt = Simplifier.the_context ss wenzelm@24017: val t = Thm.term_of ct wenzelm@24017: in wenzelm@24017: case find_double t of wenzelm@24017: (T, NONE) => NONE wenzelm@24017: | (T, SOME rhs) => wenzelm@27330: SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) wenzelm@24017: (fn _ => wenzelm@24017: rtac eq_reflection 1 THEN wenzelm@24017: rtac ext 1 THEN wenzelm@24017: simp_tac (Simplifier.inherit_context ss @{simpset}) 1)) wenzelm@24017: end wenzelm@24017: in proc end haftmann@22845: *} haftmann@22845: haftmann@22845: haftmann@32554: subsection {* Generic transfer procedure *} haftmann@32554: haftmann@32554: definition TransferMorphism:: "('b \ 'a) \ 'b set \ bool" haftmann@32554: where "TransferMorphism a B \ True" haftmann@32554: haftmann@32554: use "Tools/transfer.ML" haftmann@32554: haftmann@32554: setup Transfer.setup haftmann@32554: haftmann@32554: haftmann@21870: subsection {* Code generator setup *} haftmann@21870: berghofe@25886: types_code berghofe@25886: "fun" ("(_ ->/ _)") berghofe@25886: attach (term_of) {* berghofe@25886: fun term_of_fun_type _ aT _ bT _ = Free ("", aT --> bT); berghofe@25886: *} berghofe@25886: attach (test) {* berghofe@25886: fun gen_fun_type aF aT bG bT i = berghofe@25886: let wenzelm@32740: val tab = Unsynchronized.ref []; berghofe@25886: fun mk_upd (x, (_, y)) t = Const ("Fun.fun_upd", berghofe@25886: (aT --> bT) --> aT --> bT --> aT --> bT) $ t $ aF x $ y () berghofe@25886: in berghofe@25886: (fn x => berghofe@25886: case AList.lookup op = (!tab) x of berghofe@25886: NONE => berghofe@25886: let val p as (y, _) = bG i berghofe@25886: in (tab := (x, p) :: !tab; y) end berghofe@25886: | SOME (y, _) => y, berghofe@28711: fn () => Basics.fold mk_upd (!tab) (Const ("HOL.undefined", aT --> bT))) berghofe@25886: end; berghofe@25886: *} berghofe@25886: haftmann@21870: code_const "op \" haftmann@21870: (SML infixl 5 "o") haftmann@21870: (Haskell infixr 9 ".") haftmann@21870: haftmann@21906: code_const "id" haftmann@21906: (Haskell "id") haftmann@21906: nipkow@2912: end