# HG changeset patch # User haftmann # Date 1614543187 0 # Node ID ff24fe85ee5743a10157ec565d8f7dd4d55a1fa8 # Parent fd32f08f4fb5b020672a1d82929d9a5b36698125 lemma diffusion diff -r fd32f08f4fb5 -r ff24fe85ee57 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Sun Feb 28 20:13:07 2021 +0000 +++ b/src/HOL/Fun.thy Sun Feb 28 20:13:07 2021 +0000 @@ -891,6 +891,18 @@ lemma bij_swap_iff [simp]: "bij (swap a b f) \ bij f" by simp +subsection \Transpositions\ + +lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id" + by (rule ext) (auto simp add: Fun.swap_def) + +lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" + by (simp add: Fun.swap_def) + +lemma bij_swap_compose_bij: + \bij (Fun.swap a b id \ p)\ if \bij p\ + using that by (rule bij_comp) simp + hide_const (open) swap diff -r fd32f08f4fb5 -r ff24fe85ee57 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Sun Feb 28 20:13:07 2021 +0000 +++ b/src/HOL/Hilbert_Choice.thy Sun Feb 28 20:13:07 2021 +0000 @@ -590,6 +590,15 @@ shows "inv f = g" using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff) +lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" + by (rule inv_unique_comp) simp_all + +lemma bij_swap_comp: + assumes "bij p" + shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" + using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] + by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) + lemma subset_image_inj: "S \ f ` T \ (\U. U \ T \ inj_on f U \ S = f ` U)" proof safe diff -r fd32f08f4fb5 -r ff24fe85ee57 src/HOL/Library/Permutations.thy --- a/src/HOL/Library/Permutations.thy Sun Feb 28 20:13:07 2021 +0000 +++ b/src/HOL/Library/Permutations.thy Sun Feb 28 20:13:07 2021 +0000 @@ -8,30 +8,7 @@ imports Multiset Disjoint_Sets begin -subsection \Transpositions\ - -lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id" - by (rule ext) (auto simp add: Fun.swap_def) - -lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" - by (rule inv_unique_comp) simp_all - -lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" - by (simp add: Fun.swap_def) - -lemma bij_swap_comp: - assumes "bij p" - shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" - using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] - by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) - -lemma bij_swap_compose_bij: - assumes "bij p" - shows "bij (Fun.swap a b id \ p)" - by (simp only: bij_swap_comp[OF \bij p\] bij_swap_iff \bij p\) - - -subsection \Basic consequences of the definition\ +subsection \Basic definition and consequences\ definition permutes (infixr "permutes" 41) where "(p permutes S) \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)" @@ -426,21 +403,6 @@ qed -subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ - -lemma swap_id_common:" a \ c \ b \ c \ - Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - -lemma swap_id_common': "a \ b \ a \ c \ - Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - -lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ - Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - - subsection \Permutations as transposition sequences\ inductive swapidseq :: "nat \ ('a \ 'a) \ bool" @@ -538,6 +500,22 @@ using permutation_def swapidseq_inverse by blast + +subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ + +lemma swap_id_common:" a \ c \ b \ c \ + Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + +lemma swap_id_common': "a \ b \ a \ c \ + Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + +lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ + Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + + subsection \The identity map only has even transposition sequences\ lemma symmetry_lemma: @@ -744,13 +722,6 @@ subsection \A more abstract characterization of permutations\ -lemma bij_iff: "bij f \ (\x. \!y. f y = x)" - unfolding bij_def inj_def surj_def - apply auto - apply metis - apply metis - done - lemma permutation_bijective: assumes "permutation p" shows "bij p" @@ -791,7 +762,7 @@ lemma permutation_lemma: assumes "finite S" and "bij p" - and "\x. x\ S \ p x = x" + and "\x. x \ S \ p x = x" shows "permutation p" using assms proof (induct S arbitrary: p rule: finite_induct) @@ -872,10 +843,11 @@ subsection \Hence a sort of induction principle composing by swaps\ -lemma permutes_induct: "finite S \ P id \ - (\a b p. a \ S \ b \ S \ P p \ P p \ permutation p \ P (Fun.swap a b id \ p)) \ - (\p. p permutes S \ P p)" -proof (induct S rule: finite_induct) +lemma permutes_induct [consumes 2, case_names id swap]: + \P p\ if \p permutes S\ \finite S\ + and id: \P id\ + and swap: \\a b p. a \ S \ b \ S \ permutation p \ P p \ P (Fun.swap a b id \ p)\ +using \finite S\ \p permutes S\ id swap proof (induction S arbitrary: p) case empty then show ?case by auto next @@ -884,25 +856,26 @@ let ?q = "Fun.swap x (p x) id \ ?r" have qp: "?q = p" by (simp add: o_assoc) - from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" + from permutes_insert_lemma[OF \p permutes insert x F\] insert have Pr: "P ?r" by blast - from permutes_in_image[OF insert.prems(3), of x] + from permutes_in_image[OF \p permutes insert x F\, of x] have pxF: "p x \ insert x F" by simp have xF: "x \ insert x F" by simp have rp: "permutation ?r" unfolding permutation_permutes - using insert.hyps(1) permutes_insert_lemma[OF insert.prems(3)] + using insert.hyps(1) permutes_insert_lemma[OF \p permutes insert x F\] by blast - from insert.prems(2)[OF xF pxF Pr Pr rp] qp show ?case + from insert.prems(3)[OF xF pxF rp Pr] qp show ?case by (simp only:) qed subsection \Sign of a permutation as a real number\ -definition "sign p = (if evenperm p then (1::int) else -1)" +definition sign :: \('a \ 'a) \ int\ \ \TODO: prefer less generic name\ + where \sign p = (if evenperm p then (1::int) else -1)\ lemma sign_nz: "sign p \ 0" by (simp add: sign_def) @@ -1470,7 +1443,7 @@ lemma inverse_permutation_of_list_unique': "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x" - by (induct xs) (force simp: inverse_permutation_of_list.simps)+ + by (induct xs) (force simp: inverse_permutation_of_list.simps(2))+ lemma inverse_permutation_of_list_unique: "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x"