# HG changeset patch # User hoelzl # Date 1283417151 -7200 # Node ID a18e5946d63c9cdc4c77ee092600698e56c02162 # Parent 211e4f6aad6397983968f827bd0765c74fd9ba72 Permutation implies bij function diff -r 211e4f6aad63 -r a18e5946d63c src/HOL/Fun.thy --- a/src/HOL/Fun.thy Thu Sep 02 10:36:45 2010 +0200 +++ b/src/HOL/Fun.thy Thu Sep 02 10:45:51 2010 +0200 @@ -321,6 +321,11 @@ ultimately show ?thesis by(auto simp:bij_betw_def) qed +lemma bij_betw_combine: + assumes "bij_betw f A B" "bij_betw f C D" "B \ D = {}" + shows "bij_betw f (A \ C) (B \ D)" + using assms unfolding bij_betw_def inj_on_Un image_Un by auto + lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" by (simp add: surj_range) @@ -512,11 +517,11 @@ lemma inj_on_swap_iff [simp]: assumes A: "a \ A" "b \ A" shows "inj_on (swap a b f) A = inj_on f A" -proof +proof assume "inj_on (swap a b f) A" - with A have "inj_on (swap a b (swap a b f)) A" - by (iprover intro: inj_on_imp_inj_on_swap) - thus "inj_on f A" by simp + with A have "inj_on (swap a b (swap a b f)) A" + by (iprover intro: inj_on_imp_inj_on_swap) + thus "inj_on f A" by simp next assume "inj_on f A" with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) @@ -529,18 +534,41 @@ done lemma surj_swap_iff [simp]: "surj (swap a b f) = surj f" -proof +proof assume "surj (swap a b f)" - hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) - thus "surj f" by simp + hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) + thus "surj f" by simp next assume "surj f" - thus "surj (swap a b f)" by (rule surj_imp_surj_swap) + thus "surj (swap a b f)" by (rule surj_imp_surj_swap) qed lemma bij_swap_iff: "bij (swap a b f) = bij f" by (simp add: bij_def) +lemma bij_betw_swap: + assumes "bij_betw f A B" "x \ A" "y \ A" + shows "bij_betw (Fun.swap x y f) A B" +proof (unfold bij_betw_def, intro conjI) + show "inj_on (Fun.swap x y f) A" using assms + by (intro inj_on_imp_inj_on_swap) (auto simp: bij_betw_def) + show "Fun.swap x y f ` A = B" + proof safe + fix z assume "z \ A" + then show "Fun.swap x y f z \ B" + using assms unfolding bij_betw_def + by (auto simp: image_iff Fun.swap_def + intro!: bexI[of _ "if z = x then y else if z = y then x else z"]) + next + fix z assume "z \ B" + then obtain v where "v \ A" "z = f v" using assms unfolding bij_betw_def by auto + then show "z \ Fun.swap x y f ` A" unfolding image_iff + using assms + by (auto simp add: Fun.swap_def split: split_if_asm + intro!: bexI[of _ "if v = x then y else if v = y then x else v"]) + qed +qed + hide_const (open) swap diff -r 211e4f6aad63 -r a18e5946d63c src/HOL/Library/Permutation.thy --- a/src/HOL/Library/Permutation.thy Thu Sep 02 10:36:45 2010 +0200 +++ b/src/HOL/Library/Permutation.thy Thu Sep 02 10:45:51 2010 +0200 @@ -183,4 +183,55 @@ lemma perm_remdups_iff_eq_set: "remdups x <~~> remdups y = (set x = set y)" by (metis List.set_remdups perm_set_eq eq_set_perm_remdups) +lemma permutation_Ex_bij: + assumes "xs <~~> ys" + shows "\f. bij_betw f {.. (\ii Suc ` {.. Suc ` {..iif"] conjI allI impI) + show "bij_betw (g \ f) {.. f) i" + using trans(1,3)[THEN perm_length] perm by force + qed +qed + end