src/HOL/Fun.thy
changeset 71472 c213d067e60f
parent 71464 4a04b6bd628b
child 71616 a9de39608b1a
equal deleted inserted replaced
71471:c06604896c3d 71472:c213d067e60f
   911   then have "?X \<in> f ` A" by (simp only: f)
   911   then have "?X \<in> f ` A" by (simp only: f)
   912   then obtain x where "x \<in> A" and "f x = ?X" by blast
   912   then obtain x where "x \<in> A" and "f x = ?X" by blast
   913   then show False by blast
   913   then show False by blast
   914 qed
   914 qed
   915 
   915 
       
   916 subsection \<open>Monotonic functions over a set\<close>
       
   917 
       
   918 definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
       
   919 
       
   920 lemma mono_onI:
       
   921   "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
       
   922   unfolding mono_on_def by simp
       
   923 
       
   924 lemma mono_onD:
       
   925   "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
       
   926   unfolding mono_on_def by simp
       
   927 
       
   928 lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
       
   929   unfolding mono_def mono_on_def by auto
       
   930 
       
   931 lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
       
   932   unfolding mono_on_def by auto
       
   933 
       
   934 definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
       
   935 
       
   936 lemma strict_mono_onI:
       
   937   "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
       
   938   unfolding strict_mono_on_def by simp
       
   939 
       
   940 lemma strict_mono_onD:
       
   941   "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
       
   942   unfolding strict_mono_on_def by simp
       
   943 
       
   944 lemma mono_on_greaterD:
       
   945   assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
       
   946   shows "x > y"
       
   947 proof (rule ccontr)
       
   948   assume "\<not>x > y"
       
   949   hence "x \<le> y" by (simp add: not_less)
       
   950   from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD)
       
   951   with assms(4) show False by simp
       
   952 qed
       
   953 
       
   954 lemma strict_mono_inv:
       
   955   fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
       
   956   assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
       
   957   shows "strict_mono g"
       
   958 proof
       
   959   fix x y :: 'b assume "x < y"
       
   960   from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast
       
   961   with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less)
       
   962   with inv show "g x < g y" by simp
       
   963 qed
       
   964 
       
   965 lemma strict_mono_on_imp_inj_on:
       
   966   assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
       
   967   shows "inj_on f A"
       
   968 proof (rule inj_onI)
       
   969   fix x y assume "x \<in> A" "y \<in> A" "f x = f y"
       
   970   thus "x = y"
       
   971     by (cases x y rule: linorder_cases)
       
   972        (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
       
   973 qed
       
   974 
       
   975 lemma strict_mono_on_leD:
       
   976   assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
       
   977   shows "f x \<le> f y"
       
   978 proof (insert le_less_linear[of y x], elim disjE)
       
   979   assume "x < y"
       
   980   with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all
       
   981   thus ?thesis by (rule less_imp_le)
       
   982 qed (insert assms, simp)
       
   983 
       
   984 lemma strict_mono_on_eqD:
       
   985   fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
       
   986   assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
       
   987   shows "y = x"
       
   988   using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
       
   989 
       
   990 lemma strict_mono_on_imp_mono_on:
       
   991   "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
       
   992   by (rule mono_onI, rule strict_mono_on_leD)
       
   993 
   916 
   994 
   917 subsection \<open>Setup\<close>
   995 subsection \<open>Setup\<close>
   918 
   996 
   919 subsubsection \<open>Proof tools\<close>
   997 subsubsection \<open>Proof tools\<close>
   920 
   998