diff -r 2bc2a8599369 -r a03a63b81f44 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Mon Jan 14 18:33:53 2019 +0000 +++ b/src/HOL/Fun.thy Mon Jan 14 18:35:03 2019 +0000 @@ -599,15 +599,33 @@ context cancel_semigroup_add begin -lemma inj_add_left [simp]: \inj ((+) a)\ - by (rule injI) simp +lemma inj_on_add [simp]: + "inj_on ((+) a) A" + by (rule inj_onI) simp + +lemma inj_add_left: + \inj ((+) a)\ + by simp + +lemma inj_on_add' [simp]: + "inj_on (\b. b + a) A" + by (rule inj_onI) simp + +lemma bij_betw_add [simp]: + "bij_betw ((+) a) A B \ (+) a ` A = B" + by (simp add: bij_betw_def) end context ab_group_add begin -lemma inj_diff_right [simp]: \inj (\b. b - a)\ +lemma surj_plus [simp]: + "surj ((+) a)" + by (auto intro: range_eqI [of b "(+) a" "b - a" for b] simp add: algebra_simps) + +lemma inj_diff_right [simp]: + \inj (\b. b - a)\ proof - have \inj ((+) (- a))\ by (fact inj_add_left) @@ -616,6 +634,38 @@ finally show ?thesis . qed +lemma surj_diff_right [simp]: + "surj (\x. x - a)" + using surj_plus [of "- a"] by (simp cong: image_cong_simp) + +lemma translation_Compl: + "(+) a ` (- t) = - ((+) a ` t)" +proof (rule set_eqI) + fix b + show "b \ (+) a ` (- t) \ b \ - (+) a ` t" + by (auto simp: image_iff algebra_simps intro!: bexI [of _ "b - a"]) +qed + +lemma translation_subtract_Compl: + "(\x. x - a) ` (- t) = - ((\x. x - a) ` t)" + using translation_Compl [of "- a" t] by (simp cong: image_cong_simp) + +lemma translation_diff: + "(+) a ` (s - t) = ((+) a ` s) - ((+) a ` t)" + by auto + +lemma translation_subtract_diff: + "(\x. x - a) ` (s - t) = ((\x. x - a) ` s) - ((\x. x - a) ` t)" + using translation_diff [of "- a"] by (simp cong: image_cong_simp) + +lemma translation_Int: + "(+) a ` (s \ t) = ((+) a ` s) \ ((+) a ` t)" + by auto + +lemma translation_subtract_Int: + "(\x. x - a) ` (s \ t) = ((\x. x - a) ` s) \ ((\x. x - a) ` t)" + using translation_Int [of " -a"] by (simp cong: image_cong_simp) + end