| author | wenzelm | 
| Fri, 08 Dec 2023 15:37:46 +0100 | |
| changeset 79207 | f991d3003ec8 | 
| parent 78258 | 71366be2c647 | 
| child 79582 | 7822b55b26ce | 
| permissions | -rw-r--r-- | 
| 1475 | 1 | (* Title: HOL/Fun.thy | 
| 2 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 3 | Author: Andrei Popescu, TU Muenchen | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 4 | Copyright 1994, 2012 | 
| 18154 | 5 | *) | 
| 923 | 6 | |
| 60758 | 7 | section \<open>Notions about functions\<close> | 
| 923 | 8 | |
| 15510 | 9 | theory Fun | 
| 63575 | 10 | imports Set | 
| 69913 | 11 | keywords "functor" :: thy_goal_defn | 
| 15131 | 12 | begin | 
| 2912 | 13 | |
| 63322 | 14 | lemma apply_inverse: "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u" | 
| 26147 | 15 | by auto | 
| 2912 | 16 | |
| 63322 | 17 | text \<open>Uniqueness, so NOT the axiom of choice.\<close> | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 18 | lemma uniq_choice: "\<forall>x. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x. Q x (f x)" | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 19 | by (force intro: theI') | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 20 | |
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 21 | lemma b_uniq_choice: "\<forall>x\<in>S. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x\<in>S. Q x (f x)" | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 22 | by (force intro: theI') | 
| 12258 | 23 | |
| 63400 | 24 | |
| 61799 | 25 | subsection \<open>The Identity Function \<open>id\<close>\<close> | 
| 6171 | 26 | |
| 63322 | 27 | definition id :: "'a \<Rightarrow> 'a" | 
| 28 | where "id = (\<lambda>x. x)" | |
| 13910 | 29 | |
| 26147 | 30 | lemma id_apply [simp]: "id x = x" | 
| 31 | by (simp add: id_def) | |
| 32 | ||
| 47579 | 33 | lemma image_id [simp]: "image id = id" | 
| 34 | by (simp add: id_def fun_eq_iff) | |
| 26147 | 35 | |
| 47579 | 36 | lemma vimage_id [simp]: "vimage id = id" | 
| 37 | by (simp add: id_def fun_eq_iff) | |
| 26147 | 38 | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 39 | lemma eq_id_iff: "(\<forall>x. f x = x) \<longleftrightarrow> f = id" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 40 | by auto | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 41 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 42 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 43 | constant id \<rightharpoonup> (Haskell) "id" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 44 | |
| 26147 | 45 | |
| 61799 | 46 | subsection \<open>The Composition Operator \<open>f \<circ> g\<close>\<close> | 
| 26147 | 47 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 48 | definition comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c"  (infixl "\<circ>" 55)
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 49 | where "f \<circ> g = (\<lambda>x. f (g x))" | 
| 11123 | 50 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 51 | notation (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 52 | comp (infixl "o" 55) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 53 | |
| 63322 | 54 | lemma comp_apply [simp]: "(f \<circ> g) x = f (g x)" | 
| 49739 | 55 | by (simp add: comp_def) | 
| 13585 | 56 | |
| 63322 | 57 | lemma comp_assoc: "(f \<circ> g) \<circ> h = f \<circ> (g \<circ> h)" | 
| 49739 | 58 | by (simp add: fun_eq_iff) | 
| 13585 | 59 | |
| 63322 | 60 | lemma id_comp [simp]: "id \<circ> g = g" | 
| 49739 | 61 | by (simp add: fun_eq_iff) | 
| 13585 | 62 | |
| 63322 | 63 | lemma comp_id [simp]: "f \<circ> id = f" | 
| 49739 | 64 | by (simp add: fun_eq_iff) | 
| 65 | ||
| 63575 | 66 | lemma comp_eq_dest: "a \<circ> b = c \<circ> d \<Longrightarrow> a (b v) = c (d v)" | 
| 49739 | 67 | by (simp add: fun_eq_iff) | 
| 34150 | 68 | |
| 63575 | 69 | lemma comp_eq_elim: "a \<circ> b = c \<circ> d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R" | 
| 61204 | 70 | by (simp add: fun_eq_iff) | 
| 34150 | 71 | |
| 63322 | 72 | lemma comp_eq_dest_lhs: "a \<circ> b = c \<Longrightarrow> a (b v) = c v" | 
| 55066 | 73 | by clarsimp | 
| 74 | ||
| 63322 | 75 | lemma comp_eq_id_dest: "a \<circ> b = id \<circ> c \<Longrightarrow> a (b v) = c v" | 
| 76 | by clarsimp | |
| 77 | ||
| 78 | lemma image_comp: "f ` (g ` r) = (f \<circ> g) ` r" | |
| 33044 | 79 | by auto | 
| 80 | ||
| 63322 | 81 | lemma vimage_comp: "f -` (g -` x) = (g \<circ> f) -` x" | 
| 49739 | 82 | by auto | 
| 83 | ||
| 63322 | 84 | lemma image_eq_imp_comp: "f ` A = g ` B \<Longrightarrow> (h \<circ> f) ` A = (h \<circ> g) ` B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 85 | by (auto simp: comp_def elim!: equalityE) | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 86 | |
| 67399 | 87 | lemma image_bind: "f ` (Set.bind A g) = Set.bind A ((`) f \<circ> g)" | 
| 63322 | 88 | by (auto simp add: Set.bind_def) | 
| 59512 | 89 | |
| 90 | lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \<circ> f)" | |
| 63322 | 91 | by (auto simp add: Set.bind_def) | 
| 59512 | 92 | |
| 63322 | 93 | lemma (in group_add) minus_comp_minus [simp]: "uminus \<circ> uminus = id" | 
| 60929 | 94 | by (simp add: fun_eq_iff) | 
| 95 | ||
| 74123 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74101diff
changeset | 96 | lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \<circ> uminus = id" | 
| 60929 | 97 | by (simp add: fun_eq_iff) | 
| 98 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 99 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 100 | constant comp \<rightharpoonup> (SML) infixl 5 "o" and (Haskell) infixr 9 "." | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 101 | |
| 13585 | 102 | |
| 61799 | 103 | subsection \<open>The Forward Composition Operator \<open>fcomp\<close>\<close> | 
| 26357 | 104 | |
| 63575 | 105 | definition fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c"  (infixl "\<circ>>" 60)
 | 
| 63322 | 106 | where "f \<circ>> g = (\<lambda>x. g (f x))" | 
| 26357 | 107 | |
| 37751 | 108 | lemma fcomp_apply [simp]: "(f \<circ>> g) x = g (f x)" | 
| 26357 | 109 | by (simp add: fcomp_def) | 
| 110 | ||
| 37751 | 111 | lemma fcomp_assoc: "(f \<circ>> g) \<circ>> h = f \<circ>> (g \<circ>> h)" | 
| 26357 | 112 | by (simp add: fcomp_def) | 
| 113 | ||
| 37751 | 114 | lemma id_fcomp [simp]: "id \<circ>> g = g" | 
| 26357 | 115 | by (simp add: fcomp_def) | 
| 116 | ||
| 37751 | 117 | lemma fcomp_id [simp]: "f \<circ>> id = f" | 
| 26357 | 118 | by (simp add: fcomp_def) | 
| 119 | ||
| 63322 | 120 | lemma fcomp_comp: "fcomp f g = comp g f" | 
| 61699 
a81dc5c4d6a9
New theorems mostly from Peter Gammie
 paulson <lp15@cam.ac.uk> parents: 
61630diff
changeset | 121 | by (simp add: ext) | 
| 
a81dc5c4d6a9
New theorems mostly from Peter Gammie
 paulson <lp15@cam.ac.uk> parents: 
61630diff
changeset | 122 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 123 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 124 | constant fcomp \<rightharpoonup> (Eval) infixl 1 "#>" | 
| 31202 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
31080diff
changeset | 125 | |
| 37751 | 126 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26357diff
changeset | 127 | |
| 26357 | 128 | |
| 60758 | 129 | subsection \<open>Mapping functions\<close> | 
| 40602 | 130 | |
| 63322 | 131 | definition map_fun :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'c \<Rightarrow> 'd"
 | 
| 132 | where "map_fun f g h = g \<circ> h \<circ> f" | |
| 40602 | 133 | |
| 63322 | 134 | lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" | 
| 40602 | 135 | by (simp add: map_fun_def) | 
| 136 | ||
| 137 | ||
| 60758 | 138 | subsection \<open>Injectivity and Bijectivity\<close> | 
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 139 | |
| 63322 | 140 | definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool"  \<comment> \<open>injective\<close>
 | 
| 141 | where "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)" | |
| 26147 | 142 | |
| 63322 | 143 | definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool"  \<comment> \<open>bijective\<close>
 | 
| 144 | where "bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B" | |
| 26147 | 145 | |
| 63575 | 146 | text \<open> | 
| 147 | A common special case: functions injective, surjective or bijective over | |
| 148 | the entire domain type. | |
| 149 | \<close> | |
| 26147 | 150 | |
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 151 | abbreviation inj :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 152 | where "inj f \<equiv> inj_on f UNIV" | 
| 26147 | 153 | |
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 154 | abbreviation surj :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 63322 | 155 | where "surj f \<equiv> range f = UNIV" | 
| 13585 | 156 | |
| 67226 | 157 | translations \<comment> \<open>The negated case:\<close> | 
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 158 | "\<not> CONST surj f" \<leftharpoondown> "CONST range f \<noteq> CONST UNIV" | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 159 | |
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 160 | abbreviation bij :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 161 | where "bij f \<equiv> bij_betw f UNIV UNIV" | 
| 26147 | 162 | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 163 | lemma inj_def: "inj f \<longleftrightarrow> (\<forall>x y. f x = f y \<longrightarrow> x = y)" | 
| 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 164 | unfolding inj_on_def by blast | 
| 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 165 | |
| 63322 | 166 | lemma injI: "(\<And>x y. f x = f y \<Longrightarrow> x = y) \<Longrightarrow> inj f" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 167 | unfolding inj_def by blast | 
| 13585 | 168 | |
| 63322 | 169 | theorem range_ex1_eq: "inj f \<Longrightarrow> b \<in> range f \<longleftrightarrow> (\<exists>!x. b = f x)" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 170 | unfolding inj_def by blast | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 171 | |
| 63322 | 172 | lemma injD: "inj f \<Longrightarrow> f x = f y \<Longrightarrow> x = y" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 173 | by (simp add: inj_def) | 
| 63322 | 174 | |
| 175 | lemma inj_on_eq_iff: "inj_on f A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<longleftrightarrow> x = y" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 176 | by (auto simp: inj_on_def) | 
| 63322 | 177 | |
| 64965 | 178 | lemma inj_on_cong: "(\<And>a. a \<in> A \<Longrightarrow> f a = g a) \<Longrightarrow> inj_on f A \<longleftrightarrow> inj_on g A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 179 | by (auto simp: inj_on_def) | 
| 63322 | 180 | |
| 76281 
457f1cba78fb
renamed lemma inj_on_strict_subset to image_strict_mono for symmetry with image_mono and to distinguish from inj_on_subset
 desharna parents: 
76264diff
changeset | 181 | lemma image_strict_mono: "inj_on f B \<Longrightarrow> A \<subset> B \<Longrightarrow> f ` A \<subset> f ` B" | 
| 63322 | 182 | unfolding inj_on_def by blast | 
| 183 | ||
| 69700 
7a92cbec7030
new material about summations and powers, along with some tweaks
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 184 | lemma inj_compose: "inj f \<Longrightarrow> inj g \<Longrightarrow> inj (f \<circ> g)" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 185 | by (simp add: inj_def) | 
| 38620 | 186 | |
| 187 | lemma inj_fun: "inj f \<Longrightarrow> inj (\<lambda>x y. f x)" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 188 | by (simp add: inj_def fun_eq_iff) | 
| 38620 | 189 | |
| 63322 | 190 | lemma inj_eq: "inj f \<Longrightarrow> f x = f y \<longleftrightarrow> x = y" | 
| 191 | by (simp add: inj_on_eq_iff) | |
| 32988 | 192 | |
| 71827 | 193 | lemma inj_on_iff_Uniq: "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<exists>\<^sub>\<le>\<^sub>1y. y\<in>A \<and> f x = f y)" | 
| 194 | by (auto simp: Uniq_def inj_on_def) | |
| 195 | ||
| 26147 | 196 | lemma inj_on_id[simp]: "inj_on id A" | 
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 197 | by (simp add: inj_on_def) | 
| 13585 | 198 | |
| 63322 | 199 | lemma inj_on_id2[simp]: "inj_on (\<lambda>x. x) A" | 
| 200 | by (simp add: inj_on_def) | |
| 26147 | 201 | |
| 46586 | 202 | lemma inj_on_Int: "inj_on f A \<or> inj_on f B \<Longrightarrow> inj_on f (A \<inter> B)" | 
| 63322 | 203 | unfolding inj_on_def by blast | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 204 | |
| 40702 | 205 | lemma surj_id: "surj id" | 
| 63322 | 206 | by simp | 
| 26147 | 207 | |
| 39101 
606432dd1896
Revert bij_betw changes to simp set (Problem in afp/Ordinals_and_Cardinals)
 hoelzl parents: 
39076diff
changeset | 208 | lemma bij_id[simp]: "bij id" | 
| 63322 | 209 | by (simp add: bij_betw_def) | 
| 13585 | 210 | |
| 76252 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 211 | lemma bij_uminus: "bij (uminus :: 'a \<Rightarrow> 'a::group_add)" | 
| 63322 | 212 | unfolding bij_betw_def inj_on_def | 
| 213 | by (force intro: minus_minus [symmetric]) | |
| 63072 | 214 | |
| 72125 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 215 | lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B" | 
| 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 216 | unfolding bij_betw_def by auto | 
| 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 217 | |
| 63322 | 218 | lemma inj_onI [intro?]: "(\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> x = y) \<Longrightarrow> inj_on f A" | 
| 219 | by (simp add: inj_on_def) | |
| 13585 | 220 | |
| 78258 
71366be2c647
The sym_diff operator (symmetric difference)
 paulson <lp15@cam.ac.uk> parents: 
78099diff
changeset | 221 | text \<open>For those frequent proofs by contradiction\<close> | 
| 
71366be2c647
The sym_diff operator (symmetric difference)
 paulson <lp15@cam.ac.uk> parents: 
78099diff
changeset | 222 | lemma inj_onCI: "(\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> x \<noteq> y \<Longrightarrow> False) \<Longrightarrow> inj_on f A" | 
| 
71366be2c647
The sym_diff operator (symmetric difference)
 paulson <lp15@cam.ac.uk> parents: 
78099diff
changeset | 223 | by (force simp: inj_on_def) | 
| 
71366be2c647
The sym_diff operator (symmetric difference)
 paulson <lp15@cam.ac.uk> parents: 
78099diff
changeset | 224 | |
| 63322 | 225 | lemma inj_on_inverseI: "(\<And>x. x \<in> A \<Longrightarrow> g (f x) = x) \<Longrightarrow> inj_on f A" | 
| 64965 | 226 | by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) | 
| 13585 | 227 | |
| 63322 | 228 | lemma inj_onD: "inj_on f A \<Longrightarrow> f x = f y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x = y" | 
| 229 | unfolding inj_on_def by blast | |
| 13585 | 230 | |
| 63365 | 231 | lemma inj_on_subset: | 
| 232 | assumes "inj_on f A" | |
| 63575 | 233 | and "B \<subseteq> A" | 
| 63365 | 234 | shows "inj_on f B" | 
| 235 | proof (rule inj_onI) | |
| 236 | fix a b | |
| 237 | assume "a \<in> B" and "b \<in> B" | |
| 238 | with assms have "a \<in> A" and "b \<in> A" | |
| 239 | by auto | |
| 240 | moreover assume "f a = f b" | |
| 64965 | 241 | ultimately show "a = b" | 
| 242 | using assms by (auto dest: inj_onD) | |
| 63365 | 243 | qed | 
| 244 | ||
| 63322 | 245 | lemma comp_inj_on: "inj_on f A \<Longrightarrow> inj_on g (f ` A) \<Longrightarrow> inj_on (g \<circ> f) A" | 
| 246 | by (simp add: comp_def inj_on_def) | |
| 247 | ||
| 248 | lemma inj_on_imageI: "inj_on (g \<circ> f) A \<Longrightarrow> inj_on g (f ` A)" | |
| 63072 | 249 | by (auto simp add: inj_on_def) | 
| 15303 | 250 | |
| 63322 | 251 | lemma inj_on_image_iff: | 
| 64965 | 252 | "\<forall>x\<in>A. \<forall>y\<in>A. g (f x) = g (f y) \<longleftrightarrow> g x = g y \<Longrightarrow> inj_on f A \<Longrightarrow> inj_on g (f ` A) \<longleftrightarrow> inj_on g A" | 
| 63322 | 253 | unfolding inj_on_def by blast | 
| 15439 | 254 | |
| 63322 | 255 | lemma inj_on_contraD: "inj_on f A \<Longrightarrow> x \<noteq> y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x \<noteq> f y" | 
| 256 | unfolding inj_on_def by blast | |
| 12258 | 257 | |
| 63072 | 258 | lemma inj_singleton [simp]: "inj_on (\<lambda>x. {x}) A"
 | 
| 259 | by (simp add: inj_on_def) | |
| 13585 | 260 | |
| 15111 | 261 | lemma inj_on_empty[iff]: "inj_on f {}"
 | 
| 63322 | 262 | by (simp add: inj_on_def) | 
| 13585 | 263 | |
| 63322 | 264 | lemma subset_inj_on: "inj_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> inj_on f A" | 
| 265 | unfolding inj_on_def by blast | |
| 266 | ||
| 267 | lemma inj_on_Un: "inj_on f (A \<union> B) \<longleftrightarrow> inj_on f A \<and> inj_on f B \<and> f ` (A - B) \<inter> f ` (B - A) = {}"
 | |
| 268 | unfolding inj_on_def by (blast intro: sym) | |
| 15111 | 269 | |
| 63322 | 270 | lemma inj_on_insert [iff]: "inj_on f (insert a A) \<longleftrightarrow> inj_on f A \<and> f a \<notin> f ` (A - {a})"
 | 
| 271 | unfolding inj_on_def by (blast intro: sym) | |
| 272 | ||
| 273 | lemma inj_on_diff: "inj_on f A \<Longrightarrow> inj_on f (A - B)" | |
| 274 | unfolding inj_on_def by blast | |
| 15111 | 275 | |
| 63322 | 276 | lemma comp_inj_on_iff: "inj_on f A \<Longrightarrow> inj_on f' (f ` A) \<longleftrightarrow> inj_on (f' \<circ> f) A" | 
| 64965 | 277 | by (auto simp: comp_inj_on inj_on_def) | 
| 15111 | 278 | |
| 63322 | 279 | lemma inj_on_imageI2: "inj_on (f' \<circ> f) A \<Longrightarrow> inj_on f A" | 
| 64965 | 280 | by (auto simp: comp_inj_on inj_on_def) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 281 | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 282 | lemma inj_img_insertE: | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 283 | assumes "inj_on f A" | 
| 63322 | 284 | assumes "x \<notin> B" | 
| 285 | and "insert x B = f ` A" | |
| 286 | obtains x' A' where "x' \<notin> A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 287 | proof - | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 288 | from assms have "x \<in> f ` A" by auto | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 289 | then obtain x' where *: "x' \<in> A" "x = f x'" by auto | 
| 63322 | 290 |   then have A: "A = insert x' (A - {x'})" by auto
 | 
| 291 |   with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD)
 | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 292 |   have "x' \<notin> A - {x'}" by simp
 | 
| 63322 | 293 | from this A \<open>x = f x'\<close> B show ?thesis .. | 
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 294 | qed | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 295 | |
| 71404 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 296 | lemma linorder_inj_onI: | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 297 | fixes A :: "'a::order set" | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 298 | assumes ne: "\<And>x y. \<lbrakk>x < y; x\<in>A; y\<in>A\<rbrakk> \<Longrightarrow> f x \<noteq> f y" and lin: "\<And>x y. \<lbrakk>x\<in>A; y\<in>A\<rbrakk> \<Longrightarrow> x\<le>y \<or> y\<le>x" | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 299 | shows "inj_on f A" | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 300 | proof (rule inj_onI) | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 301 | fix x y | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 302 | assume eq: "f x = f y" and "x\<in>A" "y\<in>A" | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 303 | then show "x = y" | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 304 | using lin [of x y] ne by (force simp: dual_order.order_iff_strict) | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 305 | qed | 
| 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 306 | |
| 76722 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 307 | lemma linorder_inj_onI': | 
| 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 308 | fixes A :: "'a :: linorder set" | 
| 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 309 | assumes "\<And>i j. i \<in> A \<Longrightarrow> j \<in> A \<Longrightarrow> i < j \<Longrightarrow> f i \<noteq> f j" | 
| 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 310 | shows "inj_on f A" | 
| 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 311 | by (intro linorder_inj_onI) (auto simp add: assms) | 
| 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 312 | |
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 313 | lemma linorder_injI: | 
| 64965 | 314 | assumes "\<And>x y::'a::linorder. x < y \<Longrightarrow> f x \<noteq> f y" | 
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 315 | shows "inj f" | 
| 71404 
f2b783abfbe7
A few lemmas connected with orderings
 paulson <lp15@cam.ac.uk> parents: 
69913diff
changeset | 316 | \<comment> \<open>Courtesy of Stephan Merz\<close> | 
| 76722 
b1d57dd345e1
First round of moving material from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
76281diff
changeset | 317 | using assms by (simp add: linorder_inj_onI') | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 318 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 319 | lemma inj_on_image_Pow: "inj_on f A \<Longrightarrow>inj_on (image f) (Pow A)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 320 | unfolding Pow_def inj_on_def by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 321 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 322 | lemma bij_betw_image_Pow: "bij_betw f A B \<Longrightarrow> bij_betw (image f) (Pow A) (Pow B)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 323 | by (auto simp add: bij_betw_def inj_on_image_Pow image_Pow_surj) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 324 | |
| 40702 | 325 | lemma surj_def: "surj f \<longleftrightarrow> (\<forall>y. \<exists>x. y = f x)" | 
| 326 | by auto | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 327 | |
| 63322 | 328 | lemma surjI: | 
| 64965 | 329 | assumes "\<And>x. g (f x) = x" | 
| 63322 | 330 | shows "surj g" | 
| 64965 | 331 | using assms [symmetric] by auto | 
| 13585 | 332 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 333 | lemma surjD: "surj f \<Longrightarrow> \<exists>x. y = f x" | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 334 | by (simp add: surj_def) | 
| 13585 | 335 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 336 | lemma surjE: "surj f \<Longrightarrow> (\<And>x. y = f x \<Longrightarrow> C) \<Longrightarrow> C" | 
| 63575 | 337 | by (simp add: surj_def) blast | 
| 13585 | 338 | |
| 63322 | 339 | lemma comp_surj: "surj f \<Longrightarrow> surj g \<Longrightarrow> surj (g \<circ> f)" | 
| 69768 | 340 | using image_comp [of g f UNIV] by simp | 
| 13585 | 341 | |
| 63322 | 342 | lemma bij_betw_imageI: "inj_on f A \<Longrightarrow> f ` A = B \<Longrightarrow> bij_betw f A B" | 
| 343 | unfolding bij_betw_def by clarify | |
| 57282 | 344 | |
| 345 | lemma bij_betw_imp_surj_on: "bij_betw f A B \<Longrightarrow> f ` A = B" | |
| 346 | unfolding bij_betw_def by clarify | |
| 347 | ||
| 39074 | 348 | lemma bij_betw_imp_surj: "bij_betw f A UNIV \<Longrightarrow> surj f" | 
| 40702 | 349 | unfolding bij_betw_def by auto | 
| 39074 | 350 | |
| 63322 | 351 | lemma bij_betw_empty1: "bij_betw f {} A \<Longrightarrow> A = {}"
 | 
| 352 | unfolding bij_betw_def by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 353 | |
| 63322 | 354 | lemma bij_betw_empty2: "bij_betw f A {} \<Longrightarrow> A = {}"
 | 
| 355 | unfolding bij_betw_def by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 356 | |
| 63322 | 357 | lemma inj_on_imp_bij_betw: "inj_on f A \<Longrightarrow> bij_betw f A (f ` A)" | 
| 358 | unfolding bij_betw_def by simp | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 359 | |
| 77138 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 360 | lemma bij_betw_DiffI: | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 361 | assumes "bij_betw f A B" "bij_betw f C D" "C \<subseteq> A" "D \<subseteq> B" | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 362 | shows "bij_betw f (A - C) (B - D)" | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 363 | using assms unfolding bij_betw_def inj_on_def by auto | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 364 | |
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 365 | lemma bij_betw_singleton_iff [simp]: "bij_betw f {x} {y} \<longleftrightarrow> f x = y"
 | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 366 | by (auto simp: bij_betw_def) | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 367 | |
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 368 | lemma bij_betw_singletonI [intro]: "f x = y \<Longrightarrow> bij_betw f {x} {y}"
 | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 369 | by auto | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76722diff
changeset | 370 | |
| 71464 | 371 | lemma bij_betw_apply: "\<lbrakk>bij_betw f A B; a \<in> A\<rbrakk> \<Longrightarrow> f a \<in> B" | 
| 372 | unfolding bij_betw_def by auto | |
| 373 | ||
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 374 | lemma bij_def: "bij f \<longleftrightarrow> inj f \<and> surj f" | 
| 64965 | 375 | by (rule bij_betw_def) | 
| 39074 | 376 | |
| 63322 | 377 | lemma bijI: "inj f \<Longrightarrow> surj f \<Longrightarrow> bij f" | 
| 64965 | 378 | by (rule bij_betw_imageI) | 
| 13585 | 379 | |
| 63322 | 380 | lemma bij_is_inj: "bij f \<Longrightarrow> inj f" | 
| 381 | by (simp add: bij_def) | |
| 13585 | 382 | |
| 63322 | 383 | lemma bij_is_surj: "bij f \<Longrightarrow> surj f" | 
| 384 | by (simp add: bij_def) | |
| 13585 | 385 | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 386 | lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" | 
| 63322 | 387 | by (simp add: bij_betw_def) | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 388 | |
| 63322 | 389 | lemma bij_betw_trans: "bij_betw f A B \<Longrightarrow> bij_betw g B C \<Longrightarrow> bij_betw (g \<circ> f) A C" | 
| 390 | by (auto simp add:bij_betw_def comp_inj_on) | |
| 31438 | 391 | |
| 63322 | 392 | lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g \<circ> f)" | 
| 40702 | 393 | by (rule bij_betw_trans) | 
| 394 | ||
| 63322 | 395 | lemma bij_betw_comp_iff: "bij_betw f A A' \<Longrightarrow> bij_betw f' A' A'' \<longleftrightarrow> bij_betw (f' \<circ> f) A A''" | 
| 396 | by (auto simp add: bij_betw_def inj_on_def) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 397 | |
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 398 | lemma bij_betw_comp_iff2: | 
| 63322 | 399 | assumes bij: "bij_betw f' A' A''" | 
| 400 | and img: "f ` A \<le> A'" | |
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 401 | shows "bij_betw f A A' \<longleftrightarrow> bij_betw (f' \<circ> f) A A''" (is "?L \<longleftrightarrow> ?R") | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 402 | proof | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 403 | assume "?L" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 404 | then show "?R" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 405 | using assms by (auto simp add: bij_betw_comp_iff) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 406 | next | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 407 | assume *: "?R" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 408 | have "inj_on (f' \<circ> f) A \<Longrightarrow> inj_on f A" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 409 | using inj_on_imageI2 by blast | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 410 | moreover have "A' \<subseteq> f ` A" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 411 | proof | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 412 | fix a' | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 413 | assume **: "a' \<in> A'" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 414 | with bij have "f' a' \<in> A''" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 415 | unfolding bij_betw_def by auto | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 416 | with * obtain a where 1: "a \<in> A \<and> f' (f a) = f' a'" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 417 | unfolding bij_betw_def by force | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 418 | with img have "f a \<in> A'" by auto | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 419 | with bij ** 1 have "f a = a'" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 420 | unfolding bij_betw_def inj_on_def by auto | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 421 | with 1 show "a' \<in> f ` A" by auto | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 422 | qed | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 423 | ultimately show "?L" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 424 | using img * by (auto simp add: bij_betw_def) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 425 | qed | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 426 | |
| 63322 | 427 | lemma bij_betw_inv: | 
| 428 | assumes "bij_betw f A B" | |
| 429 | shows "\<exists>g. bij_betw g B A" | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 430 | proof - | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 431 | have i: "inj_on f A" and s: "f ` A = B" | 
| 63322 | 432 | using assms by (auto simp: bij_betw_def) | 
| 433 | let ?P = "\<lambda>b a. a \<in> A \<and> f a = b" | |
| 434 | let ?g = "\<lambda>b. The (?P b)" | |
| 435 | have g: "?g b = a" if P: "?P b a" for a b | |
| 436 | proof - | |
| 63575 | 437 | from that s have ex1: "\<exists>a. ?P b a" by blast | 
| 63322 | 438 | then have uex1: "\<exists>!a. ?P b a" by (blast dest:inj_onD[OF i]) | 
| 63575 | 439 | then show ?thesis | 
| 440 | using the1_equality[OF uex1, OF P] P by simp | |
| 63322 | 441 | qed | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 442 | have "inj_on ?g B" | 
| 63322 | 443 | proof (rule inj_onI) | 
| 444 | fix x y | |
| 445 | assume "x \<in> B" "y \<in> B" "?g x = ?g y" | |
| 446 | from s \<open>x \<in> B\<close> obtain a1 where a1: "?P x a1" by blast | |
| 447 | from s \<open>y \<in> B\<close> obtain a2 where a2: "?P y a2" by blast | |
| 448 | from g [OF a1] a1 g [OF a2] a2 \<open>?g x = ?g y\<close> show "x = y" by simp | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 449 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 450 | moreover have "?g ` B = A" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 451 | proof safe | 
| 63322 | 452 | fix b | 
| 453 | assume "b \<in> B" | |
| 56077 | 454 | with s obtain a where P: "?P b a" by blast | 
| 63575 | 455 | with g[OF P] show "?g b \<in> A" by auto | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 456 | next | 
| 63322 | 457 | fix a | 
| 458 | assume "a \<in> A" | |
| 63575 | 459 | with s obtain b where P: "?P b a" by blast | 
| 460 | with s have "b \<in> B" by blast | |
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 461 | with g[OF P] have "\<exists>b\<in>B. a = ?g b" by blast | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 462 | then show "a \<in> ?g ` B" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 463 | by auto | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 464 | qed | 
| 63575 | 465 | ultimately show ?thesis | 
| 466 | by (auto simp: bij_betw_def) | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 467 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 468 | |
| 63588 | 469 | lemma bij_betw_cong: "(\<And>a. a \<in> A \<Longrightarrow> f a = g a) \<Longrightarrow> bij_betw f A A' = bij_betw g A A'" | 
| 63591 | 470 | unfolding bij_betw_def inj_on_def by safe force+ (* somewhat slow *) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 471 | |
| 63322 | 472 | lemma bij_betw_id[intro, simp]: "bij_betw id A A" | 
| 473 | unfolding bij_betw_def id_def by auto | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 474 | |
| 63322 | 475 | lemma bij_betw_id_iff: "bij_betw id A B \<longleftrightarrow> A = B" | 
| 476 | by (auto simp add: bij_betw_def) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 477 | |
| 39075 | 478 | lemma bij_betw_combine: | 
| 63400 | 479 |   "bij_betw f A B \<Longrightarrow> bij_betw f C D \<Longrightarrow> B \<inter> D = {} \<Longrightarrow> bij_betw f (A \<union> C) (B \<union> D)"
 | 
| 480 | unfolding bij_betw_def inj_on_Un image_Un by auto | |
| 39075 | 481 | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 482 | lemma bij_betw_subset: "bij_betw f A A' \<Longrightarrow> B \<subseteq> A \<Longrightarrow> f ` B = B' \<Longrightarrow> bij_betw f B B'" | 
| 63322 | 483 | by (auto simp add: bij_betw_def inj_on_def) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 484 | |
| 75624 | 485 | lemma bij_betw_ball: "bij_betw f A B \<Longrightarrow> (\<forall>b \<in> B. phi b) = (\<forall>a \<in> A. phi (f a))" | 
| 486 | unfolding bij_betw_def inj_on_def by blast | |
| 487 | ||
| 58195 | 488 | lemma bij_pointE: | 
| 489 | assumes "bij f" | |
| 490 | obtains x where "y = f x" and "\<And>x'. y = f x' \<Longrightarrow> x' = x" | |
| 491 | proof - | |
| 492 | from assms have "inj f" by (rule bij_is_inj) | |
| 493 | moreover from assms have "surj f" by (rule bij_is_surj) | |
| 494 | then have "y \<in> range f" by simp | |
| 495 | ultimately have "\<exists>!x. y = f x" by (simp add: range_ex1_eq) | |
| 496 | with that show thesis by blast | |
| 497 | qed | |
| 498 | ||
| 73326 | 499 | lemma bij_iff: \<^marker>\<open>contributor \<open>Amine Chaieb\<close>\<close> | 
| 500 | \<open>bij f \<longleftrightarrow> (\<forall>x. \<exists>!y. f y = x)\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 501 | proof | |
| 502 | assume ?P | |
| 503 | then have \<open>inj f\<close> \<open>surj f\<close> | |
| 504 | by (simp_all add: bij_def) | |
| 505 | show ?Q | |
| 506 | proof | |
| 507 | fix y | |
| 508 | from \<open>surj f\<close> obtain x where \<open>y = f x\<close> | |
| 509 | by (auto simp add: surj_def) | |
| 510 | with \<open>inj f\<close> show \<open>\<exists>!x. f x = y\<close> | |
| 511 | by (auto simp add: inj_def) | |
| 512 | qed | |
| 513 | next | |
| 514 | assume ?Q | |
| 515 | then have \<open>inj f\<close> | |
| 516 | by (auto simp add: inj_def) | |
| 517 | moreover have \<open>\<exists>x. y = f x\<close> for y | |
| 518 | proof - | |
| 519 | from \<open>?Q\<close> obtain x where \<open>f x = y\<close> | |
| 520 | by blast | |
| 521 | then have \<open>y = f x\<close> | |
| 522 | by simp | |
| 523 | then show ?thesis .. | |
| 524 | qed | |
| 525 | then have \<open>surj f\<close> | |
| 526 | by (auto simp add: surj_def) | |
| 527 | ultimately show ?P | |
| 528 | by (rule bijI) | |
| 529 | qed | |
| 530 | ||
| 73466 | 531 | lemma bij_betw_partition: | 
| 532 | \<open>bij_betw f A B\<close> | |
| 533 |   if \<open>bij_betw f (A \<union> C) (B \<union> D)\<close> \<open>bij_betw f C D\<close> \<open>A \<inter> C = {}\<close> \<open>B \<inter> D = {}\<close>
 | |
| 534 | proof - | |
| 535 | from that have \<open>inj_on f (A \<union> C)\<close> \<open>inj_on f C\<close> \<open>f ` (A \<union> C) = B \<union> D\<close> \<open>f ` C = D\<close> | |
| 536 | by (simp_all add: bij_betw_def) | |
| 537 |   then have \<open>inj_on f A\<close> and \<open>f ` (A - C) \<inter> f ` (C - A) = {}\<close>
 | |
| 538 | by (simp_all add: inj_on_Un) | |
| 539 |   with \<open>A \<inter> C = {}\<close> have \<open>f ` A \<inter> f ` C = {}\<close>
 | |
| 540 | by auto | |
| 541 |   with \<open>f ` (A \<union> C) = B \<union> D\<close> \<open>f ` C = D\<close>  \<open>B \<inter> D = {}\<close>
 | |
| 542 | have \<open>f ` A = B\<close> | |
| 543 | by blast | |
| 544 | with \<open>inj_on f A\<close> show ?thesis | |
| 545 | by (simp add: bij_betw_def) | |
| 546 | qed | |
| 547 | ||
| 63322 | 548 | lemma surj_image_vimage_eq: "surj f \<Longrightarrow> f ` (f -` A) = A" | 
| 549 | by simp | |
| 13585 | 550 | |
| 42903 | 551 | lemma surj_vimage_empty: | 
| 63322 | 552 | assumes "surj f" | 
| 553 |   shows "f -` A = {} \<longleftrightarrow> A = {}"
 | |
| 554 | using surj_image_vimage_eq [OF \<open>surj f\<close>, of A] | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44860diff
changeset | 555 | by (intro iffI) fastforce+ | 
| 42903 | 556 | |
| 63322 | 557 | lemma inj_vimage_image_eq: "inj f \<Longrightarrow> f -` (f ` A) = A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 558 | unfolding inj_def by blast | 
| 13585 | 559 | |
| 63322 | 560 | lemma vimage_subsetD: "surj f \<Longrightarrow> f -` B \<subseteq> A \<Longrightarrow> B \<subseteq> f ` A" | 
| 561 | by (blast intro: sym) | |
| 13585 | 562 | |
| 63322 | 563 | lemma vimage_subsetI: "inj f \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> f -` B \<subseteq> A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 564 | unfolding inj_def by blast | 
| 13585 | 565 | |
| 63322 | 566 | lemma vimage_subset_eq: "bij f \<Longrightarrow> f -` B \<subseteq> A \<longleftrightarrow> B \<subseteq> f ` A" | 
| 567 | unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) | |
| 13585 | 568 | |
| 63322 | 569 | lemma inj_on_image_eq_iff: "inj_on f C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 64965 | 570 | by (fastforce simp: inj_on_def) | 
| 53927 | 571 | |
| 31438 | 572 | lemma inj_on_Un_image_eq_iff: "inj_on f (A \<union> B) \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 63322 | 573 | by (erule inj_on_image_eq_iff) simp_all | 
| 31438 | 574 | |
| 63322 | 575 | lemma inj_on_image_Int: "inj_on f C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` (A \<inter> B) = f ` A \<inter> f ` B" | 
| 576 | unfolding inj_on_def by blast | |
| 577 | ||
| 578 | lemma inj_on_image_set_diff: "inj_on f C \<Longrightarrow> A - B \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` (A - B) = f ` A - f ` B" | |
| 579 | unfolding inj_on_def by blast | |
| 13585 | 580 | |
| 63322 | 581 | lemma image_Int: "inj f \<Longrightarrow> f ` (A \<inter> B) = f ` A \<inter> f ` B" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 582 | unfolding inj_def by blast | 
| 13585 | 583 | |
| 63322 | 584 | lemma image_set_diff: "inj f \<Longrightarrow> f ` (A - B) = f ` A - f ` B" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 585 | unfolding inj_def by blast | 
| 13585 | 586 | |
| 63322 | 587 | lemma inj_on_image_mem_iff: "inj_on f B \<Longrightarrow> a \<in> B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> f a \<in> f ` A \<longleftrightarrow> a \<in> A" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 588 | by (auto simp: inj_on_def) | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 589 | |
| 63322 | 590 | lemma inj_image_mem_iff: "inj f \<Longrightarrow> f a \<in> f ` A \<longleftrightarrow> a \<in> A" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 591 | by (blast dest: injD) | 
| 13585 | 592 | |
| 63322 | 593 | lemma inj_image_subset_iff: "inj f \<Longrightarrow> f ` A \<subseteq> f ` B \<longleftrightarrow> A \<subseteq> B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 594 | by (blast dest: injD) | 
| 13585 | 595 | |
| 63322 | 596 | lemma inj_image_eq_iff: "inj f \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 597 | by (blast dest: injD) | 
| 13585 | 598 | |
| 63322 | 599 | lemma surj_Compl_image_subset: "surj f \<Longrightarrow> - (f ` A) \<subseteq> f ` (- A)" | 
| 600 | by auto | |
| 5852 | 601 | |
| 63322 | 602 | lemma inj_image_Compl_subset: "inj f \<Longrightarrow> f ` (- A) \<subseteq> - (f ` A)" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 603 | by (auto simp: inj_def) | 
| 63322 | 604 | |
| 605 | lemma bij_image_Compl_eq: "bij f \<Longrightarrow> f ` (- A) = - (f ` A)" | |
| 606 | by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) | |
| 13585 | 607 | |
| 41657 | 608 | lemma inj_vimage_singleton: "inj f \<Longrightarrow> f -` {a} \<subseteq> {THE x. f x = a}"
 | 
| 63322 | 609 | \<comment> \<open>The inverse image of a singleton under an injective function is included in a singleton.\<close> | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 610 | by (simp add: inj_def) (blast intro: the_equality [symmetric]) | 
| 41657 | 611 | |
| 63322 | 612 | lemma inj_on_vimage_singleton: "inj_on f A \<Longrightarrow> f -` {a} \<inter> A \<subseteq> {THE x. x \<in> A \<and> f x = a}"
 | 
| 43991 | 613 | by (auto simp add: inj_on_def intro: the_equality [symmetric]) | 
| 614 | ||
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 615 | lemma bij_betw_byWitness: | 
| 63322 | 616 | assumes left: "\<forall>a \<in> A. f' (f a) = a" | 
| 617 | and right: "\<forall>a' \<in> A'. f (f' a') = a'" | |
| 63575 | 618 | and "f ` A \<subseteq> A'" | 
| 619 | and img2: "f' ` A' \<subseteq> A" | |
| 63322 | 620 | shows "bij_betw f A A'" | 
| 621 | using assms | |
| 63400 | 622 | unfolding bij_betw_def inj_on_def | 
| 623 | proof safe | |
| 63322 | 624 | fix a b | 
| 63575 | 625 | assume "a \<in> A" "b \<in> A" | 
| 626 | with left have "a = f' (f a) \<and> b = f' (f b)" by simp | |
| 627 | moreover assume "f a = f b" | |
| 628 | ultimately show "a = b" by simp | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 629 | next | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 630 | fix a' assume *: "a' \<in> A'" | 
| 63575 | 631 | with img2 have "f' a' \<in> A" by blast | 
| 632 | moreover from * right have "a' = f (f' a')" by simp | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 633 | ultimately show "a' \<in> f ` A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 634 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 635 | |
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 636 | corollary notIn_Un_bij_betw: | 
| 63322 | 637 | assumes "b \<notin> A" | 
| 638 | and "f b \<notin> A'" | |
| 639 | and "bij_betw f A A'" | |
| 640 |   shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | |
| 641 | proof - | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 642 |   have "bij_betw f {b} {f b}"
 | 
| 63322 | 643 | unfolding bij_betw_def inj_on_def by simp | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 644 | with assms show ?thesis | 
| 63322 | 645 |     using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
 | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 646 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 647 | |
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 648 | lemma notIn_Un_bij_betw3: | 
| 63322 | 649 | assumes "b \<notin> A" | 
| 650 | and "f b \<notin> A'" | |
| 651 |   shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 652 | proof | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 653 | assume "bij_betw f A A'" | 
| 63322 | 654 |   then show "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | 
| 655 | using assms notIn_Un_bij_betw [of b A f A'] by blast | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 656 | next | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 657 |   assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 658 | have "f ` A = A'" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 659 | proof safe | 
| 63322 | 660 | fix a | 
| 661 | assume **: "a \<in> A" | |
| 662 |     then have "f a \<in> A' \<union> {f b}"
 | |
| 663 | using * unfolding bij_betw_def by blast | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 664 | moreover | 
| 63322 | 665 | have False if "f a = f b" | 
| 666 | proof - | |
| 63575 | 667 | have "a = b" | 
| 668 | using * ** that unfolding bij_betw_def inj_on_def by blast | |
| 63322 | 669 | with \<open>b \<notin> A\<close> ** show ?thesis by blast | 
| 670 | qed | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 671 | ultimately show "f a \<in> A'" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 672 | next | 
| 63322 | 673 | fix a' | 
| 674 | assume **: "a' \<in> A'" | |
| 675 |     then have "a' \<in> f ` (A \<union> {b})"
 | |
| 676 | using * by (auto simp add: bij_betw_def) | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 677 |     then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
 | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 678 | moreover | 
| 63322 | 679 | have False if "a = b" using 1 ** \<open>f b \<notin> A'\<close> that by blast | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 680 | ultimately have "a \<in> A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 681 | with 1 show "a' \<in> f ` A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 682 | qed | 
| 63322 | 683 | then show "bij_betw f A A'" | 
| 684 |     using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
 | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 685 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 686 | |
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 687 | lemma inj_on_disjoint_Un: | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 688 | assumes "inj_on f A" and "inj_on g B" | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 689 |   and "f ` A \<inter> g ` B = {}"
 | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 690 | shows "inj_on (\<lambda>x. if x \<in> A then f x else g x) (A \<union> B)" | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 691 | using assms by (simp add: inj_on_def disjoint_iff) (blast) | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 692 | |
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 693 | lemma bij_betw_disjoint_Un: | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 694 | assumes "bij_betw f A C" and "bij_betw g B D" | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 695 |   and "A \<inter> B = {}"
 | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 696 |   and "C \<inter> D = {}"
 | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 697 | shows "bij_betw (\<lambda>x. if x \<in> A then f x else g x) (A \<union> B) (C \<union> D)" | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 698 | using assms by (auto simp: inj_on_disjoint_Un bij_betw_def) | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 699 | |
| 73594 | 700 | lemma involuntory_imp_bij: | 
| 701 | \<open>bij f\<close> if \<open>\<And>x. f (f x) = x\<close> | |
| 702 | proof (rule bijI) | |
| 703 | from that show \<open>surj f\<close> | |
| 704 | by (rule surjI) | |
| 705 | show \<open>inj f\<close> | |
| 706 | proof (rule injI) | |
| 707 | fix x y | |
| 708 | assume \<open>f x = f y\<close> | |
| 709 | then have \<open>f (f x) = f (f y)\<close> | |
| 710 | by simp | |
| 711 | then show \<open>x = y\<close> | |
| 712 | by (simp add: that) | |
| 713 | qed | |
| 714 | qed | |
| 715 | ||
| 716 | ||
| 76261 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 717 | subsubsection \<open>Inj/surj/bij of Algebraic Operations\<close> | 
| 69502 | 718 | |
| 719 | context cancel_semigroup_add | |
| 720 | begin | |
| 721 | ||
| 69661 | 722 | lemma inj_on_add [simp]: | 
| 723 | "inj_on ((+) a) A" | |
| 724 | by (rule inj_onI) simp | |
| 725 | ||
| 726 | lemma inj_on_add' [simp]: | |
| 727 | "inj_on (\<lambda>b. b + a) A" | |
| 728 | by (rule inj_onI) simp | |
| 729 | ||
| 730 | lemma bij_betw_add [simp]: | |
| 731 | "bij_betw ((+) a) A B \<longleftrightarrow> (+) a ` A = B" | |
| 732 | by (simp add: bij_betw_def) | |
| 69502 | 733 | |
| 734 | end | |
| 735 | ||
| 76252 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 736 | context group_add | 
| 69502 | 737 | begin | 
| 738 | ||
| 76261 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 739 | lemma diff_left_imp_eq: "a - b = a - c \<Longrightarrow> b = c" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 740 | unfolding add_uminus_conv_diff[symmetric] | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 741 | by(drule local.add_left_imp_eq) simp | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 742 | |
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 743 | lemma inj_uminus[simp, intro]: "inj_on uminus A" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 744 | by (auto intro!: inj_onI) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 745 | |
| 76264 | 746 | lemma surj_uminus[simp]: "surj uminus" | 
| 747 | using surjI minus_minus by blast | |
| 748 | ||
| 69661 | 749 | lemma surj_plus [simp]: | 
| 750 | "surj ((+) a)" | |
| 76252 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 751 | proof (standard, simp, standard, simp) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 752 | fix x | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 753 | have "x = a + (-a + x)" by (simp add: add.assoc) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 754 | thus "x \<in> range ((+) a)" by blast | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 755 | qed | 
| 69661 | 756 | |
| 76261 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 757 | lemma surj_plus_right [simp]: | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 758 | "surj (\<lambda>b. b+a)" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 759 | proof (standard, simp, standard, simp) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 760 | fix b show "b \<in> range (\<lambda>b. b+a)" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 761 | using diff_add_cancel[of b a, symmetric] by blast | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 762 | qed | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 763 | |
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 764 | lemma inj_on_diff_left [simp]: | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 765 | \<open>inj_on ((-) a) A\<close> | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 766 | by (auto intro: inj_onI dest!: diff_left_imp_eq) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 767 | |
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 768 | lemma inj_on_diff_right [simp]: | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 769 | \<open>inj_on (\<lambda>b. b - a) A\<close> | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 770 | by (auto intro: inj_onI simp add: algebra_simps) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 771 | |
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 772 | lemma surj_diff [simp]: | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 773 | "surj ((-) a)" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 774 | proof (standard, simp, standard, simp) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 775 | fix x | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 776 | have "x = a - (- x + a)" by (simp add: algebra_simps) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 777 | thus "x \<in> range ((-) a)" by blast | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 778 | qed | 
| 69502 | 779 | |
| 69661 | 780 | lemma surj_diff_right [simp]: | 
| 781 | "surj (\<lambda>x. x - a)" | |
| 76252 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 782 | proof (standard, simp, standard, simp) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 783 | fix x | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 784 | have "x = x + a - a" by simp | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 785 | thus "x \<in> range (\<lambda>x. x - a)" by fast | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 786 | qed | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 787 | |
| 76261 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 788 | lemma shows bij_plus: "bij ((+) a)" and bij_plus_right: "bij (\<lambda>x. x + a)" | 
| 76264 | 789 | and bij_uminus: "bij uminus" | 
| 76261 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 790 | and bij_diff: "bij ((-) a)" and bij_diff_right: "bij (\<lambda>x. x - a)" | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 791 | by(simp_all add: bij_def) | 
| 
26524d0b4395
added and reorganized lemmas (some suggested by Jeremy Sylvestre)
 nipkow parents: 
76260diff
changeset | 792 | |
| 76252 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 793 | lemma translation_subtract_Compl: | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 794 | "(\<lambda>x. x - a) ` (- t) = - ((\<lambda>x. x - a) ` t)" | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 795 | by(rule bij_image_Compl_eq) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 796 | (auto simp add: bij_def surj_def inj_def diff_eq_eq intro!: add_diff_cancel[symmetric]) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 797 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 798 | lemma translation_diff: | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 799 | "(+) a ` (s - t) = ((+) a ` s) - ((+) a ` t)" | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 800 | by auto | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 801 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 802 | lemma translation_subtract_diff: | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 803 | "(\<lambda>x. x - a) ` (s - t) = ((\<lambda>x. x - a) ` s) - ((\<lambda>x. x - a) ` t)" | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 804 | by(rule image_set_diff)(simp add: inj_on_def diff_eq_eq) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 805 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 806 | lemma translation_Int: | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 807 | "(+) a ` (s \<inter> t) = ((+) a ` s) \<inter> ((+) a ` t)" | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 808 | by auto | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 809 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 810 | lemma translation_subtract_Int: | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 811 | "(\<lambda>x. x - a) ` (s \<inter> t) = ((\<lambda>x. x - a) ` s) \<inter> ((\<lambda>x. x - a) ` t)" | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 812 | by(rule image_Int)(simp add: inj_on_def diff_eq_eq) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 813 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 814 | end | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 815 | |
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 816 | (* TODO: prove in group_add *) | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 817 | context ab_group_add | 
| 
d123d9f67514
generalized type classes as suggested by Jeremy Sylvestre
 nipkow parents: 
76056diff
changeset | 818 | begin | 
| 69661 | 819 | |
| 820 | lemma translation_Compl: | |
| 821 | "(+) a ` (- t) = - ((+) a ` t)" | |
| 822 | proof (rule set_eqI) | |
| 823 | fix b | |
| 824 | show "b \<in> (+) a ` (- t) \<longleftrightarrow> b \<in> - (+) a ` t" | |
| 825 | by (auto simp: image_iff algebra_simps intro!: bexI [of _ "b - a"]) | |
| 826 | qed | |
| 827 | ||
| 69502 | 828 | end | 
| 829 | ||
| 41657 | 830 | |
| 63322 | 831 | subsection \<open>Function Updating\<close> | 
| 13585 | 832 | |
| 63322 | 833 | definition fun_upd :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 63324 | 834 | where "fun_upd f a b = (\<lambda>x. if x = a then b else f x)" | 
| 26147 | 835 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40969diff
changeset | 836 | nonterminal updbinds and updbind | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40969diff
changeset | 837 | |
| 26147 | 838 | syntax | 
| 63322 | 839 |   "_updbind" :: "'a \<Rightarrow> 'a \<Rightarrow> updbind"             ("(2_ :=/ _)")
 | 
| 840 |   ""         :: "updbind \<Rightarrow> updbinds"             ("_")
 | |
| 841 |   "_updbinds":: "updbind \<Rightarrow> updbinds \<Rightarrow> updbinds" ("_,/ _")
 | |
| 842 |   "_Update"  :: "'a \<Rightarrow> updbinds \<Rightarrow> 'a"            ("_/'((_)')" [1000, 0] 900)
 | |
| 26147 | 843 | |
| 844 | translations | |
| 63322 | 845 | "_Update f (_updbinds b bs)" \<rightleftharpoons> "_Update (_Update f b) bs" | 
| 846 | "f(x:=y)" \<rightleftharpoons> "CONST fun_upd f x y" | |
| 26147 | 847 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55066diff
changeset | 848 | (* Hint: to define the sum of two functions (or maps), use case_sum. | 
| 58111 | 849 | A nice infix syntax could be defined by | 
| 35115 | 850 | notation | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55066diff
changeset | 851 | case_sum (infixr "'(+')"80) | 
| 26147 | 852 | *) | 
| 853 | ||
| 63322 | 854 | lemma fun_upd_idem_iff: "f(x:=y) = f \<longleftrightarrow> f x = y" | 
| 855 | unfolding fun_upd_def | |
| 856 | apply safe | |
| 63575 | 857 | apply (erule subst) | 
| 858 | apply auto | |
| 63322 | 859 | done | 
| 13585 | 860 | |
| 63322 | 861 | lemma fun_upd_idem: "f x = y \<Longrightarrow> f(x := y) = f" | 
| 45603 | 862 | by (simp only: fun_upd_idem_iff) | 
| 13585 | 863 | |
| 45603 | 864 | lemma fun_upd_triv [iff]: "f(x := f x) = f" | 
| 865 | by (simp only: fun_upd_idem) | |
| 13585 | 866 | |
| 63322 | 867 | lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" | 
| 868 | by (simp add: fun_upd_def) | |
| 13585 | 869 | |
| 63322 | 870 | (* fun_upd_apply supersedes these two, but they are useful | 
| 13585 | 871 | if fun_upd_apply is intentionally removed from the simpset *) | 
| 63322 | 872 | lemma fun_upd_same: "(f(x := y)) x = y" | 
| 873 | by simp | |
| 13585 | 874 | |
| 63322 | 875 | lemma fun_upd_other: "z \<noteq> x \<Longrightarrow> (f(x := y)) z = f z" | 
| 876 | by simp | |
| 13585 | 877 | |
| 63322 | 878 | lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" | 
| 879 | by (simp add: fun_eq_iff) | |
| 13585 | 880 | |
| 63322 | 881 | lemma fun_upd_twist: "a \<noteq> c \<Longrightarrow> (m(a := b))(c := d) = (m(c := d))(a := b)" | 
| 71616 
a9de39608b1a
more tidying up of old apply-proofs
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 882 | by auto | 
| 63322 | 883 | |
| 884 | lemma inj_on_fun_updI: "inj_on f A \<Longrightarrow> y \<notin> f ` A \<Longrightarrow> inj_on (f(x := y)) A" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 885 | by (auto simp: inj_on_def) | 
| 15303 | 886 | |
| 63322 | 887 | lemma fun_upd_image: "f(x := y) ` A = (if x \<in> A then insert y (f ` (A - {x})) else f ` A)"
 | 
| 888 | by auto | |
| 15510 | 889 | |
| 31080 | 890 | lemma fun_upd_comp: "f \<circ> (g(x := y)) = (f \<circ> g)(x := f y)" | 
| 44921 | 891 | by auto | 
| 31080 | 892 | |
| 61630 | 893 | lemma fun_upd_eqD: "f(x := y) = g(x := z) \<Longrightarrow> y = z" | 
| 63322 | 894 | by (simp add: fun_eq_iff split: if_split_asm) | 
| 895 | ||
| 26147 | 896 | |
| 61799 | 897 | subsection \<open>\<open>override_on\<close>\<close> | 
| 26147 | 898 | |
| 63322 | 899 | definition override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 900 | where "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)" | |
| 13910 | 901 | |
| 15691 | 902 | lemma override_on_emptyset[simp]: "override_on f g {} = f"
 | 
| 64965 | 903 | by (simp add: override_on_def) | 
| 13910 | 904 | |
| 63322 | 905 | lemma override_on_apply_notin[simp]: "a \<notin> A \<Longrightarrow> (override_on f g A) a = f a" | 
| 64965 | 906 | by (simp add: override_on_def) | 
| 13910 | 907 | |
| 63322 | 908 | lemma override_on_apply_in[simp]: "a \<in> A \<Longrightarrow> (override_on f g A) a = g a" | 
| 64965 | 909 | by (simp add: override_on_def) | 
| 13910 | 910 | |
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 911 | lemma override_on_insert: "override_on f g (insert x X) = (override_on f g X)(x:=g x)" | 
| 64965 | 912 | by (simp add: override_on_def fun_eq_iff) | 
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 913 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 914 | lemma override_on_insert': "override_on f g (insert x X) = (override_on (f(x:=g x)) g X)" | 
| 64965 | 915 | by (simp add: override_on_def fun_eq_iff) | 
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 916 | |
| 26147 | 917 | |
| 60758 | 918 | subsection \<open>Inversion of injective functions\<close> | 
| 31949 | 919 | |
| 63322 | 920 | definition the_inv_into :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
 | 
| 63324 | 921 | where "the_inv_into A f = (\<lambda>x. THE y. y \<in> A \<and> f y = x)" | 
| 63322 | 922 | |
| 923 | lemma the_inv_into_f_f: "inj_on f A \<Longrightarrow> x \<in> A \<Longrightarrow> the_inv_into A f (f x) = x" | |
| 924 | unfolding the_inv_into_def inj_on_def by blast | |
| 32961 | 925 | |
| 63322 | 926 | lemma f_the_inv_into_f: "inj_on f A \<Longrightarrow> y \<in> f ` A \<Longrightarrow> f (the_inv_into A f y) = y" | 
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 927 | unfolding the_inv_into_def | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 928 | by (rule the1I2; blast dest: inj_onD) | 
| 32961 | 929 | |
| 72125 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 930 | lemma f_the_inv_into_f_bij_betw: | 
| 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 931 | "bij_betw f A B \<Longrightarrow> (bij_betw f A B \<Longrightarrow> x \<in> B) \<Longrightarrow> f (the_inv_into A f x) = x" | 
| 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 932 | unfolding bij_betw_def by (blast intro: f_the_inv_into_f) | 
| 
cf8399df4d76
elimination of some needless assumptions
 paulson <lp15@cam.ac.uk> parents: 
71857diff
changeset | 933 | |
| 63322 | 934 | lemma the_inv_into_into: "inj_on f A \<Longrightarrow> x \<in> f ` A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> the_inv_into A f x \<in> B" | 
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 935 | unfolding the_inv_into_def | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 936 | by (rule the1I2; blast dest: inj_onD) | 
| 32961 | 937 | |
| 63322 | 938 | lemma the_inv_into_onto [simp]: "inj_on f A \<Longrightarrow> the_inv_into A f ` (f ` A) = A" | 
| 939 | by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) | |
| 32961 | 940 | |
| 63322 | 941 | lemma the_inv_into_f_eq: "inj_on f A \<Longrightarrow> f x = y \<Longrightarrow> x \<in> A \<Longrightarrow> the_inv_into A f y = x" | 
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 942 | by (force simp add: the_inv_into_f_f) | 
| 32961 | 943 | |
| 33057 | 944 | lemma the_inv_into_comp: | 
| 63322 | 945 | "inj_on f (g ` A) \<Longrightarrow> inj_on g A \<Longrightarrow> x \<in> f ` g ` A \<Longrightarrow> | 
| 946 | the_inv_into A (f \<circ> g) x = (the_inv_into A g \<circ> the_inv_into (g ` A) f) x" | |
| 947 | apply (rule the_inv_into_f_eq) | |
| 948 | apply (fast intro: comp_inj_on) | |
| 949 | apply (simp add: f_the_inv_into_f the_inv_into_into) | |
| 950 | apply (simp add: the_inv_into_into) | |
| 951 | done | |
| 32961 | 952 | |
| 63322 | 953 | lemma inj_on_the_inv_into: "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)" | 
| 954 | by (auto intro: inj_onI simp: the_inv_into_f_f) | |
| 32961 | 955 | |
| 63322 | 956 | lemma bij_betw_the_inv_into: "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A" | 
| 957 | by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) | |
| 32961 | 958 | |
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 959 | lemma bij_betw_iff_bijections: | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 960 | "bij_betw f A B \<longleftrightarrow> (\<exists>g. (\<forall>x \<in> A. f x \<in> B \<and> g(f x) = x) \<and> (\<forall>y \<in> B. g y \<in> A \<and> f(g y) = y))" | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 961 | (is "?lhs = ?rhs") | 
| 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 962 | proof | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 963 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 964 | by (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 965 | exI[where ?x="the_inv_into A f"]) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 966 | next | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 967 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 968 | by (force intro: bij_betw_byWitness) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 969 | qed | 
| 71857 
d73955442df5
a few new lemmas about functions
 paulson <lp15@cam.ac.uk> parents: 
71827diff
changeset | 970 | |
| 63322 | 971 | abbreviation the_inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
 | 
| 972 | where "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 | 973 | |
| 64965 | 974 | lemma the_inv_f_f: "the_inv f (f x) = x" if "inj f" | 
| 975 | using that UNIV_I 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 | 976 | |
| 44277 
bcb696533579
moved fundamental lemma fun_eq_iff to theory HOL; tuned whitespace
 haftmann parents: 
43991diff
changeset | 977 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 978 | subsection \<open>Monotonicity\<close> | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 979 | |
| 75582 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 980 | definition monotone_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 981 | where "monotone_on A orda ordb f \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. orda x y \<longrightarrow> ordb (f x) (f y))" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 982 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 983 | abbreviation monotone :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 984 | where "monotone \<equiv> monotone_on UNIV" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 985 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 986 | lemma monotone_def[no_atp]: "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 987 | by (simp add: monotone_on_def) | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 988 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 989 | text \<open>Lemma @{thm [source] monotone_def} is provided for backward compatibility.\<close>
 | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 990 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 991 | lemma monotone_onI: | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 992 | "(\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)) \<Longrightarrow> monotone_on A orda ordb f" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 993 | by (simp add: monotone_on_def) | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 994 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 995 | lemma monotoneI[intro?]: "(\<And>x y. orda x y \<Longrightarrow> ordb (f x) (f y)) \<Longrightarrow> monotone orda ordb f" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 996 | by (rule monotone_onI) | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 997 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 998 | lemma monotone_onD: | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 999 | "monotone_on A orda ordb f \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)" | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 1000 | by (simp add: monotone_on_def) | 
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 1001 | |
| 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 1002 | lemma monotoneD[dest?]: "monotone orda ordb f \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)" | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1003 | by (rule monotone_onD[of UNIV, simplified]) | 
| 75582 
6fb4a0829cc4
added predicate monotone_on and redefined monotone to be an abbreviation.
 desharna parents: 
74123diff
changeset | 1004 | |
| 75583 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1005 | lemma monotone_on_subset: "monotone_on A orda ordb f \<Longrightarrow> B \<subseteq> A \<Longrightarrow> monotone_on B orda ordb f" | 
| 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1006 | by (auto intro: monotone_onI dest: monotone_onD) | 
| 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1007 | |
| 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1008 | lemma monotone_on_empty[simp]: "monotone_on {} orda ordb f"
 | 
| 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1009 | by (auto intro: monotone_onI dest: monotone_onD) | 
| 
451e17e0ba9d
added lemmas monotone_on_empty[simp] and monotone_on_subset
 desharna parents: 
75582diff
changeset | 1010 | |
| 75609 | 1011 | lemma monotone_on_o: | 
| 1012 | assumes | |
| 1013 | mono_f: "monotone_on A orda ordb f" and | |
| 1014 | mono_g: "monotone_on B ordc orda g" and | |
| 1015 | "g ` B \<subseteq> A" | |
| 1016 | shows "monotone_on B ordc ordb (f \<circ> g)" | |
| 1017 | proof (rule monotone_onI) | |
| 1018 | fix x y assume "x \<in> B" and "y \<in> B" and "ordc x y" | |
| 1019 | hence "orda (g x) (g y)" | |
| 1020 | by (rule mono_g[THEN monotone_onD]) | |
| 1021 | moreover from \<open>g ` B \<subseteq> A\<close> \<open>x \<in> B\<close> \<open>y \<in> B\<close> have "g x \<in> A" and "g y \<in> A" | |
| 1022 | unfolding image_subset_iff by simp_all | |
| 1023 | ultimately show "ordb ((f \<circ> g) x) ((f \<circ> g) y)" | |
| 1024 | using mono_f[THEN monotone_onD] by simp | |
| 1025 | qed | |
| 1026 | ||
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1027 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1028 | subsubsection \<open>Specializations For @{class ord} Type Class And More\<close>
 | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1029 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1030 | context ord begin | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1031 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1032 | abbreviation mono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
 | 
| 75608 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1033 | where "mono_on A \<equiv> monotone_on A (\<le>) (\<le>)" | 
| 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1034 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1035 | abbreviation strict_mono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
 | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1036 | where "strict_mono_on A \<equiv> monotone_on A (<) (<)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1037 | |
| 77934 | 1038 | abbreviation antimono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
 | 
| 1039 | where "antimono_on A \<equiv> monotone_on A (\<le>) (\<ge>)" | |
| 1040 | ||
| 1041 | abbreviation strict_antimono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
 | |
| 1042 | where "strict_antimono_on A \<equiv> monotone_on A (<) (>)" | |
| 1043 | ||
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1044 | lemma mono_on_def[no_atp]: "mono_on A f \<longleftrightarrow> (\<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s)" | 
| 75608 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1045 | by (auto simp add: monotone_on_def) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1046 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1047 | lemma strict_mono_on_def[no_atp]: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1048 | "strict_mono_on A f \<longleftrightarrow> (\<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1049 | by (auto simp add: monotone_on_def) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1050 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1051 | text \<open>Lemmas @{thm [source] mono_on_def} and @{thm [source] strict_mono_on_def} are provided for
 | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1052 | backward compatibility.\<close> | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1053 | |
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1054 | lemma mono_onI: | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1055 | "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on A f" | 
| 75608 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1056 | by (rule monotone_onI) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1057 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1058 | lemma strict_mono_onI: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1059 | "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on A f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1060 | by (rule monotone_onI) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1061 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1062 | lemma mono_onD: "\<lbrakk>mono_on A f; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s" | 
| 75608 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1063 | by (rule monotone_onD) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1064 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1065 | lemma strict_mono_onD: "\<lbrakk>strict_mono_on A f; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1066 | by (rule monotone_onD) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1067 | |
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1068 | lemma mono_on_subset: "mono_on A f \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on B f" | 
| 75608 
6c542e152b8a
redefined mono_on and strict_mono_on as an abbreviation of monotone_on
 desharna parents: 
75607diff
changeset | 1069 | by (rule monotone_on_subset) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1070 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1071 | end | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1072 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1073 | lemma mono_on_greaterD: | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1074 | assumes "mono_on A g" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1075 | shows "x > y" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1076 | proof (rule ccontr) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1077 | assume "\<not>x > y" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1078 | hence "x \<le> y" by (simp add: not_less) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1079 | from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1080 | with assms(4) show False by simp | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1081 | qed | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1082 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1083 | context order begin | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1084 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1085 | abbreviation mono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool"
 | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1086 | where "mono \<equiv> mono_on UNIV" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1087 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1088 | abbreviation strict_mono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool"
 | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1089 | where "strict_mono \<equiv> strict_mono_on UNIV" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1090 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1091 | abbreviation antimono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool"
 | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1092 | where "antimono \<equiv> monotone (\<le>) (\<lambda>x y. y \<le> x)" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1093 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1094 | lemma mono_def[no_atp]: "mono f \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1095 | by (simp add: monotone_on_def) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1096 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1097 | lemma strict_mono_def[no_atp]: "strict_mono f \<longleftrightarrow> (\<forall>x y. x < y \<longrightarrow> f x < f y)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1098 | by (simp add: monotone_on_def) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1099 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1100 | lemma antimono_def[no_atp]: "antimono f \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> f x \<ge> f y)" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1101 | by (simp add: monotone_on_def) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1102 | |
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1103 | text \<open>Lemmas @{thm [source] mono_def}, @{thm [source] strict_mono_def}, and
 | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1104 | @{thm [source] antimono_def} are provided for backward compatibility.\<close>
 | 
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1105 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1106 | lemma monoI [intro?]: "(\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y) \<Longrightarrow> mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1107 | by (rule monotoneI) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1108 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1109 | lemma strict_monoI [intro?]: "(\<And>x y. x < y \<Longrightarrow> f x < f y) \<Longrightarrow> strict_mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1110 | by (rule monotoneI) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1111 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1112 | lemma antimonoI [intro?]: "(\<And>x y. x \<le> y \<Longrightarrow> f x \<ge> f y) \<Longrightarrow> antimono f" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1113 | by (rule monotoneI) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1114 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1115 | lemma monoD [dest?]: "mono f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1116 | by (rule monotoneD) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1117 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1118 | lemma strict_monoD [dest?]: "strict_mono f \<Longrightarrow> x < y \<Longrightarrow> f x < f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1119 | by (rule monotoneD) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1120 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1121 | lemma antimonoD [dest?]: "antimono f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<ge> f y" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1122 | by (rule monotoneD) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1123 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1124 | lemma monoE: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1125 | assumes "mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1126 | assumes "x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1127 | obtains "f x \<le> f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1128 | proof | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1129 | from assms show "f x \<le> f y" by (simp add: mono_def) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1130 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1131 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1132 | lemma antimonoE: | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1133 | fixes f :: "'a \<Rightarrow> 'b::order" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1134 | assumes "antimono f" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1135 | assumes "x \<le> y" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1136 | obtains "f x \<ge> f y" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1137 | proof | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1138 | from assms show "f x \<ge> f y" by (simp add: antimono_def) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1139 | qed | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1140 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1141 | lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on A f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1142 | by (rule monotone_on_subset[OF _ subset_UNIV]) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1143 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1144 | lemma strict_mono_mono [dest?]: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1145 | assumes "strict_mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1146 | shows "mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1147 | proof (rule monoI) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1148 | fix x y | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1149 | assume "x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1150 | show "f x \<le> f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1151 | proof (cases "x = y") | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1152 | case True then show ?thesis by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1153 | next | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1154 | case False with \<open>x \<le> y\<close> have "x < y" by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1155 | with assms strict_monoD have "f x < f y" by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1156 | then show ?thesis by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1157 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1158 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1159 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1160 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1161 | end | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1162 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1163 | context linorder begin | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1164 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1165 | lemma mono_invE: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1166 | fixes f :: "'a \<Rightarrow> 'b::order" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1167 | assumes "mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1168 | assumes "f x < f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1169 | obtains "x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1170 | proof | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1171 | show "x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1172 | proof (rule ccontr) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1173 | assume "\<not> x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1174 | then have "y \<le> x" by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1175 | with \<open>mono f\<close> obtain "f y \<le> f x" by (rule monoE) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1176 | with \<open>f x < f y\<close> show False by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1177 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1178 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1179 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1180 | lemma mono_strict_invE: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1181 | fixes f :: "'a \<Rightarrow> 'b::order" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1182 | assumes "mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1183 | assumes "f x < f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1184 | obtains "x < y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1185 | proof | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1186 | show "x < y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1187 | proof (rule ccontr) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1188 | assume "\<not> x < y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1189 | then have "y \<le> x" by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1190 | with \<open>mono f\<close> obtain "f y \<le> f x" by (rule monoE) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1191 | with \<open>f x < f y\<close> show False by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1192 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1193 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1194 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1195 | lemma strict_mono_eq: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1196 | assumes "strict_mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1197 | shows "f x = f y \<longleftrightarrow> x = y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1198 | proof | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1199 | assume "f x = f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1200 | show "x = y" proof (cases x y rule: linorder_cases) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1201 | case less with assms strict_monoD have "f x < f y" by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1202 | with \<open>f x = f y\<close> show ?thesis by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1203 | next | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1204 | case equal then show ?thesis . | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1205 | next | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1206 | case greater with assms strict_monoD have "f y < f x" by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1207 | with \<open>f x = f y\<close> show ?thesis by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1208 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1209 | qed simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1210 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1211 | lemma strict_mono_less_eq: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1212 | assumes "strict_mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1213 | shows "f x \<le> f y \<longleftrightarrow> x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1214 | proof | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1215 | assume "x \<le> y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1216 | with assms strict_mono_mono monoD show "f x \<le> f y" by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1217 | next | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1218 | assume "f x \<le> f y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1219 | show "x \<le> y" proof (rule ccontr) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1220 | assume "\<not> x \<le> y" then have "y < x" by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1221 | with assms strict_monoD have "f y < f x" by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1222 | with \<open>f x \<le> f y\<close> show False by simp | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1223 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1224 | qed | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1225 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1226 | lemma strict_mono_less: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1227 | assumes "strict_mono f" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1228 | shows "f x < f y \<longleftrightarrow> x < y" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1229 | using assms | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1230 | by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1231 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1232 | end | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1233 | |
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1234 | lemma strict_mono_inv: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1235 |   fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
 | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1236 | assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1237 | shows "strict_mono g" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1238 | proof | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1239 | fix x y :: 'b assume "x < y" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1240 | from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1241 | with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1242 | with inv show "g x < g y" by simp | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1243 | qed | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1244 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1245 | lemma strict_mono_on_imp_inj_on: | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1246 | assumes "strict_mono_on A (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder))" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1247 | shows "inj_on f A" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1248 | proof (rule inj_onI) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1249 | fix x y assume "x \<in> A" "y \<in> A" "f x = f y" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1250 | thus "x = y" | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1251 | by (cases x y rule: linorder_cases) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1252 | (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1253 | qed | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1254 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1255 | lemma strict_mono_on_leD: | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1256 | assumes "strict_mono_on A (f :: (_ :: linorder) \<Rightarrow> _ :: preorder)" "x \<in> A" "y \<in> A" "x \<le> y" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1257 | shows "f x \<le> f y" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1258 | proof (cases "x = y") | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1259 | case True | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1260 | then show ?thesis by simp | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1261 | next | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1262 | case False | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1263 | with assms have "f x < f y" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1264 | using strict_mono_onD[OF assms(1)] by simp | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1265 | then show ?thesis by (rule less_imp_le) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1266 | qed | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1267 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1268 | lemma strict_mono_on_eqD: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1269 | fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)" | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1270 | assumes "strict_mono_on A f" "f x = f y" "x \<in> A" "y \<in> A" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1271 | shows "y = x" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75624diff
changeset | 1272 | using assms by (cases rule: linorder_cases) (auto dest: strict_mono_onD) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1273 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1274 | lemma strict_mono_on_imp_mono_on: | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
75583diff
changeset | 1275 | "strict_mono_on A (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) \<Longrightarrow> mono_on A f" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1276 | by (rule mono_onI, rule strict_mono_on_leD) | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 1277 | |
| 77934 | 1278 | lemma mono_imp_strict_mono: | 
| 1279 | fixes f :: "'a::order \<Rightarrow> 'b::order" | |
| 1280 | shows "\<lbrakk>mono_on S f; inj_on f S\<rbrakk> \<Longrightarrow> strict_mono_on S f" | |
| 1281 | by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) | |
| 1282 | ||
| 1283 | lemma strict_mono_iff_mono: | |
| 1284 | fixes f :: "'a::linorder \<Rightarrow> 'b::order" | |
| 1285 | shows "strict_mono_on S f \<longleftrightarrow> mono_on S f \<and> inj_on f S" | |
| 1286 | proof | |
| 1287 | show "strict_mono_on S f \<Longrightarrow> mono_on S f \<and> inj_on f S" | |
| 1288 | by (simp add: strict_mono_on_imp_inj_on strict_mono_on_imp_mono_on) | |
| 1289 | qed (auto intro: mono_imp_strict_mono) | |
| 1290 | ||
| 1291 | lemma antimono_imp_strict_antimono: | |
| 1292 | fixes f :: "'a::order \<Rightarrow> 'b::order" | |
| 1293 | shows "\<lbrakk>antimono_on S f; inj_on f S\<rbrakk> \<Longrightarrow> strict_antimono_on S f" | |
| 1294 | by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) | |
| 1295 | ||
| 1296 | lemma strict_antimono_iff_antimono: | |
| 1297 | fixes f :: "'a::linorder \<Rightarrow> 'b::order" | |
| 1298 | shows "strict_antimono_on S f \<longleftrightarrow> antimono_on S f \<and> inj_on f S" | |
| 1299 | proof | |
| 1300 | show "strict_antimono_on S f \<Longrightarrow> antimono_on S f \<and> inj_on f S" | |
| 1301 | by (force simp add: monotone_on_def intro: linorder_inj_onI) | |
| 1302 | qed (auto intro: antimono_imp_strict_antimono) | |
| 1303 | ||
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1304 | lemma mono_compose: "mono Q \<Longrightarrow> mono (\<lambda>i x. Q i (f x))" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1305 | unfolding mono_def le_fun_def by auto | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1306 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1307 | lemma mono_add: | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1308 | fixes a :: "'a::ordered_ab_semigroup_add" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1309 | shows "mono ((+) a)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1310 | by (simp add: add_left_mono monoI) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1311 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1312 | lemma (in semilattice_inf) mono_inf: "mono f \<Longrightarrow> f (A \<sqinter> B) \<le> f A \<sqinter> f B" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1313 | for f :: "'a \<Rightarrow> 'b::semilattice_inf" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1314 | by (auto simp add: mono_def intro: Lattices.inf_greatest) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1315 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1316 | lemma (in semilattice_sup) mono_sup: "mono f \<Longrightarrow> f A \<squnion> f B \<le> f (A \<squnion> B)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1317 | for f :: "'a \<Rightarrow> 'b::semilattice_sup" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1318 | by (auto simp add: mono_def intro: Lattices.sup_least) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1319 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1320 | lemma (in linorder) min_of_mono: "mono f \<Longrightarrow> min (f m) (f n) = f (min m n)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1321 | by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1322 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1323 | lemma (in linorder) max_of_mono: "mono f \<Longrightarrow> max (f m) (f n) = f (max m n)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1324 | by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1325 | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1326 | lemma (in linorder) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1327 | max_of_antimono: "antimono f \<Longrightarrow> max (f x) (f y) = f (min x y)" and | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1328 | min_of_antimono: "antimono f \<Longrightarrow> min (f x) (f y) = f (max x y)" | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1329 | by (auto simp: antimono_def Orderings.max_def max_def Orderings.min_def min_def intro!: antisym) | 
| 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 1330 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1331 | lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \<Longrightarrow> inj_on f A" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1332 | by (auto intro!: inj_onI dest: strict_mono_eq) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1333 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1334 | lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1335 | by (fact mono_inf) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1336 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1337 | lemma mono_Un: "mono f \<Longrightarrow> f A \<union> f B \<subseteq> f (A \<union> B)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1338 | by (fact mono_sup) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1339 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1340 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1341 | subsubsection \<open>Least value operator\<close> | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1342 | |
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1343 | lemma Least_mono: "mono f \<Longrightarrow> \<exists>x\<in>S. \<forall>y\<in>S. x \<le> y \<Longrightarrow> (LEAST y. y \<in> f ` S) = f (LEAST x. x \<in> S)" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1344 | for f :: "'a::order \<Rightarrow> 'b::order" | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1345 | \<comment> \<open>Courtesy of Stephan Merz\<close> | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1346 | apply clarify | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1347 | apply (erule_tac P = "\<lambda>x. x \<in> S" in LeastI2_order) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1348 | apply fast | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1349 | apply (rule LeastI2_order) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1350 | apply (auto elim: monoD intro!: order_antisym) | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1351 | done | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
75624diff
changeset | 1352 | |
| 63322 | 1353 | |
| 61204 | 1354 | subsection \<open>Setup\<close> | 
| 40969 | 1355 | |
| 60758 | 1356 | subsubsection \<open>Proof tools\<close> | 
| 22845 | 1357 | |
| 63400 | 1358 | text \<open>Simplify terms of the form \<open>f(\<dots>,x:=y,\<dots>,x:=z,\<dots>)\<close> to \<open>f(\<dots>,x:=z,\<dots>)\<close>\<close> | 
| 22845 | 1359 | |
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77934diff
changeset | 1360 | simproc_setup fun_upd2 ("f(v := w, x := y)") = \<open>
 | 
| 63322 | 1361 | let | 
| 1362 | fun gen_fun_upd NONE T _ _ = NONE | |
| 69593 | 1363 | | gen_fun_upd (SOME f) T x y = SOME (Const (\<^const_name>\<open>fun_upd\<close>, T) $ f $ x $ y) | 
| 63322 | 1364 | fun dest_fun_T1 (Type (_, T :: Ts)) = T | 
| 69593 | 1365 | fun find_double (t as Const (\<^const_name>\<open>fun_upd\<close>,T) $ f $ x $ y) = | 
| 63322 | 1366 | let | 
| 69593 | 1367 | fun find (Const (\<^const_name>\<open>fun_upd\<close>,T) $ g $ v $ w) = | 
| 63322 | 1368 | if v aconv x then SOME g else gen_fun_upd (find g) T v w | 
| 1369 | | find t = NONE | |
| 1370 | in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end | |
| 24017 | 1371 | |
| 69593 | 1372 | val ss = simpset_of \<^context> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51598diff
changeset | 1373 | |
| 63322 | 1374 | fun proc ctxt ct = | 
| 1375 | let | |
| 1376 | val t = Thm.term_of ct | |
| 1377 | in | |
| 63400 | 1378 | (case find_double t of | 
| 63322 | 1379 | (T, NONE) => NONE | 
| 1380 | | (T, SOME rhs) => | |
| 1381 | SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) | |
| 1382 | (fn _ => | |
| 1383 | resolve_tac ctxt [eq_reflection] 1 THEN | |
| 1384 |                 resolve_tac ctxt @{thms ext} 1 THEN
 | |
| 63400 | 1385 | simp_tac (put_simpset ss ctxt) 1))) | 
| 63322 | 1386 | end | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77934diff
changeset | 1387 | in K proc end | 
| 60758 | 1388 | \<close> | 
| 22845 | 1389 | |
| 1390 | ||
| 60758 | 1391 | subsubsection \<open>Functorial structure of types\<close> | 
| 40969 | 1392 | |
| 69605 | 1393 | ML_file \<open>Tools/functor.ML\<close> | 
| 40969 | 1394 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 1395 | functor map_fun: map_fun | 
| 47488 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 1396 | by (simp_all add: fun_eq_iff) | 
| 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 1397 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 1398 | functor vimage | 
| 49739 | 1399 | by (simp_all add: fun_eq_iff vimage_comp) | 
| 1400 | ||
| 63322 | 1401 | |
| 60758 | 1402 | text \<open>Legacy theorem names\<close> | 
| 49739 | 1403 | |
| 1404 | lemmas o_def = comp_def | |
| 1405 | lemmas o_apply = comp_apply | |
| 1406 | lemmas o_assoc = comp_assoc [symmetric] | |
| 1407 | lemmas id_o = id_comp | |
| 1408 | lemmas o_id = comp_id | |
| 1409 | lemmas o_eq_dest = comp_eq_dest | |
| 1410 | lemmas o_eq_elim = comp_eq_elim | |
| 55066 | 1411 | lemmas o_eq_dest_lhs = comp_eq_dest_lhs | 
| 1412 | lemmas o_eq_id_dest = comp_eq_id_dest | |
| 47488 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 1413 | |
| 2912 | 1414 | end |