| author | nipkow | 
| Wed, 03 Jul 2024 19:42:13 +0200 | |
| changeset 80484 | 0ca36dcdcbd3 | 
| parent 70817 | dd675800469d | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 68933 | 1 | (* Title: HOL/Library/Nonpos_Ints.thy | 
| 62055 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 2 | Author: Manuel Eberl, TU München | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 3 | *) | 
| 62055 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 4 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 5 | section \<open>Non-negative, non-positive integers and reals\<close> | 
| 62055 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 6 | |
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 7 | theory Nonpos_Ints | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 8 | imports Complex_Main | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 9 | begin | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 10 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 11 | subsection\<open>Non-positive integers\<close> | 
| 62055 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 12 | text \<open> | 
| 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 13 | The set of non-positive integers on a ring. (in analogy to the set of non-negative | 
| 69593 | 14 | integers \<^term>\<open>\<nat>\<close>) This is useful e.g. for the Gamma function. | 
| 62055 
755fda743c49
Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
 hoelzl parents: 
62049diff
changeset | 15 | \<close> | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 16 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 17 | definition nonpos_Ints ("\<int>\<^sub>\<le>\<^sub>0") where "\<int>\<^sub>\<le>\<^sub>0 = {of_int n |n. n \<le> 0}"
 | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 18 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 19 | lemma zero_in_nonpos_Ints [simp,intro]: "0 \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 20 | unfolding nonpos_Ints_def by (auto intro!: exI[of _ "0::int"]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 21 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 22 | lemma neg_one_in_nonpos_Ints [simp,intro]: "-1 \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 23 | unfolding nonpos_Ints_def by (auto intro!: exI[of _ "-1::int"]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 24 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 25 | lemma neg_numeral_in_nonpos_Ints [simp,intro]: "-numeral n \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 26 | unfolding nonpos_Ints_def by (auto intro!: exI[of _ "-numeral n::int"]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 27 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 28 | lemma one_notin_nonpos_Ints [simp]: "(1 :: 'a :: ring_char_0) \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 29 | by (auto simp: nonpos_Ints_def) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 30 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 31 | lemma numeral_notin_nonpos_Ints [simp]: "(numeral n :: 'a :: ring_char_0) \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 32 | by (auto simp: nonpos_Ints_def) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 33 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 34 | lemma minus_of_nat_in_nonpos_Ints [simp, intro]: "- of_nat n \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 35 | proof - | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 36 | have "- of_nat n = of_int (-int n)" by simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 37 | also have "-int n \<le> 0" by simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 38 | hence "of_int (-int n) \<in> \<int>\<^sub>\<le>\<^sub>0" unfolding nonpos_Ints_def by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 39 | finally show ?thesis . | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 40 | qed | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 41 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 42 | lemma of_nat_in_nonpos_Ints_iff: "(of_nat n :: 'a :: {ring_1,ring_char_0}) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n = 0"
 | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 43 | proof | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 44 | assume "(of_nat n :: 'a) \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 45 | then obtain m where "of_nat n = (of_int m :: 'a)" "m \<le> 0" by (auto simp: nonpos_Ints_def) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 46 | hence "(of_int m :: 'a) = of_nat n" by simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 47 | also have "... = of_int (int n)" by simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 48 | finally have "m = int n" by (subst (asm) of_int_eq_iff) | 
| 62072 | 49 | with \<open>m \<le> 0\<close> show "n = 0" by auto | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 50 | qed simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 51 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 52 | lemma nonpos_Ints_of_int: "n \<le> 0 \<Longrightarrow> of_int n \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 53 | unfolding nonpos_Ints_def by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 54 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 55 | lemma nonpos_IntsI: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 56 | "x \<in> \<int> \<Longrightarrow> x \<le> 0 \<Longrightarrow> (x :: 'a :: linordered_idom) \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 63092 | 57 | unfolding nonpos_Ints_def Ints_def by auto | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 58 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 59 | lemma nonpos_Ints_subset_Ints: "\<int>\<^sub>\<le>\<^sub>0 \<subseteq> \<int>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 60 | unfolding nonpos_Ints_def Ints_def by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 61 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 62 | lemma nonpos_Ints_nonpos [dest]: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x \<le> (0 :: 'a :: linordered_idom)" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 63 | unfolding nonpos_Ints_def by auto | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 64 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 65 | lemma nonpos_Ints_Int [dest]: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x \<in> \<int>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 66 | unfolding nonpos_Ints_def Ints_def by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 67 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 68 | lemma nonpos_Ints_cases: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 69 | assumes "x \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 70 | obtains n where "x = of_int n" "n \<le> 0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 71 | using assms unfolding nonpos_Ints_def by (auto elim!: Ints_cases) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 72 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 73 | lemma nonpos_Ints_cases': | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 74 | assumes "x \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 75 | obtains n where "x = -of_nat n" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 76 | proof - | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 77 | from assms obtain m where "x = of_int m" and m: "m \<le> 0" by (auto elim!: nonpos_Ints_cases) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 78 | hence "x = - of_int (-m)" by auto | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 79 | also from m have "(of_int (-m) :: 'a) = of_nat (nat (-m))" by simp_all | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 80 | finally show ?thesis by (rule that) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 81 | qed | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 82 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 83 | lemma of_real_in_nonpos_Ints_iff: "(of_real x :: 'a :: real_algebra_1) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 84 | proof | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 85 | assume "of_real x \<in> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 86 | then obtain n where "(of_real x :: 'a) = of_int n" "n \<le> 0" by (erule nonpos_Ints_cases) | 
| 62072 | 87 | note \<open>of_real x = of_int n\<close> | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 88 | also have "of_int n = of_real (of_int n)" by (rule of_real_of_int_eq [symmetric]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 89 | finally have "x = of_int n" by (subst (asm) of_real_eq_iff) | 
| 62072 | 90 | with \<open>n \<le> 0\<close> show "x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: nonpos_Ints_of_int) | 
| 62049 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 91 | qed (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 92 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 93 | lemma nonpos_Ints_altdef: "\<int>\<^sub>\<le>\<^sub>0 = {n \<in> \<int>. (n :: 'a :: linordered_idom) \<le> 0}"
 | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 94 | by (auto intro!: nonpos_IntsI elim!: nonpos_Ints_cases) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 95 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 96 | lemma uminus_in_Nats_iff: "-x \<in> \<nat> \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 97 | proof | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 98 | assume "-x \<in> \<nat>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 99 | then obtain n where "n \<ge> 0" "-x = of_int n" by (auto simp: Nats_altdef1) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 100 | hence "-n \<le> 0" "x = of_int (-n)" by (simp_all add: eq_commute minus_equation_iff[of x]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 101 | thus "x \<in> \<int>\<^sub>\<le>\<^sub>0" unfolding nonpos_Ints_def by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 102 | next | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 103 | assume "x \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 104 | then obtain n where "n \<le> 0" "x = of_int n" by (auto simp: nonpos_Ints_def) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 105 | hence "-n \<ge> 0" "-x = of_int (-n)" by (simp_all add: eq_commute minus_equation_iff[of x]) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 106 | thus "-x \<in> \<nat>" unfolding Nats_altdef1 by blast | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 107 | qed | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 108 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 109 | lemma uminus_in_nonpos_Ints_iff: "-x \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> x \<in> \<nat>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 110 | using uminus_in_Nats_iff[of "-x"] by simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 111 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 112 | lemma nonpos_Ints_mult: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x * y \<in> \<nat>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 113 | using Nats_mult[of "-x" "-y"] by (simp add: uminus_in_Nats_iff) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 114 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 115 | lemma Nats_mult_nonpos_Ints: "x \<in> \<nat> \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x * y \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 116 | using Nats_mult[of x "-y"] by (simp add: uminus_in_Nats_iff) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 117 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 118 | lemma nonpos_Ints_mult_Nats: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 119 | "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<nat> \<Longrightarrow> x * y \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 120 | using Nats_mult[of "-x" y] by (simp add: uminus_in_Nats_iff) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 121 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 122 | lemma nonpos_Ints_add: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 123 | "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x + y \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 124 | using Nats_add[of "-x" "-y"] uminus_in_Nats_iff[of "y+x", simplified minus_add] | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 125 | by (simp add: uminus_in_Nats_iff add.commute) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 126 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 127 | lemma nonpos_Ints_diff_Nats: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 128 | "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<nat> \<Longrightarrow> x - y \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 129 | using Nats_add[of "-x" "y"] uminus_in_Nats_iff[of "x-y", simplified minus_add] | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 130 | by (simp add: uminus_in_Nats_iff add.commute) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 131 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 132 | lemma Nats_diff_nonpos_Ints: | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 133 | "x \<in> \<nat> \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x - y \<in> \<nat>" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 134 | using Nats_add[of "x" "-y"] by (simp add: uminus_in_Nats_iff add.commute) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 135 | |
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 136 | lemma plus_of_nat_eq_0_imp: "z + of_nat n = 0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 137 | proof - | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 138 | assume "z + of_nat n = 0" | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 139 | hence A: "z = - of_nat n" by (simp add: eq_neg_iff_add_eq_0) | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 140 | show "z \<in> \<int>\<^sub>\<le>\<^sub>0" by (subst A) simp | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 141 | qed | 
| 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 eberlm parents: diff
changeset | 142 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 143 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 144 | subsection\<open>Non-negative reals\<close> | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 145 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 146 | definition nonneg_Reals :: "'a::real_algebra_1 set"  ("\<real>\<^sub>\<ge>\<^sub>0")
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 147 |   where "\<real>\<^sub>\<ge>\<^sub>0 = {of_real r | r. r \<ge> 0}"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 148 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 149 | lemma nonneg_Reals_of_real_iff [simp]: "of_real r \<in> \<real>\<^sub>\<ge>\<^sub>0 \<longleftrightarrow> r \<ge> 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 150 | by (force simp add: nonneg_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 151 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 152 | lemma nonneg_Reals_subset_Reals: "\<real>\<^sub>\<ge>\<^sub>0 \<subseteq> \<real>" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 153 | unfolding nonneg_Reals_def Reals_def by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 154 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 155 | lemma nonneg_Reals_Real [dest]: "x \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> x \<in> \<real>" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 156 | unfolding nonneg_Reals_def Reals_def by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 157 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 158 | lemma nonneg_Reals_of_nat_I [simp]: "of_nat n \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 159 | by (metis nonneg_Reals_of_real_iff of_nat_0_le_iff of_real_of_nat_eq) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 160 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 161 | lemma nonneg_Reals_cases: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 162 | assumes "x \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 163 | obtains r where "x = of_real r" "r \<ge> 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 164 | using assms unfolding nonneg_Reals_def by (auto elim!: Reals_cases) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 165 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 166 | lemma nonneg_Reals_zero_I [simp]: "0 \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 167 | unfolding nonneg_Reals_def by auto | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 168 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 169 | lemma nonneg_Reals_one_I [simp]: "1 \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 170 | by (metis (mono_tags, lifting) nonneg_Reals_of_nat_I of_nat_1) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 171 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 172 | lemma nonneg_Reals_minus_one_I [simp]: "-1 \<notin> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 173 | by (metis nonneg_Reals_of_real_iff le_minus_one_simps(3) of_real_1 of_real_def real_vector.scale_minus_left) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 174 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 175 | lemma nonneg_Reals_numeral_I [simp]: "numeral w \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 176 | by (metis (no_types) nonneg_Reals_of_nat_I of_nat_numeral) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 177 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 178 | lemma nonneg_Reals_minus_numeral_I [simp]: "- numeral w \<notin> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 179 | using nonneg_Reals_of_real_iff not_zero_le_neg_numeral by fastforce | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 180 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 181 | lemma nonneg_Reals_add_I [simp]: "\<lbrakk>a \<in> \<real>\<^sub>\<ge>\<^sub>0; b \<in> \<real>\<^sub>\<ge>\<^sub>0\<rbrakk> \<Longrightarrow> a + b \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 182 | apply (simp add: nonneg_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 183 | apply clarify | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 184 | apply (rename_tac r s) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 185 | apply (rule_tac x="r+s" in exI, auto) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 186 | done | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 187 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 188 | lemma nonneg_Reals_mult_I [simp]: "\<lbrakk>a \<in> \<real>\<^sub>\<ge>\<^sub>0; b \<in> \<real>\<^sub>\<ge>\<^sub>0\<rbrakk> \<Longrightarrow> a * b \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 189 | unfolding nonneg_Reals_def by (auto simp: of_real_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 190 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 191 | lemma nonneg_Reals_inverse_I [simp]: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 192 | fixes a :: "'a::real_div_algebra" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 193 | shows "a \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> inverse a \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 194 | by (simp add: nonneg_Reals_def image_iff) (metis inverse_nonnegative_iff_nonnegative of_real_inverse) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 195 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 196 | lemma nonneg_Reals_divide_I [simp]: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 197 | fixes a :: "'a::real_div_algebra" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 198 | shows "\<lbrakk>a \<in> \<real>\<^sub>\<ge>\<^sub>0; b \<in> \<real>\<^sub>\<ge>\<^sub>0\<rbrakk> \<Longrightarrow> a / b \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 199 | by (simp add: divide_inverse) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 200 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 201 | lemma nonneg_Reals_pow_I [simp]: "a \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> a^n \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 202 | by (induction n) auto | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 203 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 204 | lemma complex_nonneg_Reals_iff: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<longleftrightarrow> Re z \<ge> 0 \<and> Im z = 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 205 | by (auto simp: nonneg_Reals_def) (metis complex_of_real_def complex_surj) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 206 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 207 | lemma ii_not_nonneg_Reals [iff]: "\<i> \<notin> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 208 | by (simp add: complex_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 209 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 210 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 211 | subsection\<open>Non-positive reals\<close> | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 212 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 213 | definition nonpos_Reals :: "'a::real_algebra_1 set"  ("\<real>\<^sub>\<le>\<^sub>0")
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 214 |   where "\<real>\<^sub>\<le>\<^sub>0 = {of_real r | r. r \<le> 0}"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 215 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 216 | lemma nonpos_Reals_of_real_iff [simp]: "of_real r \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> r \<le> 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 217 | by (force simp add: nonpos_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 218 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 219 | lemma nonpos_Reals_subset_Reals: "\<real>\<^sub>\<le>\<^sub>0 \<subseteq> \<real>" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 220 | unfolding nonpos_Reals_def Reals_def by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 221 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 222 | lemma nonpos_Ints_subset_nonpos_Reals: "\<int>\<^sub>\<le>\<^sub>0 \<subseteq> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 223 | by (metis nonpos_Ints_cases nonpos_Ints_nonpos nonpos_Ints_of_int | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 224 | nonpos_Reals_of_real_iff of_real_of_int_eq subsetI) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 225 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 226 | lemma nonpos_Reals_of_nat_iff [simp]: "of_nat n \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n=0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 227 | by (metis nonpos_Reals_of_real_iff of_nat_le_0_iff of_real_of_nat_eq) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 228 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 229 | lemma nonpos_Reals_Real [dest]: "x \<in> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x \<in> \<real>" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 230 | unfolding nonpos_Reals_def Reals_def by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 231 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 232 | lemma nonpos_Reals_cases: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 233 | assumes "x \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 234 | obtains r where "x = of_real r" "r \<le> 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 235 | using assms unfolding nonpos_Reals_def by (auto elim!: Reals_cases) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 236 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 237 | lemma uminus_nonneg_Reals_iff [simp]: "-x \<in> \<real>\<^sub>\<ge>\<^sub>0 \<longleftrightarrow> x \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 238 | apply (auto simp: nonpos_Reals_def nonneg_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 239 | apply (metis nonpos_Reals_of_real_iff minus_minus neg_le_0_iff_le of_real_minus) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 240 | done | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 241 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 242 | lemma uminus_nonpos_Reals_iff [simp]: "-x \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> x \<in> \<real>\<^sub>\<ge>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 243 | by (metis (no_types) minus_minus uminus_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 244 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 245 | lemma nonpos_Reals_zero_I [simp]: "0 \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 246 | unfolding nonpos_Reals_def by force | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 247 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 248 | lemma nonpos_Reals_one_I [simp]: "1 \<notin> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 249 | using nonneg_Reals_minus_one_I uminus_nonneg_Reals_iff by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 250 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 251 | lemma nonpos_Reals_numeral_I [simp]: "numeral w \<notin> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 252 | using nonneg_Reals_minus_numeral_I uminus_nonneg_Reals_iff by blast | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 253 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 254 | lemma nonpos_Reals_add_I [simp]: "\<lbrakk>a \<in> \<real>\<^sub>\<le>\<^sub>0; b \<in> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> a + b \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 255 | by (metis nonneg_Reals_add_I add_uminus_conv_diff minus_diff_eq minus_minus uminus_nonpos_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 256 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 257 | lemma nonpos_Reals_mult_I1: "\<lbrakk>a \<in> \<real>\<^sub>\<ge>\<^sub>0; b \<in> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> a * b \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 258 | by (metis nonneg_Reals_mult_I mult_minus_right uminus_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 259 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 260 | lemma nonpos_Reals_mult_I2: "\<lbrakk>a \<in> \<real>\<^sub>\<le>\<^sub>0; b \<in> \<real>\<^sub>\<ge>\<^sub>0\<rbrakk> \<Longrightarrow> a * b \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 261 | by (metis nonneg_Reals_mult_I mult_minus_left uminus_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 262 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 263 | lemma nonpos_Reals_mult_of_nat_iff: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 264 | fixes a:: "'a :: real_div_algebra" shows "a * of_nat n \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> a \<in> \<real>\<^sub>\<le>\<^sub>0 \<or> n=0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 265 | apply (auto intro: nonpos_Reals_mult_I2) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 266 | apply (auto simp: nonpos_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 267 | apply (rule_tac x="r/n" in exI) | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
69593diff
changeset | 268 | apply (auto simp: field_split_simps) | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 269 | done | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 270 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 271 | lemma nonpos_Reals_inverse_I: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 272 | fixes a :: "'a::real_div_algebra" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 273 | shows "a \<in> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> inverse a \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 274 | using nonneg_Reals_inverse_I uminus_nonneg_Reals_iff by fastforce | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 275 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 276 | lemma nonpos_Reals_divide_I1: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 277 | fixes a :: "'a::real_div_algebra" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 278 | shows "\<lbrakk>a \<in> \<real>\<^sub>\<ge>\<^sub>0; b \<in> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> a / b \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 279 | by (simp add: nonpos_Reals_inverse_I nonpos_Reals_mult_I1 divide_inverse) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 280 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 281 | lemma nonpos_Reals_divide_I2: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 282 | fixes a :: "'a::real_div_algebra" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 283 | shows "\<lbrakk>a \<in> \<real>\<^sub>\<le>\<^sub>0; b \<in> \<real>\<^sub>\<ge>\<^sub>0\<rbrakk> \<Longrightarrow> a / b \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 284 | by (metis nonneg_Reals_divide_I minus_divide_left uminus_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 285 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 286 | lemma nonpos_Reals_divide_of_nat_iff: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 287 | fixes a:: "'a :: real_div_algebra" shows "a / of_nat n \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> a \<in> \<real>\<^sub>\<le>\<^sub>0 \<or> n=0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 288 | apply (auto intro: nonpos_Reals_divide_I2) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 289 | apply (auto simp: nonpos_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 290 | apply (rule_tac x="r*n" in exI) | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
69593diff
changeset | 291 | apply (auto simp: field_split_simps mult_le_0_iff) | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 292 | done | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 293 | |
| 68499 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
67135diff
changeset | 294 | lemma nonpos_Reals_inverse_iff [simp]: | 
| 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
67135diff
changeset | 295 | fixes a :: "'a::real_div_algebra" | 
| 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
67135diff
changeset | 296 | shows "inverse a \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> a \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
67135diff
changeset | 297 | using nonpos_Reals_inverse_I by fastforce | 
| 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
67135diff
changeset | 298 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 299 | lemma nonpos_Reals_pow_I: "\<lbrakk>a \<in> \<real>\<^sub>\<le>\<^sub>0; odd n\<rbrakk> \<Longrightarrow> a^n \<in> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 300 | by (metis nonneg_Reals_pow_I power_minus_odd uminus_nonneg_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 301 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 302 | lemma complex_nonpos_Reals_iff: "z \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> Re z \<le> 0 \<and> Im z = 0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 303 | using complex_is_Real_iff by (force simp add: nonpos_Reals_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 304 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 305 | lemma ii_not_nonpos_Reals [iff]: "\<i> \<notin> \<real>\<^sub>\<le>\<^sub>0" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 306 | by (simp add: complex_nonpos_Reals_iff) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62072diff
changeset | 307 | |
| 62390 | 308 | end |