| author | wenzelm | 
| Tue, 28 Jul 2015 18:57:10 +0200 | |
| changeset 60815 | c93a83472eab | 
| parent 59190 | 3a594fd13ca4 | 
| child 62225 | c8c48906b858 | 
| permissions | -rw-r--r-- | 
| 59190 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 1 | (* Title: HOL/ex/Cubic_Quartic.thy | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 2 | Author: Amine Chaieb | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 3 | *) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 4 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 5 | header "The Cubic and Quartic Root Formulas" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 6 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 7 | theory Cubic_Quartic | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 8 | imports Complex_Main | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 9 | begin | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 10 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 11 | section "The Cubic Formula" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 12 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 13 | definition "ccbrt z = (SOME (w::complex). w^3 = z)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 14 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 15 | lemma ccbrt: "(ccbrt z) ^ 3 = z" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 16 | proof- | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 17 | from rcis_Ex obtain r a where ra: "z = rcis r a" by blast | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 18 | let ?r' = "if r < 0 then - root 3 (-r) else root 3 r" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 19 | let ?a' = "a/3" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 20 | have "rcis ?r' ?a' ^ 3 = rcis r a" by (cases "r<0", simp_all add: DeMoivre2) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 21 | hence th: "\<exists>w. w^3 = z" unfolding ra by blast | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 22 | from someI_ex[OF th] show ?thesis unfolding ccbrt_def by blast | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 23 | qed | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 24 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 25 | text "The reduction to a simpler form:" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 26 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 27 | lemma cubic_reduction: | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 28 | fixes a :: complex | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 29 | assumes H: "a \<noteq> 0 \<and> x = y - b / (3 * a) \<and> p = (3* a * c - b^2) / (9 * a^2) \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 30 | q = (9 * a * b * c - 2 * b^3 - 27 * a^2 * d) / (54 * a^3)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 31 | shows "a * x^3 + b * x^2 + c * x + d = 0 \<longleftrightarrow> y^3 + 3 * p * y - 2 * q = 0" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 32 | proof- | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 33 | from H have "3*a \<noteq> 0" "9*a^2 \<noteq> 0" "54*a^3 \<noteq> 0" by auto | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 34 | hence th: "x = y - b / (3 * a) \<longleftrightarrow> (3*a) * x = (3*a) * y - b" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 35 | "p = (3* a * c - b^2) / (9 * a^2) \<longleftrightarrow> (9 * a^2) * p = (3* a * c - b^2)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 36 | "q = (9 * a * b * c - 2 * b^3 - 27 * a^2 * d) / (54 * a^3) \<longleftrightarrow> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 37 | (54 * a^3) * q = (9 * a * b * c - 2 * b^3 - 27 * a^2 * d)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 38 | by (simp_all add: field_simps) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 39 | from H[unfolded th] show ?thesis by algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 40 | qed | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 41 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 42 | text "The solutions of the special form:" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 43 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 44 | lemma cubic_basic: | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 45 | fixes s :: complex | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 46 | assumes H: "s^2 = q^2 + p^3 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 47 | s1^3 = (if p = 0 then 2 * q else q + s) \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 48 | s2 = -s1 * (1 + i * t) / 2 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 49 | s3 = -s1 * (1 - i * t) / 2 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 50 | i^2 + 1 = 0 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 51 | t^2 = 3" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 52 | shows | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 53 | "if p = 0 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 54 | then y^3 + 3 * p * y - 2 * q = 0 \<longleftrightarrow> y = s1 \<or> y = s2 \<or> y = s3 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 55 | else s1 \<noteq> 0 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 56 | (y^3 + 3 * p * y - 2 * q = 0 \<longleftrightarrow> (y = s1 - p / s1 \<or> y = s2 - p / s2 \<or> y = s3 - p / s3))" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 57 | proof- | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 58 |  { assume p0: "p = 0"
 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 59 | with H have ?thesis by (simp add: field_simps) algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 60 | } | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 61 | moreover | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 62 |  { assume p0: "p \<noteq> 0"
 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 63 | with H have th1: "s1 \<noteq> 0" by (simp add: field_simps) algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 64 | from p0 H th1 have th0: "s2 \<noteq> 0" "s3 \<noteq>0" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 65 | by (simp_all add: field_simps) algebra+ | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 66 | from th1 th0 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 67 | have th: "y = s1 - p / s1 \<longleftrightarrow> s1*y = s1^2 - p" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 68 | "y = s2 - p / s2 \<longleftrightarrow> s2*y = s2^2 - p" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 69 | "y = s3 - p / s3 \<longleftrightarrow> s3*y = s3^2 - p" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 70 | by (simp_all add: field_simps power2_eq_square) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 71 | from p0 H have ?thesis unfolding th by (simp add: field_simps) algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 72 | } | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 73 | ultimately show ?thesis by blast | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 74 | qed | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 75 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 76 | text "Explicit formula for the roots:" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 77 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 78 | lemma cubic: | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 79 | assumes a0: "a \<noteq> 0" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 80 | shows " | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 81 | let p = (3 * a * c - b^2) / (9 * a^2) ; | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 82 | q = (9 * a * b * c - 2 * b^3 - 27 * a^2 * d) / (54 * a^3); | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 83 | s = csqrt(q^2 + p^3); | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 84 | s1 = (if p = 0 then ccbrt(2 * q) else ccbrt(q + s)); | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 85 | s2 = -s1 * (1 + ii * csqrt 3) / 2; | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 86 | s3 = -s1 * (1 - ii * csqrt 3) / 2 | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 87 | in if p = 0 then | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 88 | a * x^3 + b * x^2 + c * x + d = 0 \<longleftrightarrow> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 89 | x = s1 - b / (3 * a) \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 90 | x = s2 - b / (3 * a) \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 91 | x = s3 - b / (3 * a) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 92 | else | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 93 | s1 \<noteq> 0 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 94 | (a * x^3 + b * x^2 + c * x + d = 0 \<longleftrightarrow> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 95 | x = s1 - p / s1 - b / (3 * a) \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 96 | x = s2 - p / s2 - b / (3 * a) \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 97 | x = s3 - p / s3 - b / (3 * a))" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 98 | proof- | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 99 | let ?p = "(3 * a * c - b^2) / (9 * a^2)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 100 | let ?q = "(9 * a * b * c - 2 * b^3 - 27 * a^2 * d) / (54 * a^3)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 101 | let ?s = "csqrt(?q^2 + ?p^3)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 102 | let ?s1 = "if ?p = 0 then ccbrt(2 * ?q) else ccbrt(?q + ?s)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 103 | let ?s2 = "- ?s1 * (1 + ii * csqrt 3) / 2" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 104 | let ?s3 = "- ?s1 * (1 - ii * csqrt 3) / 2" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 105 | let ?y = "x + b / (3 * a)" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 106 | from a0 have zero: "9 * a^2 \<noteq> 0" "a^3 * 54 \<noteq> 0" "(a*3)\<noteq> 0" by auto | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 107 | have eq:"a * x^3 + b * x^2 + c * x + d = 0 \<longleftrightarrow> ?y^3 + 3 * ?p * ?y - 2 * ?q = 0" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 108 | by (rule cubic_reduction) (auto simp add: field_simps zero a0) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 109 | have "csqrt 3^2 = 3" by (rule power2_csqrt) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 110 | hence th0: "?s^2 = ?q^2 + ?p ^ 3 \<and> ?s1^ 3 = (if ?p = 0 then 2 * ?q else ?q + ?s) \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 111 | ?s2 = - ?s1 * (1 + ii * csqrt 3) / 2 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 112 | ?s3 = - ?s1 * (1 - ii * csqrt 3) / 2 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 113 | ii^2 + 1 = 0 \<and> csqrt 3^2 = 3" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 114 | using zero by (simp add: field_simps power2_csqrt ccbrt) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 115 | from cubic_basic[OF th0, of ?y] | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 116 | show ?thesis | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 117 | apply (simp only: Let_def eq) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 118 | using zero apply (simp add: field_simps ccbrt power2_csqrt) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 119 | using zero | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 120 | apply (cases "a * (c * 3) = b^2", simp_all add: field_simps) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 121 | done | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 122 | qed | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 123 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 124 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 125 | section "The Quartic Formula" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 126 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 127 | lemma quartic: | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 128 | "(y::real)^3 - b * y^2 + (a * c - 4 * d) * y - a^2 * d + 4 * b * d - c^2 = 0 \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 129 | R^2 = a^2 / 4 - b + y \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 130 | s^2 = y^2 - 4 * d \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 131 | (D^2 = (if R = 0 then 3 * a^2 / 4 - 2 * b + 2 * s | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 132 | else 3 * a^2 / 4 - R^2 - 2 * b + (4 * a * b - 8 * c - a^3) / (4 * R))) \<and> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 133 | (E^2 = (if R = 0 then 3 * a^2 / 4 - 2 * b - 2 * s | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 134 | else 3 * a^2 / 4 - R^2 - 2 * b - (4 * a * b - 8 * c - a^3) / (4 * R))) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 135 | \<Longrightarrow> x^4 + a * x^3 + b * x^2 + c * x + d = 0 \<longleftrightarrow> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 136 | x = -a / 4 + R / 2 + D / 2 \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 137 | x = -a / 4 + R / 2 - D / 2 \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 138 | x = -a / 4 - R / 2 + E / 2 \<or> | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 139 | x = -a / 4 - R / 2 - E / 2" | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 140 | apply (cases "R=0", simp_all add: field_simps divide_minus_left[symmetric] del: divide_minus_left) | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 141 | apply algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 142 | apply algebra | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 143 | done | 
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 144 | |
| 
3a594fd13ca4
3 old example lemmas by Amine listed in the top 100 theorems
 kleing parents: diff
changeset | 145 | end |