diff --git a/thys/Virtual_Substitution/ExecutiblePolyProps.thy b/thys/Virtual_Substitution/ExecutiblePolyProps.thy --- a/thys/Virtual_Substitution/ExecutiblePolyProps.thy +++ b/thys/Virtual_Substitution/ExecutiblePolyProps.thy @@ -1,1083 +1,1084 @@ text "Executable Polynomial Properties" theory ExecutiblePolyProps imports Polynomials.MPoly_Type_Univariate MPolyExtension begin text \(Univariate) Polynomial hiding\ lifting_update poly.lifting lifting_forget poly.lifting text \\ no_notation MPoly_Type.div (infixl "div" 70) no_notation MPoly_Type.mod (infixl "mod" 70) subsection "Lemmas with Monomial and Monomials" lemma of_nat_monomial: "of_nat p = monomial p 0" by (auto simp: poly_mapping_eq_iff lookup_of_nat fun_eq_iff lookup_single) lemma of_nat_times_monomial: "of_nat p * monomial c i = monomial (p*c) i" by (auto simp: poly_mapping_eq_iff prod_fun_def fun_eq_iff of_nat_monomial lookup_single mult_single) lemma monomial_adds_nat_iff: "monomial p i adds c \ lookup c i \ p" for p::"nat" apply (auto simp: adds_def lookup_add) by (metis add.left_commute nat_le_iff_add remove_key_sum single_add) lemma update_minus_monomial: "Poly_Mapping.update k i (m - monomial i k) = Poly_Mapping.update k i m" by (auto simp: poly_mapping_eq_iff lookup_update update.rep_eq fun_eq_iff lookup_minus lookup_single) lemma monomials_Var: "monomials (Var x::'a::zero_neq_one mpoly) = {Poly_Mapping.single x 1}" by transfer (auto simp: Var\<^sub>0_def) lemma monomials_Const: "monomials (Const x) = (if x = 0 then {} else {0})" by transfer' (auto simp: Const\<^sub>0_def) lemma coeff_eq_zero_iff: "MPoly_Type.coeff c p = 0 \ p \ monomials c" by transfer (simp add: not_in_keys_iff_lookup_eq_zero) lemma monomials_1[simp]: "monomials 1 = {0}" by transfer auto lemma monomials_and_monoms: shows "(k \ monomials m) = (\ (a::nat). a \ 0 \ (monomials (MPoly_Type.monom k a)) \ monomials m)" proof - show ?thesis using monomials_monom by auto qed lemma mult_monomials_dir_one: shows "monomials (p*q) \ {a+b | a b . a \ monomials p \ b \ monomials q}" using monomials_and_monoms mult_monom by (simp add: keys_mult monomials.rep_eq times_mpoly.rep_eq) lemma monom_eq_zero_iff[simp]: "MPoly_Type.monom a b = 0 \ b = 0" by (metis MPolyExtension.coeff_monom MPolyExtension.monom_zero) lemma update_eq_plus_monomial: "v \ lookup m k \ Poly_Mapping.update k v m = m + monomial (v - lookup m k) k" for v n::nat by transfer auto lemma coeff_monom_mult': "MPoly_Type.coeff ((MPoly_Type.monom m' a) * q) (m'm) = a * MPoly_Type.coeff q (m'm - m')" if *: "m'm = m' + (m'm - m')" by (subst *) (rule More_MPoly_Type.coeff_monom_mult) lemma monomials_zero[simp]: "monomials 0 = {}" by transfer auto lemma in_monomials_iff: "x \ monomials m \ MPoly_Type.coeff m x \ 0" using coeff_eq_zero_iff[of m x] by auto lemma monomials_monom_mult: "monomials (MPoly_Type.monom mon a * q) = (if a = 0 then {} else (+) mon ` monomials q)" for q::"'a::semiring_no_zero_divisors mpoly" apply auto subgoal by transfer' (auto elim!: in_keys_timesE) subgoal by (simp add: in_monomials_iff More_MPoly_Type.coeff_monom_mult) done subsection "Simplification Lemmas for Const 0 and Const 1" lemma add_zero : "P + Const 0 = P" proof - have h:"P + 0 = P" using add_0_right by auto show ?thesis unfolding Const_def using h by (simp add: Const\<^sub>0_zero zero_mpoly.abs_eq) qed (* example *) lemma add_zero_example : "((Var 0)^2 - (Const 1)) + Const 0 = ((Var 0)^2 - (Const 1))" proof - show ?thesis by (simp add : add_zero) qed lemma mult_zero_left : "Const 0 * P = Const 0" proof - have h:"0*P = 0" by simp show ?thesis unfolding Const_def using h by (simp add: Const\<^sub>0_zero zero_mpoly_def) qed lemma mult_zero_right : "P * Const 0 = Const 0" by (metis mult_zero_left mult_zero_right) lemma mult_one_left : "Const 1 * (P :: real mpoly) = P" by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly_def) lemma mult_one_right : "(P::real mpoly) * Const 1 = P" by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly_def) subsection "Coefficient Lemmas" lemma coeff_zero[simp]: "MPoly_Type.coeff 0 x = 0" by transfer auto lemma coeff_sum: "MPoly_Type.coeff (sum f S) x = sum (\i. MPoly_Type.coeff (f i) x) S" apply (induction S rule: infinite_finite_induct) apply (auto) by (metis More_MPoly_Type.coeff_add) lemma coeff_mult_Var: "MPoly_Type.coeff (x * Var i ^ p) c = (MPoly_Type.coeff x (c - monomial p i) when lookup c i \ p)" by transfer' (auto simp: Var\<^sub>0_def pprod.monomial_power lookup_times_monomial_right of_nat_times_monomial monomial_adds_nat_iff) lemma lookup_update_self[simp]: "Poly_Mapping.update i (lookup m i) m = m" by (auto simp: lookup_update intro!: poly_mapping_eqI) lemma coeff_Const: "MPoly_Type.coeff (Const p) m = (p when m = 0)" by transfer' (auto simp: Const\<^sub>0_def lookup_single) lemma coeff_Var: "MPoly_Type.coeff (Var p) m = (1 when m = monomial 1 p)" by transfer' (auto simp: Var\<^sub>0_def lookup_single when_def) subsection "Miscellaneous" lemma update_0_0[simp]: "Poly_Mapping.update x 0 0 = 0" by (metis lookup_update_self lookup_zero) lemma mpoly_eq_iff: "p = q \ (\m. MPoly_Type.coeff p m = MPoly_Type.coeff q m)" by transfer (auto simp: poly_mapping_eqI) lemma power_both_sides : assumes Ah : "(A::real) \0" assumes Bh : "(B::real) \0" shows "(A\B) = (A^2\B^2)" using Ah Bh by (meson power2_le_imp_le power_mono) lemma in_list_induct_helper: assumes "set(xs)\X" assumes "P []" assumes "(\x. x\X \ ( \xs. P xs \ P (x # xs)))" shows "P xs" using assms(1) proof(induction xs) case Nil then show ?case using assms by auto next case (Cons a xs) then show ?case using assms(3) by auto qed theorem in_list_induct [case_names Nil Cons]: assumes "P []" assumes "(\x. x\set(xs) \ ( \xs. P xs \ P (x # xs)))" shows "P xs" using in_list_induct_helper[of xs "set(xs)" P] using assms by auto subsubsection "Keys and vars" lemma inKeys_inVars : "a\0 \ x \ keys m \ x \ vars(MPoly_Type.monom m a)" by(simp add: vars_def) lemma notInKeys_notInVars : "x \ keys m \ x \ vars(MPoly_Type.monom m a)" by(simp add: vars_def) lemma lookupNotIn : "x \ keys m \ lookup m x = 0" by (simp add: in_keys_iff) subsection "Degree Lemmas" lemma degree_le_iff: "MPoly_Type.degree p x \ k \ (\m\monomials p. lookup m x \ k)" by transfer simp lemma degree_less_iff: "MPoly_Type.degree p x < k \ (k>0 \ (\m\monomials p. lookup m x < k))" by (transfer fixing: k) (cases "k = 0"; simp) lemma degree_ge_iff: "k > 0 \ MPoly_Type.degree p x \ k \ (\m\monomials p. lookup m x \ k)" using Max_ge_iff by (meson degree_less_iff not_less) lemma degree_greater_iff: "MPoly_Type.degree p x > k \ (\m\monomials p. lookup m x > k)" by transfer' (auto simp: Max_gr_iff) lemma degree_eq_iff: "MPoly_Type.degree p x = k \ (if k = 0 then (\m\monomials p. lookup m x = 0) else (\m\monomials p. lookup m x = k) \ (\m\monomials p. lookup m x \ k))" by (subst eq_iff) (force simp: degree_le_iff degree_ge_iff intro!: antisym) declare poly_mapping.lookup_inject[simp del] lemma lookup_eq_and_update_lemma: "lookup m var = deg \ Poly_Mapping.update var 0 m = x \ m = Poly_Mapping.update var deg x \ lookup x var = 0" unfolding poly_mapping_eq_iff by (force simp: update.rep_eq fun_eq_iff) lemma degree_const : "MPoly_Type.degree (Const (z::real)) (x::nat) = 0" by (simp add: degree_eq_iff monomials_Const) lemma degree_one : "MPoly_Type.degree (Var x :: real mpoly) x = 1" by(simp add: degree_eq_iff monomials_Var) subsection "Lemmas on treating a multivariate polynomial as univariate " lemma coeff_isolate_variable_sparse: "MPoly_Type.coeff (isolate_variable_sparse p var deg) x = (if lookup x var = 0 then MPoly_Type.coeff p (Poly_Mapping.update var deg x) else 0)" apply (transfer fixing: x var deg p) unfolding lookup_sum unfolding lookup_single apply (auto simp: when_def) apply (subst sum.inter_filter[symmetric]) subgoal by simp subgoal by (simp add: lookup_eq_and_update_lemma Collect_conv_if) by (auto intro!: sum.neutral simp add: lookup_update) lemma isovarspar_sum: "isolate_variable_sparse (p+q) var deg = isolate_variable_sparse (p) var deg + isolate_variable_sparse (q) var deg" apply (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse ) apply (metis More_MPoly_Type.coeff_add coeff_isolate_variable_sparse) by (metis More_MPoly_Type.coeff_add add.comm_neutral coeff_isolate_variable_sparse less_numeral_extra(3)) lemma isolate_zero[simp]: "isolate_variable_sparse 0 var n = 0" by transfer' (auto simp: mpoly_eq_iff) lemma coeff_isolate_variable_sparse_minus_monomial: "MPoly_Type.coeff (isolate_variable_sparse mp var i) (m - monomial i var) = (if lookup m var \ i then MPoly_Type.coeff mp (Poly_Mapping.update var i m) else 0)" by (simp add: coeff_isolate_variable_sparse lookup_minus update_minus_monomial) lemma sum_over_zero: "(mp::real mpoly) = (\i::nat \MPoly_Type.degree mp x. isolate_variable_sparse mp x i * Var x^i)" by (auto simp add: mpoly_eq_iff coeff_sum coeff_mult_Var if_if_eq_conj not_le coeff_isolate_variable_sparse_minus_monomial when_def lookup_update degree_less_iff simp flip: eq_iff intro!: coeff_not_in_monomials) lemma const_lookup_zero : "isolate_variable_sparse (Const p ::real mpoly) (x::nat) (0::nat) = Const p" by (auto simp: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Const when_def) (metis lookup_update_self) lemma const_lookup_suc : "isolate_variable_sparse (Const p :: real mpoly) x (Suc i) = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Const when_def) by (metis lookup_update lookup_zero nat.distinct(1)) lemma isovar_greater_degree : "\i > MPoly_Type.degree p var. isolate_variable_sparse p var i = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse degree_less_iff) by (metis coeff_not_in_monomials less_irrefl_nat lookup_update) lemma isolate_var_0 : "isolate_variable_sparse (Var x::real mpoly) x 0 = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Var when_def) by (metis gr0I lookup_single_eq lookup_update_self n_not_Suc_n) lemma isolate_var_one : "isolate_variable_sparse (Var x :: real mpoly) x 1 = 1" by (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Var coeff_eq_zero_iff) (smt More_MPoly_Type.coeff_monom One_nat_def add_diff_cancel_left' lookup_eq_and_update_lemma lookup_single_eq lookup_update_self monom_one plus_1_eq_Suc single_diff single_zero update_minus_monomial) lemma isovarspase_monom : assumes notInKeys : "x \ keys m" assumes notZero : "a \ 0" shows "isolate_variable_sparse (MPoly_Type.monom m a) x 0 = (MPoly_Type.monom m a :: real mpoly)" using assms by (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_eq_zero_iff in_keys_iff) (metis lookup_update_self) lemma isolate_variable_spase_zero : "lookup m x = 0 \ insertion (nth L) ((MPoly_Type.monom m a)::real mpoly) = 0 \ a \ 0 \ insertion (nth L) (isolate_variable_sparse (MPoly_Type.monom m a) x 0) = 0" by (simp add: isovarspase_monom lookup_eq_zero_in_keys_contradict notInKeys_notInVars) lemma isovarsparNotIn : "x \ vars (p::real mpoly) \ isolate_variable_sparse p x 0 = p" proof(induction p rule: mpoly_induct) case (monom m a) then show ?case apply(cases "a=0") by (simp_all add: isovarspase_monom vars_monom_keys) next case (sum p1 p2 m a) then show ?case by (simp add: monomials.rep_eq vars_add_monom isovarspar_sum) qed lemma degree0isovarspar : assumes deg0 : "MPoly_Type.degree (p::real mpoly) x = 0" shows "isolate_variable_sparse p x 0 = p" proof - have h1 : "p = (\i::nat \MPoly_Type.degree p x. isolate_variable_sparse p x i * Var x ^ i)" using sum_over_zero by auto have h2a : "\f. (\i::nat \0. f i) = f 0" apply(simp add: sum_def) by (metis add.right_neutral comm_monoid_add_class.sum_def finite.emptyI insert_absorb insert_not_empty sum.empty sum.insert) have h2 : "p = isolate_variable_sparse p x 0 * Var x ^ 0" using deg0 h1 h2a by auto show ?thesis using h2 by auto qed subsection "Summation Lemmas" lemma summation_normalized : assumes nonzero : "(B ::real) \0" shows "(\i = 0..<((n::nat)+1). (f i :: real) * B ^ (n - i)) = (\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n)" proof - have h1a : "\x::real. ((\i = 0..<(n+1). (f i) / (B ^ i)) * x = (\i = 0..<(n+1). ((f i) / (B ^ i)) * x))" apply(induction n ) apply(auto) by (simp add: distrib_right) have h1 : "(\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n) = (\i = 0..<(n+1). ((f i) / (B ^ i)) * (B^n))" using h1a by auto have h2 : "(\i = 0..<(n+1). ((f i) / (B ^ i)) * (B^n)) = (\i = 0..<(n+1). (f i) * ((B^n) / (B ^ i)))" by auto have h3 : "(\i = 0..<(n+1). (f i) * ((B^n) / (B ^ i))) = (\i = 0..<(n+1). (f i) * B ^ (n - i))" using add.right_neutral nonzero power_diff by fastforce show ?thesis using h1 h2 h3 by auto qed lemma normalize_summation : assumes nonzero : "(B::real)\0" shows "(\i = 0.. (\i = 0..<(n::nat)+1. (f i::real) / (B ^ i)) = 0" proof - have pow_zero : "\i. B^(i :: nat)\0" using nonzero by(auto) have single_division_zero : "\X. B*(X::real)=0 \ X=0" using nonzero by(auto) have h1: "(\i = 0.. ((\i = 0..i = 0..i = 0..0" shows "(\i = 0..<(n+1). (f i) * B ^ (n - i)) * B ^ (n mod 2) < 0 \ (\i = 0..<((n::nat)+1). (f i::real) / (B ^ i)) < 0" proof - have h1 : "(\i = 0..<(n+1). (f i) * B ^ (n - i)) * B ^ (n mod 2) < 0 \ (\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n) * B ^ (n mod 2) < 0" using summation_normalized nonzero by auto have h2a : "n mod 2 = 0 \ n mod 2 = 1" by auto have h2b : "n mod 2 = 1 \ odd n" by auto have h2c : "(B^n) * B ^ (n mod 2) > 0" using h2a h2b apply auto using nonzero apply presburger by (metis even_Suc mult.commute nonzero power_Suc zero_less_power_eq) have h2 : "\x. ((x * (B^n) * B ^ (n mod 2) < 0) = (x<0))" using h2c using mult.assoc by (metis mult_less_0_iff not_square_less_zero) show ?thesis using h1 h2 by auto qed subsection "Additional Lemmas with Vars" lemma not_in_isovarspar : "isolate_variable_sparse (p::real mpoly) var x = (q::real mpoly) \ var\(vars q)" apply(simp add: isolate_variable_sparse_def vars_def) proof - assume a1: "MPoly (\m | m \ monomials p \ lookup m var = x. monomial (MPoly_Type.coeff p m) (Poly_Mapping.update var 0 m)) = q" { fix pp :: "nat \\<^sub>0 nat" have "isolate_variable_sparse p var x = q" using a1 isolate_variable_sparse.abs_eq by blast then have "var \ keys pp \ pp \ keys (mapping_of q)" by (metis (no_types) coeff_def coeff_isolate_variable_sparse in_keys_iff) } then show "\p\keys (mapping_of q). var \ keys p" by meson qed lemma not_in_add : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p+q))" by (meson UnE in_mono vars_add) lemma not_in_mult : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p*q))" by (meson UnE in_mono vars_mult) lemma not_in_neg : "var\(vars(p::real mpoly)) \ var\(vars(-p))" proof - have h: "var \ (vars (-1::real mpoly))" by (metis add_diff_cancel_right' add_neg_numeral_special(8) isolate_var_one isolate_zero isovarsparNotIn isovarspar_sum not_in_isovarspar) show ?thesis using not_in_mult using h by fastforce qed lemma not_in_sub : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p-q))" using not_in_add not_in_neg by fastforce lemma not_in_pow : "var\(vars(p::real mpoly)) \ var\(vars(p^i))" proof (induction i) case 0 then show ?case using isolate_var_one not_in_isovarspar by (metis power_0) next case (Suc i) then show ?case using not_in_mult by simp qed lemma not_in_sum_var: "(\i. var\(vars(f(i)::real mpoly))) \ var\(vars(\i\{0..<(n::nat)}.f(i)))" proof (induction n) case 0 then show ?case using isolate_zero not_in_isovarspar by fastforce next case (Suc n) have h1: "(sum f {0.. vars (f n)" by (simp add: Suc.prems) then show ?case using h1 vars_add by (simp add: Suc.IH Suc.prems not_in_add) qed lemma not_in_sum : "(\i. var\(vars(f(i)::real mpoly))) \ \(n::nat). var\(vars(\i\{0..x\keys (mapping_of p). var \ keys x \ (\k f. (k \ keys f) = (lookup f k = 0)) \ lookup (mapping_of p) a = 0 \ (\aa. (if aa < length L then L[var := y] ! aa else 0) ^ lookup a aa) = (\aa. (if aa < length L then L[var := x] ! aa else 0) ^ lookup a aa)" apply(cases "lookup (mapping_of p) a = 0") apply auto apply(rule Prod_any.cong) apply auto using lookupNotIn nth_list_update_neq power_0 by smt lemma not_contains_insertion : assumes notIn : "var \ vars (p:: real mpoly)" assumes exists : "insertion (nth_default 0 (list_update L var x)) p = val" shows "insertion (nth_default 0 (list_update L var y)) p = val" using notIn exists apply(simp add: insertion_def insertion_aux_def insertion_fun_def) unfolding vars_def nth_default_def using not_in_keys_iff_lookup_eq_zero apply auto apply(rule Sum_any.cong) apply simp using not_contains_insertion_helper[of p var _ L y x] proof - fix a :: "nat \\<^sub>0 nat" assume a1: "\x\keys (mapping_of p). var \ keys x" assume "\k f. ((k::'a) \ keys f) = (lookup f k = 0)" then show "lookup (mapping_of p) a = 0 \ (\n. (if n < length L then L[var := y] ! n else 0) ^ lookup a n) = (\n. (if n < length L then L[var := x] ! n else 0) ^ lookup a n)" using a1 \\a. \\x\keys (mapping_of p). var \ keys x; \k f. (k \ keys f) = (lookup f k = 0)\ \ lookup (mapping_of p) a = 0 \ (\aa. (if aa < length L then L[var := y] ! aa else 0) ^ lookup a aa) = (\aa. (if aa < length L then L[var := x] ! aa else 0) ^ lookup a aa)\ by blast qed subsection "Insertion Lemmas" lemma insertion_sum_var : "((insertion f (\i\{0..<(n::nat)}.g(i))) = (\i\{0..(n::nat). ((insertion f (\i\{0..i\{0..(n::nat). ((insertion f (\i\n. g(i))) = (\i\n. insertion f (g i)))" by (metis (no_types, lifting) fun_sum_commute insertion_add insertion_zero sum.cong) lemma insertion_pow : "insertion f (p^i) = (insertion f p)^i" proof (induction i) case 0 then show ?case by auto next case (Suc n) then show ?case by (simp add: insertion_mult) qed lemma insertion_neg : "insertion f (-p) = -insertion f p" by (metis add.inverse_inverse insertionNegative) lemma insertion_var : "length L > var \ insertion (nth_default 0 (list_update L var x)) (Var var) = x" by (auto simp: monomials_Var coeff_Var insertion_code nth_default_def) lemma insertion_var_zero : "insertion (nth_default 0 (x#xs)) (Var 0) = x" using insertion_var by fastforce lemma notIn_insertion_sub : "x\vars(p::real mpoly) \ x\vars(q::real mpoly) \ insertion f (p-q) = insertion f p - insertion f q" by (metis ab_group_add_class.ab_diff_conv_add_uminus insertion_add insertion_neg) lemma insertion_sub : "insertion f (A-B :: real mpoly) = insertion f A - insertion f B" using insertion_neg insertion_add by (metis uminus_add_conv_diff) lemma insertion_four : "insertion ((nth_default 0) xs) 4 = 4" by (metis (no_types, lifting) insertion_add insertion_one numeral_plus_numeral one_add_one semiring_norm(2) semiring_norm(6)) lemma insertion_add_ind_basecase: "insertion (nth (list_update L var x)) ((\i::nat \ 0. isolate_variable_sparse p var i * (Var var)^i)) = (\i = 0..<(0+1). insertion (nth (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" proof - have h1: "((\i::nat \ 0. isolate_variable_sparse p var i * (Var var)^i)) = (isolate_variable_sparse p var 0 * (Var var)^0)" by auto show ?thesis using h1 by auto qed lemma insertion_add_ind: "insertion (nth_default 0 (list_update L var x)) ((\i::nat \ d. isolate_variable_sparse p var i * (Var var)^i)) = (\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" proof (induction d) case 0 then show ?case using insertion_add_ind_basecase nth_default_def by auto next case (Suc n) then show ?case using insertion_add apply auto by (simp add: insertion_add) qed lemma sum_over_degree_insertion : assumes lLength : "length L > var" assumes deg : "MPoly_Type.degree (p::real mpoly) var = d" shows "(\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i) * (x^i)) = insertion (nth_default 0 (list_update L var x)) p" proof - have h1: "(p::real mpoly) = (\i::nat \(MPoly_Type.degree p var). isolate_variable_sparse p var i * (Var var)^i)" using sum_over_zero by auto have h2: "insertion (nth_default 0 (list_update L var x)) p = insertion (nth_default 0 (list_update L var x)) ((\i::nat \ d. isolate_variable_sparse p var i * (Var var)^i))" using h1 deg by auto have h3: "insertion (nth_default 0 (list_update L var x)) p = (\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" using h2 insertion_add_ind nth_default_def by (simp add: ) show ?thesis using h3 insertion_var insertion_pow by (metis (no_types, lifting) insertion_mult lLength sum.cong) qed lemma insertion_isovarspars_free : "insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse (p::real mpoly) var (i::nat)) =insertion (nth_default 0 (list_update L var y)) (isolate_variable_sparse (p::real mpoly) var (i::nat))" proof - have h : "var \ vars(isolate_variable_sparse (p::real mpoly) var (i::nat))" by (simp add: not_in_isovarspar) then show ?thesis using not_contains_insertion by blast qed lemma insertion_zero : "insertion f (Const 0 ::real mpoly) = 0" by (metis add_cancel_right_right add_zero insertion_zero) lemma insertion_one : "insertion f (Const 1 ::real mpoly) = 1" by (metis insertion_one mult.right_neutral mult_one_left) lemma insertion_const : "insertion f (Const c::real mpoly) = (c::real)" by (auto simp: monomials_Const coeff_Const insertion_code) subsection "Putting Things Together" subsubsection "More Degree Lemmas" lemma degree_add_leq : assumes h1 : "MPoly_Type.degree a var \ x" assumes h2 : "MPoly_Type.degree b var \ x" shows "MPoly_Type.degree (a+b) var \ x" using degree_eq_iff coeff_add coeff_not_in_monomials by (smt (z3) More_MPoly_Type.coeff_add add.left_neutral coeff_eq_zero_iff degree_le_iff h1 h2) lemma degree_add_less : assumes h1 : "MPoly_Type.degree a var < x" assumes h2 : "MPoly_Type.degree b var < x" shows "MPoly_Type.degree (a+b) var < x" proof - obtain pp :: "nat \ nat \ 'a mpoly \ nat \\<^sub>0 nat" where "\x0 x1 x2. (\v3. v3 \ monomials x2 \ \ lookup v3 x1 < x0) = (pp x0 x1 x2 \ monomials x2 \ \ lookup (pp x0 x1 x2) x1 < x0)" by moura then have f1: "\m n na. (\ MPoly_Type.degree m n < na \ 0 < na \ (\p. p \ monomials m \ lookup p n < na)) \ (MPoly_Type.degree m n < na \ \ 0 < na \ pp na n m \ monomials m \ \ lookup (pp na n m) n < na)" by (metis (no_types) degree_less_iff) then have "0 < x \ (\p. p \ monomials a \ lookup p var < x)" using assms(1) by presburger then show ?thesis using f1 by (metis MPolyExtension.coeff_add add.left_neutral assms(2) coeff_eq_zero_iff) qed lemma degree_sum : "(\i\{0..n::nat}. MPoly_Type.degree (f i :: real mpoly) var \ x) \ (MPoly_Type.degree (\x\{0..n}. f x) var) \ x" proof(induction n) case 0 then show ?case by auto next case (Suc n) then show ?case by(simp add: degree_add_leq) qed lemma degree_sum_less : "(\i\{0..n::nat}. MPoly_Type.degree (f i :: real mpoly) var < x) \ (MPoly_Type.degree (\x\{0..n}. f x) var) < x" proof(induction n) case 0 then show ?case by auto next case (Suc n) then show ?case by(simp add: degree_add_less) qed lemma varNotIn_degree : assumes "var \ vars p" shows "MPoly_Type.degree p var = 0" proof- have "\m\monomials p. lookup m var = 0" using assms unfolding vars_def keys_def using monomials.rep_eq by fastforce then show ?thesis using degree_less_iff by blast qed lemma degree_mult_leq : assumes pnonzero: "(p::real mpoly)\0" assumes qnonzero: "(q::real mpoly)\0" shows "MPoly_Type.degree ((p::real mpoly) * (q::real mpoly)) var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var)" proof(cases "MPoly_Type.degree (p*q) var =0") case True then show ?thesis by simp next case False have hp: "\m\monomials p. lookup m var \ MPoly_Type.degree p var" using degree_eq_iff by (metis zero_le) have hq: "\m\monomials q. lookup m var \ MPoly_Type.degree q var" using degree_eq_iff by (metis zero_le) have hpq: "\m\{a+b | a b . a \ monomials p \ b \ monomials q}. lookup m var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var)" by (smt add_le_mono hp hq mem_Collect_eq plus_poly_mapping.rep_eq) have h1: "(\m\monomials (p*q). lookup m var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var))" using mult_monomials_dir_one hpq by blast then show ?thesis using h1 degree_eq_iff False by (simp add: degree_le_iff) qed lemma degree_exists_monom: assumes "p\0" shows "\m\monomials p. lookup m var = MPoly_Type.degree p var" proof(cases "MPoly_Type.degree p var =0") case True have h1: "\m\monomials p. lookup m var = 0" unfolding monomials_def by (metis True assms(1) aux degree_eq_iff in_keys_iff mapping_of_inject monomials.rep_eq monomials_def zero_mpoly.rep_eq) then show ?thesis using h1 using True by simp next case False then show ?thesis using degree_eq_iff assms(1) apply(auto) by (metis degree_eq_iff dual_order.strict_iff_order) qed lemma degree_not_exists_monom: assumes "p\0" shows "\ (\ m\monomials p. lookup m var > MPoly_Type.degree p var)" proof - show ?thesis using degree_less_iff by blast qed lemma degree_monom: "MPoly_Type.degree (MPoly_Type.monom x y) v = (if y = 0 then 0 else lookup x v)" by (auto simp: degree_eq_iff) lemma degree_plus_disjoint: "MPoly_Type.degree (p + MPoly_Type.monom m c) v = max (MPoly_Type.degree p v) (MPoly_Type.degree (MPoly_Type.monom m c) v)" if "m \ monomials p" for p::"real mpoly" using that apply (subst degree_eq_iff) apply (auto simp: monomials_add_disjoint) apply (auto simp: degree_eq_iff degree_monom) apply (metis MPoly_Type.degree_zero degree_exists_monom less_numeral_extra(3)) using degree_le_iff apply blast using degree_eq_iff apply (metis max_def neq0_conv) apply (metis degree_eq_iff max.coboundedI1 neq0_conv) apply (metis MPoly_Type.degree_zero degree_exists_monom max_def zero_le) using degree_le_iff max.cobounded1 by blast subsubsection "More isolate\\_variable\\_sparse lemmas" lemma isolate_variable_sparse_ne_zeroD: "isolate_variable_sparse mp v x \ 0 \ x \ MPoly_Type.degree mp v" using isovar_greater_degree leI by blast context includes poly.lifting begin lift_definition mpoly_to_nested_poly::"'a::comm_monoid_add mpoly \ nat \ 'a mpoly Polynomial.poly" is "\(mp::'a mpoly) (v::nat) (p::nat). isolate_variable_sparse mp v p" \ \note \<^const>\extract_var\ nests the other way around\ unfolding MOST_iff_cofinite proof - fix mp::"'a mpoly" and v::nat have "{p. isolate_variable_sparse mp v p \ 0} \ {0..MPoly_Type.degree mp v}" (is "?s \ _") by (auto dest!: isolate_variable_sparse_ne_zeroD) also have "finite \" by simp finally (finite_subset) show "finite ?s" . qed lemma degree_eq_0_mpoly_to_nested_polyI: "mpoly_to_nested_poly mp v = 0 \ MPoly_Type.degree mp v = 0" apply transfer' apply (simp add: degree_eq_iff) apply transfer' apply (auto simp: fun_eq_iff) proof - fix mpa :: "'a mpoly" and va :: nat and m :: "nat \\<^sub>0 nat" assume a1: "\x. (\m | m \ monomials mpa \ lookup m va = x. monomial (MPoly_Type.coeff mpa m) (Poly_Mapping.update va 0 m)) = 0" assume a2: "m \ monomials mpa" have f3: "\m p. lookup (mapping_of m) p \ (0::'a) \ p \ monomials m" by (metis (full_types) coeff_def coeff_eq_zero_iff) have f4: "\n. mapping_of (isolate_variable_sparse mpa va n) = 0" using a1 by (simp add: isolate_variable_sparse.rep_eq) have "\p n. lookup 0 (p::nat \\<^sub>0 nat) = (0::'a) \ (0::nat) = n" by simp then show "lookup m va = 0" using f4 f3 a2 by (metis coeff_def coeff_isolate_variable_sparse lookup_eq_and_update_lemma) qed lemma coeff_eq_zero_mpoly_to_nested_polyD: "mpoly_to_nested_poly mp v = 0 \ MPoly_Type.coeff mp 0 = 0" apply transfer' apply transfer' by (metis (no_types) coeff_def coeff_isolate_variable_sparse isolate_variable_sparse.rep_eq lookup_zero update_0_0) lemma mpoly_to_nested_poly_eq_zero_iff[simp]: "mpoly_to_nested_poly mp v = 0 \ mp = 0" apply (auto simp: coeff_eq_zero_mpoly_to_nested_polyD degree_eq_0_mpoly_to_nested_polyI) proof - show "mpoly_to_nested_poly mp v = 0 \ mp = 0" apply (frule degree_eq_0_mpoly_to_nested_polyI) apply (frule coeff_eq_zero_mpoly_to_nested_polyD) apply (transfer' fixing: mp v) apply (transfer' fixing: mp v) apply (auto simp: fun_eq_iff mpoly_eq_iff intro!: sum.neutral) proof - fix m :: "nat \\<^sub>0 nat" assume a1: "\x. (\m | m \ monomials mp \ lookup m v = x. monomial (MPoly_Type.coeff mp m) (Poly_Mapping.update v 0 m)) = 0" assume a2: "MPoly_Type.degree mp v = 0" have "\n. isolate_variable_sparse mp v n = 0" using a1 by (simp add: isolate_variable_sparse.abs_eq zero_mpoly.abs_eq) then have f3: "\p. MPoly_Type.coeff mp p = MPoly_Type.coeff 0 p \ lookup p v \ 0" by (metis (no_types) coeff_isolate_variable_sparse lookup_update_self) then show "MPoly_Type.coeff mp m = 0" using a2 coeff_zero by (metis coeff_not_in_monomials degree_eq_iff) qed show "mp = 0 \ mpoly_to_nested_poly 0 v = 0" subgoal apply transfer' apply transfer' by (auto simp: fun_eq_iff intro!: sum.neutral) done qed lemma isolate_variable_sparse_degree_eq_zero_iff: "isolate_variable_sparse p v (MPoly_Type.degree p v) = 0 \ p = 0" apply (transfer') apply auto proof - fix pa :: "'a mpoly" and va :: nat assume "(\m | m \ monomials pa \ lookup m va = MPoly_Type.degree pa va. monomial (MPoly_Type.coeff pa m) (Poly_Mapping.update va 0 m)) = 0" then have "mapping_of (isolate_variable_sparse pa va (MPoly_Type.degree pa va)) = 0" by (simp add: isolate_variable_sparse.rep_eq) then show "pa = 0" by (metis (no_types) coeff_def coeff_eq_zero_iff coeff_isolate_variable_sparse degree_exists_monom lookup_eq_and_update_lemma lookup_zero) qed lemma degree_eq_univariate_degree: "MPoly_Type.degree p v = (if p = 0 then 0 else Polynomial.degree (mpoly_to_nested_poly p v))" apply auto apply (rule antisym) subgoal apply (rule Polynomial.le_degree) apply (auto simp: ) apply transfer' by (simp add: isolate_variable_sparse_degree_eq_zero_iff) subgoal apply (rule Polynomial.degree_le) apply (auto simp: elim!: degree_eq_zeroE) apply transfer' by (simp add: isovar_greater_degree) done lemma compute_mpoly_to_nested_poly[code]: "coeffs (mpoly_to_nested_poly mp v) = (if mp = 0 then [] else map (isolate_variable_sparse mp v) [0.. lookup m v \ i then 0 else MPoly_Type.monom (Poly_Mapping.update v 0 m) a)" proof - have *: "{x. x = m \ lookup x v = i} = (if lookup m v = i then {m} else {})" by auto show ?thesis by (transfer' fixing: m a v i) (simp add:*) qed lemma isolate_variable_sparse_monom_mult: "isolate_variable_sparse (MPoly_Type.monom m a * q) v n = (if n \ lookup m v then MPoly_Type.monom (Poly_Mapping.update v 0 m) a * isolate_variable_sparse q v (n - lookup m v) else 0)" for q::"'a::semiring_no_zero_divisors mpoly" apply (auto simp: MPoly_Type.mult_monom) subgoal apply transfer' subgoal for mon v i a q apply (auto simp add: monomials_monom_mult sum_distrib_left) apply (rule sum.reindex_bij_witness_not_neutral[where j="\a. a - mon" and i="\a. mon + a" and S'="{}" and T'="{}" ]) apply (auto simp: lookup_add) apply (auto simp: poly_mapping_eq_iff fun_eq_iff single.rep_eq Poly_Mapping.mult_single) apply (auto simp: when_def More_MPoly_Type.coeff_monom_mult) by (auto simp: lookup_update lookup_add split: if_splits) done subgoal apply transfer' apply (auto intro!: sum.neutral simp: monomials_monom_mult ) apply (rule poly_mapping_eqI) apply (auto simp: lookup_single when_def) by (simp add: lookup_add) done lemma isolate_variable_sparse_mult: "isolate_variable_sparse (p * q) v n = (\i\n. isolate_variable_sparse p v i * isolate_variable_sparse q v (n - i))" for p q::"'a::semiring_no_zero_divisors mpoly" proof (induction p rule: mpoly_induct) case (monom m a) then show ?case by (cases "a = 0") (auto simp add: mpoly_eq_iff coeff_sum coeff_mult if_conn if_distrib if_distribR isolate_variable_sparse_monom isolate_variable_sparse_monom_mult cong: if_cong) next case (sum p1 p2 m a) then show ?case by (simp add: distrib_right isovarspar_sum sum.distrib) qed subsubsection "More Miscellaneous" lemma var_not_in_Const : "var\vars (Const x :: real mpoly)" unfolding vars_def keys_def by (smt UN_iff coeff_def coeff_isolate_variable_sparse const_lookup_zero keys_def lookup_eq_zero_in_keys_contradict) lemma mpoly_to_nested_poly_mult: "mpoly_to_nested_poly (p * q) v = mpoly_to_nested_poly p v * mpoly_to_nested_poly q v" for p q::"'a::{comm_semiring_0, semiring_no_zero_divisors} mpoly" by (auto simp: poly_eq_iff coeff_mult mpoly_to_nested_poly.rep_eq isolate_variable_sparse_mult) lemma get_if_const_insertion : assumes "get_if_const (p::real mpoly) = Some r" shows "insertion f p = r" proof- have "Set.is_empty (vars p)" apply(cases "Set.is_empty (vars p)") apply(simp) using assms by(simp) then have "(MPoly_Type.coeff p 0) = r" using assms by simp then show ?thesis by (metis Set.is_empty_def \Set.is_empty (vars p)\ empty_iff insertion_irrelevant_vars insertion_trivial) qed subsubsection "Yet more Degree Lemmas" lemma degree_mult: fixes p q::"'a::{comm_semiring_0, ring_1_no_zero_divisors} mpoly" assumes "p \ 0" assumes "q \ 0" shows "MPoly_Type.degree (p * q) v = MPoly_Type.degree p v + MPoly_Type.degree q v" using assms by (auto simp add: degree_eq_univariate_degree mpoly_to_nested_poly_mult Polynomial.degree_mult_eq) lemma degree_eq: assumes "(p::real mpoly) = (q:: real mpoly)" shows "MPoly_Type.degree p x = MPoly_Type.degree q x" by (simp add: assms) lemma degree_var_i : "MPoly_Type.degree (((Var x)^i:: real mpoly)) x = i" proof (induct i) case 0 then show ?case using degree_const by simp next case (Suc i) have multh: "(Var x)^(Suc i) = ((Var x)^i::real mpoly) * (Var x:: real mpoly)" using power_Suc2 by blast have deg0h: "MPoly_Type.degree 0 x = 0" by simp have deg1h: "MPoly_Type.degree (Var x::real mpoly) x = 1" using degree_one by blast have nonzeroh1: "(Var x:: real mpoly) \ 0" using degree_eq deg0h deg1h by auto have nonzeroh2: "((Var x)^i:: real mpoly) \ 0" using degree_eq deg0h Suc.hyps by (metis one_neq_zero power_0) have sumh: "(MPoly_Type.degree (((Var x)^i:: real mpoly) * (Var x:: real mpoly)) x) = (MPoly_Type.degree ((Var x)^i::real mpoly) x) + (MPoly_Type.degree (Var x::real mpoly) x)" using degree_mult[where p = "(Var x)^i::real mpoly", where q = "Var x::real mpoly"] nonzeroh1 nonzeroh2 by blast then show ?case using sumh deg1h by (metis Suc.hyps Suc_eq_plus1 multh) qed lemma degree_less_sum: assumes "MPoly_Type.degree (p::real mpoly) var = n" assumes "MPoly_Type.degree (q::real mpoly) var = m" assumes "m < n" shows "MPoly_Type.degree (p + q) var = n" proof - have h1: "n > 0" using assms(3) neq0_conv by blast have h2: "(\m\monomials p. lookup m var = n) \ (\m\monomials p. lookup m var \ n)" using degree_eq_iff assms(1) by (metis degree_ge_iff h1 order_refl) have h3: "(\m\monomials (p + q). lookup m var = n)" using h2 by (metis MPolyExtension.coeff_add add.right_neutral assms(2) assms(3) coeff_eq_zero_iff degree_not_exists_monom) have h4: "(\m\monomials (p + q). lookup m var \ n)" using h2 assms(3) assms(2) using degree_add_leq degree_le_iff dual_order.strict_iff_order by blast show ?thesis using degree_eq_iff h3 h4 by (metis assms(3) gr_implies_not0) qed lemma degree_less_sum': assumes "MPoly_Type.degree (p::real mpoly) var = n" assumes "MPoly_Type.degree (q::real mpoly) var = m" assumes "n < m" shows "MPoly_Type.degree (p + q) var = m" using degree_less_sum[OF assms(2) assms(1) assms(3)] by (simp add: add.commute) (* Result on the degree of the derivative *) lemma nonzero_const_is_nonzero: assumes "(k::real) \ 0" shows "((Const k)::real mpoly) \ 0" by (metis MPoly_Type.insertion_zero assms insertion_const) lemma degree_derivative : assumes "MPoly_Type.degree p x > 0" shows "MPoly_Type.degree p x = MPoly_Type.degree (derivative x p) x + 1" proof - define f where "f i = (isolate_variable_sparse p x (i+1) * (Var x)^(i) * (Const (i+1)))" for i define n where "n = MPoly_Type.degree p x-1" have d : "\m\monomials p. lookup m x = n+1" using n_def degree_eq_iff assms by (metis add.commute less_not_refl2 less_one linordered_semidom_class.add_diff_inverse) have h1a : "\i. MPoly_Type.degree (isolate_variable_sparse p x i) x = 0" by (simp add: not_in_isovarspar varNotIn_degree) have h1b : "\i::nat. MPoly_Type.degree ((Var x)^i:: real mpoly) x = i" using degree_var_i by auto have h1c1 : "\i. (Var(x)^(i)::real mpoly)\0" by (metis MPoly_Type.degree_zero h1b power_0 zero_neq_one) + fix i have h1c1var: "((Var x)^i:: real mpoly) \ 0" using h1c1 by blast have h1c2 : "((Const ((i::nat) + 1))::real mpoly)\0" using nonzero_const_is_nonzero by auto have h1c3: "(isolate_variable_sparse p x (n + 1)) \ 0" using d by (metis One_nat_def Suc_pred add.commute assms isolate_variable_sparse_degree_eq_zero_iff n_def not_gr_zero not_in_isovarspar plus_1_eq_Suc varNotIn_degree) have h3: "(isolate_variable_sparse p x (i+1) = 0) \ (MPoly_Type.degree (f i) x) = 0" by (simp add: f_def) have degh1: "(MPoly_Type.degree (((Const ((i::nat)+1))::real mpoly)*(Var x)^i) x) = ((MPoly_Type.degree ((Const (i+1))::real mpoly) x) + (MPoly_Type.degree ((Var x)^i:: real mpoly) x))" using h1c2 h1c1var degree_mult[where p="((Const ((i::nat)+1))::real mpoly)", where q="((Var x)^i:: real mpoly)"] by blast then have degh2: "(MPoly_Type.degree (((Const ((i::nat)+1))::real mpoly)*(Var x)^i) x) = i" using degree_var_i degree_const by simp have nonzerohyp: "(((Const ((i::nat)+1))::real mpoly)*(Var x)^i) \ 0" proof (induct i) case 0 then show ?case by (simp add: nonzero_const_is_nonzero) next case (Suc i) then show ?case using degree_eq degh2 by (metis Suc_eq_plus1 h1c1 mult_eq_0_iff nat.simps(3) nonzero_const_is_nonzero of_nat_eq_0_iff) qed have h4a1: "(isolate_variable_sparse p x (i+1) \ 0) \ (MPoly_Type.degree (isolate_variable_sparse p x (i+1) * ((Var x)^(i) * (Const (i+1)))::real mpoly) x = (MPoly_Type.degree (isolate_variable_sparse p x (i+1):: real mpoly) x) + (MPoly_Type.degree (((Const (i+1)) * (Var x)^(i))::real mpoly) x))" using nonzerohyp degree_mult[where p = "(isolate_variable_sparse p x (i+1))::real mpoly", where q = "((Const (i+1)) * (Var x)^(i)):: real mpoly", where v = "x"] by (simp add: mult.commute) have h4: "(isolate_variable_sparse p x (i+1) \ 0) \ (MPoly_Type.degree (f i) x) = i" using h4a1 f_def degh2 h1a - by (metis (no_types, hide_lams) add.left_neutral mult.commute mult.left_commute of_nat_1 of_nat_add) + by (metis (no_types, lifting) degh1 degree_const h1b mult.commute mult.left_commute of_nat_1 of_nat_add) have h5: "(MPoly_Type.degree (f i) x) \ i" using h3 h4 by auto have h6: "(MPoly_Type.degree (f n) x) = n" using h1c3 h4 by (smt One_nat_def add.right_neutral add_Suc_right degree_const degree_mult divisors_zero f_def h1a h1b h1c1 mult.commute nonzero_const_is_nonzero of_nat_1 of_nat_add of_nat_neq_0) have h7a: "derivative x p = (\i\{0..MPoly_Type.degree p x-1}. isolate_variable_sparse p x (i+1) * (Var x)^i * (Const (i+1)))" using derivative_def by auto have h7b: "(\i\{0..MPoly_Type.degree p x-1}. isolate_variable_sparse p x (i+1) * (Var x)^i * (Const (i+1))) = (\i\{0..MPoly_Type.degree p x-1}. (f i))" using f_def by (metis Suc_eq_plus1 add.commute semiring_1_class.of_nat_simps(2)) have h7c: "derivative x p = (\i\{0..MPoly_Type.degree p x-1}. (f i))" using h7a h7b by auto have h7: "derivative x p = (\i\{0..n}. (f i))" using n_def h7c by blast have h8: "n > 0 \ ((MPoly_Type.degree (\i\{0..(n-1)}. (f i)) x) < n)" proof (induct n) case 0 then show ?case by blast next case (Suc n) then show ?case using h5 degree_less_sum by (smt add_cancel_right_right atLeastAtMost_iff degree_const degree_mult degree_sum_less degree_var_i diff_Suc_1 f_def h1a le_imp_less_Suc mult.commute mult_eq_0_iff) qed have h9a: "n = 0 \ MPoly_Type.degree (\i\{0..n}. (f i)) x = n" using h6 by auto have h9b: "n > 0 \ MPoly_Type.degree (\i\{0..n}. (f i)) x = n" proof - have h9bhyp: "n > 0 \ (MPoly_Type.degree (\i\{0..n}. (f i)) x = MPoly_Type.degree ((\i\{0..(n-1)}. (f i)) + (f n)) x)" by (metis Suc_diff_1 sum.atLeast0_atMost_Suc) have h9bhyp2: "n > 0 \ ((MPoly_Type.degree ((\i\{0..(n-1)}. (f i)) + (f n)) x) = n)" using h6 h8 degree_less_sum by (simp add: add.commute) then show ?thesis using h9bhyp2 h9bhyp by linarith qed have h9: "MPoly_Type.degree (\i\{0..n}. (f i)) x = n" using h9a h9b by blast have h10: "MPoly_Type.degree (derivative x p) x = n" using h9 h7 by simp show ?thesis using h10 n_def using assms by linarith qed lemma express_poly : assumes h : "MPoly_Type.degree (p::real mpoly) var = 1 \ MPoly_Type.degree p var = 2" shows "p = (isolate_variable_sparse p var 2)*(Var var)^2 +(isolate_variable_sparse p var 1)*(Var var) +(isolate_variable_sparse p var 0)" proof- have h1a: "MPoly_Type.degree p var = 1 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var" using sum_over_zero[where mp="p",where x="var"] by auto have h1b: "MPoly_Type.degree p var = 1 \ isolate_variable_sparse p var 2 = 0" using isovar_greater_degree by (simp add: isovar_greater_degree) have h1: "MPoly_Type.degree p var = 1 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2" using h1a h1b by auto have h2a: "MPoly_Type.degree p var = 2 \ p = (\i::nat \ 2. isolate_variable_sparse p var i * Var var^i)" using sum_over_zero[where mp="p", where x="var"] by auto have h2b: "(\i::nat \ 2. isolate_variable_sparse p var i * Var var^i) = (\i::nat \ 1. isolate_variable_sparse p var i * Var var^i) + isolate_variable_sparse p var 2 * (Var var)^2" apply auto by (simp add: numerals(2)) have h2: "MPoly_Type.degree p var = 2 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2" using h2a h2b by auto have h3: "isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2 = isolate_variable_sparse p var 2 * (Var var)^2 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 0" by (simp add: add.commute) show ?thesis using h h1 h2 h3 by presburger qed lemma degree_isovarspar : "MPoly_Type.degree (isolate_variable_sparse (p::real mpoly) x i) x = 0" using not_in_isovarspar varNotIn_degree by blast end diff --git a/thys/Virtual_Substitution/GeneralVSProofs.thy b/thys/Virtual_Substitution/GeneralVSProofs.thy --- a/thys/Virtual_Substitution/GeneralVSProofs.thy +++ b/thys/Virtual_Substitution/GeneralVSProofs.thy @@ -1,2390 +1,2390 @@ theory GeneralVSProofs imports DNFUni EqualityVS VSAlgos begin fun separateAtoms :: "atomUni list \ (real * real * real) list * (real * real * real) list * (real * real * real) list * (real * real * real) list" where "separateAtoms [] = ([],[],[],[])"| "separateAtoms (EqUni p # L) = (let (a,b,c,d) = separateAtoms(L) in (p#a,b,c,d))"| "separateAtoms (LessUni p # L) = (let (a,b,c,d) = separateAtoms(L) in (a,p#b,c,d))"| "separateAtoms (LeqUni p # L) = (let (a,b,c,d) = separateAtoms(L) in (a,b,p#c,d))"| "separateAtoms (NeqUni p # L) = (let (a,b,c,d) = separateAtoms(L) in (a,b,c,p#d))" lemma separate_aEval : assumes "separateAtoms L = (a,b,c,d)" shows "(\l\set L. aEvalUni l x) = ((\(a,b,c)\set a. a*x^2+b*x+c=0) \ (\(a,b,c)\set b. a*x^2+b*x+c<0) \ (\(a,b,c)\set c. a*x^2+b*x+c\0) \ (\(a,b,c)\set d. a*x^2+b*x+c\0))" using assms proof(induction L arbitrary :a b c d) case Nil then show ?case by auto next case (Cons At L) then have Cons1 : "\a b c d. separateAtoms L = (a, b, c, d) \ (\l\set L. aEvalUni l x) = ((\a\set a. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c = 0) \ (\a\set b. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0)\ (\a\set c. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0) \ (\a\set d. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0))" " separateAtoms (At # L) = (a, b,c,d)" by auto then show ?case proof(cases At) case (LessUni p) show ?thesis proof(cases b) case Nil show ?thesis using Cons(2) unfolding LessUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' b') then have p_def : "p' = p" using Cons1(2) unfolding LessUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b',c,d)" using Cons Cons1(2) unfolding LessUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # b'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0) = ( (\a\set (b'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0)\ (case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0))" by auto have h3 : "(\l\set (LessUni p # L). aEvalUni l x) = ((\l\set (L). aEvalUni l x)\(case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0))" by auto show ?thesis unfolding Cons LessUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (EqUni p) show ?thesis proof(cases a) case Nil show ?thesis using Cons(2) unfolding EqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' a') then have p_def : "p' = p" using Cons1(2) unfolding EqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a',b,c,d)" using Cons Cons1(2) unfolding EqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c = 0) = ( (\a\set (a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c = 0)\ (case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c = 0))" by auto have h3 : "(\l\set (EqUni p # L). aEvalUni l x) = ((\l\set (L). aEvalUni l x)\(case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c = 0))" by auto show ?thesis unfolding Cons EqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (LeqUni p) then show ?thesis proof(cases c) case Nil show ?thesis using Cons(2) unfolding LeqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' a') then have p_def : "p' = p" using Cons1(2) unfolding LeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b,a',d)" using Cons Cons1(2) unfolding LeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0) = ( (\a\set (a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0)\ (case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0))" by auto have h3 : "(\l\set (LeqUni p # L). aEvalUni l x) = ((\l\set (L). aEvalUni l x)\(case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0))" by auto show ?thesis unfolding Cons LeqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (NeqUni p) then show ?thesis proof(cases d) case Nil show ?thesis using Cons(2) unfolding NeqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' a') then have p_def : "p' = p" using Cons1(2) unfolding NeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b,c,a')" using Cons Cons1(2) unfolding NeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0) = ( (\a\set (a'). case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0)\ (case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0))" by auto have h3 : "(\l\set (NeqUni p # L). aEvalUni l x) = ((\l\set (L). aEvalUni l x)\(case p of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0))" by auto show ?thesis unfolding Cons NeqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed qed qed lemma splitAtoms_negInfinity : assumes "separateAtoms L = (a,b,c,d)" shows "(\l\set L. evalUni (substNegInfinityUni l) x) = ( (\(a,b,c)\set a.(\x. \y (\(a,b,c)\set b.(\x. \y (\(a,b,c)\set c.(\x. \y0))\ (\(a,b,c)\set d.(\x. \y0)))" using assms proof(induction L arbitrary :a b c d) case Nil then show ?case by auto next case (Cons At L) then have Cons1 : "\a b c d. separateAtoms L = (a, b, c, d) \ (\l\set L. evalUni (substNegInfinityUni l) x) = ((\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" "separateAtoms (At # L) = (a, b, c, d)" by auto then show ?case proof(cases At) case (LessUni p) show ?thesis using LessUni Cons proof(induction b rule : list.induct) case Nil then have Nil : "b = []" using Cons.prems by auto show ?case using Cons(2) unfolding LessUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' b') then have p_def : "p' = p" using Cons1(2) unfolding LessUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b',c,d)" using Cons Cons1(2) unfolding LessUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # b'). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0) = ( (\a\set ( b'). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (case p of (a, ba, c) \ \x. \y2 + ba * y + c < 0))" by auto have one: "(\x. \y \x. \y2 + ba * y + c < 0)" apply(cases p) by simp have "(\l\set (LessUni p # L). evalUni (substNegInfinityUni l) x) = ((evalUni (substNegInfinityUni (LessUni p)) x)\(\l\set ( L). evalUni (substNegInfinityUni l) x))" by auto also have "... = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c < 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b'. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" unfolding infinity_evalUni[of "LessUni p" x, symmetric] Cons(3)[OF h1] LessUni one by simp finally have h3 : "(\l\set (LessUni p # L). evalUni (substNegInfinityUni l) x) = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c < 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b'. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0) )" by auto show ?case unfolding Cons LessUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (EqUni p) show ?thesis using EqUni Cons proof(induction a rule : list.induct) case Nil then have Nil : "a = []" using Cons.prems by auto show ?case using Cons(2) unfolding EqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' a') then have p_def : "p' = p" using Cons1(2) unfolding EqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a',b,c,d)" using Cons Cons1(2) unfolding EqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # a'). case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) = ( (\a\set ( a'). case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0)\ (case p of (a, ba, c) \ \x. \y2 + ba * y + c = 0))" by auto have one: "(\x. \y \x. \y2 + ba * y + c = 0)" apply(cases p) by simp have "(\l\set (EqUni p # L). evalUni (substNegInfinityUni l) x) = ((evalUni (substNegInfinityUni (EqUni p)) x)\(\l\set ( L). evalUni (substNegInfinityUni l) x))" by auto also have "... = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c = 0)\ (\a\set a'. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" unfolding infinity_evalUni[of "EqUni p" x, symmetric] Cons(3)[OF h1] EqUni one by simp finally have h3 : "(\l\set (EqUni p # L). evalUni (substNegInfinityUni l) x) = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c = 0)\ (\a\set a'. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" by auto show ?case unfolding Cons EqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (LeqUni p) show ?thesis using LeqUni Cons proof(induction c rule : list.induct) case Nil then have Nil : "c = []" using Cons.prems by auto show ?case using Cons(2) unfolding LeqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' c') then have p_def : "p' = p" using Cons1(2) unfolding LeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b,c',d)" using Cons Cons1(2) unfolding LeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # c'). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0) = ( (\a\set ( c'). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (case p of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" by auto have one: "(\x. \y \x. \y2 + ba * y + c \ 0)" apply(cases p) by simp have "(\l\set (LeqUni p # L). evalUni (substNegInfinityUni l) x) = ((evalUni (substNegInfinityUni (LeqUni p)) x)\(\l\set ( L). evalUni (substNegInfinityUni l) x))" by auto also have "... = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c'. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" unfolding infinity_evalUni[of "LeqUni p" x, symmetric] Cons(3)[OF h1] LeqUni one by simp finally have h3 : "(\l\set (LeqUni p # L). evalUni (substNegInfinityUni l) x) = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c'. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0) )" by auto show ?case unfolding Cons LeqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed next case (NeqUni p) show ?thesis using NeqUni Cons proof(induction d rule : list.induct) case Nil then have Nil : "d = []" using Cons.prems by auto show ?case using Cons(2) unfolding NeqUni separateAtoms.simps Nil apply(cases "separateAtoms L") by simp next case (Cons p' d') then have p_def : "p' = p" using Cons1(2) unfolding NeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h1 : "separateAtoms L = (a,b,c,d')" using Cons Cons1(2) unfolding NeqUni separateAtoms.simps apply(cases "separateAtoms L") by simp have h2 : "(\a\set (p # d'). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0) = ( (\a\set ( d'). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (case p of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" by auto have one: "(\x. \y \x. \y2 + ba * y + c \ 0)" apply(cases p) by simp have "(\l\set (NeqUni p # L). evalUni (substNegInfinityUni l) x) = ((evalUni (substNegInfinityUni (NeqUni p)) x)\(\l\set ( L). evalUni (substNegInfinityUni l) x))" by auto also have "... = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d'. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0))" unfolding infinity_evalUni[of "NeqUni p" x, symmetric] Cons(3)[OF h1] NeqUni one by simp finally have h3 : "(\l\set (NeqUni p # L). evalUni (substNegInfinityUni l) x) = ( (case p of (a,ba,c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set a. case a of (a, ba, c) \ \x. \y2 + ba * y + c = 0) \ (\a\set b. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0)\ (\a\set c. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0)\ (\a\set d'. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0) )" by auto show ?case unfolding Cons NeqUni p_def h2 h3 using Cons1(1)[OF h1] by auto qed qed qed lemma set_split : assumes "separateAtoms L = (eq,les,leq,neq)" shows "set L = set (map EqUni eq @ map LessUni les @ map LeqUni leq @ map NeqUni neq)" using assms proof(induction L arbitrary :eq les leq neq) case Nil then show ?case by auto next case (Cons At L) then show ?case proof(cases At) case (LessUni p) have "\les'. p#les' = les \ separateAtoms L = (eq, les', leq, neq)" using Cons(2) unfolding LessUni apply (cases "separateAtoms L") by auto then obtain les' where les' : "p#les' = les" "separateAtoms L = (eq, les', leq, neq)" by auto show ?thesis unfolding LessUni les'(1)[symmetric] using Cons(1)[OF les'(2)] by auto next case (EqUni p) have "\eq'. p#eq' = eq \ separateAtoms L = (eq', les, leq, neq)" using Cons(2) unfolding EqUni apply (cases "separateAtoms L") by auto then obtain eq' where eq' : "p#eq' = eq" "separateAtoms L = (eq', les, leq, neq)" by auto show ?thesis unfolding EqUni eq'(1)[symmetric] using Cons(1)[OF eq'(2)] by auto next case (LeqUni p) have "\leq'. p#leq' = leq \ separateAtoms L = (eq, les, leq', neq)" using Cons(2) unfolding LeqUni apply (cases "separateAtoms L") by auto then obtain leq' where leq' : "p#leq' = leq" "separateAtoms L = (eq, les, leq', neq)" by auto show ?thesis unfolding LeqUni leq'(1)[symmetric] using Cons(1)[OF leq'(2)] by auto next case (NeqUni p) have "\neq'. p#neq' = neq \ separateAtoms L = (eq, les, leq, neq')" using Cons(2) unfolding NeqUni apply (cases "separateAtoms L") by auto then obtain neq' where neq' : "p#neq' = neq" "separateAtoms L = (eq, les, leq, neq')" by auto show ?thesis unfolding NeqUni neq'(1)[symmetric] using Cons(1)[OF neq'(2)] by auto qed qed lemma set_split' : assumes "separateAtoms L = (eq,les,leq,neq)" shows "set (map P L) = set (map (P o EqUni) eq @ map (P o LessUni) les @ map (P o LeqUni) leq @ map (P o NeqUni) neq)" unfolding image_set[symmetric] set_split[OF assms] unfolding image_set map_append map_map by auto lemma split_elimVar : assumes "separateAtoms L = (eq,les,leq,neq)" shows "(\l\set L. evalUni (elimVarUni_atom L' l) x) = ((\(a',b',c')\set eq. (evalUni (elimVarUni_atom L' (EqUni(a',b',c'))) x)) \ (\(a',b',c')\set les. (evalUni (elimVarUni_atom L' (LessUni(a',b',c'))) x)) \ (\(a',b',c')\set leq. (evalUni (elimVarUni_atom L' (LeqUni(a',b',c'))) x)) \ (\(a',b',c')\set neq. (evalUni (elimVarUni_atom L' (NeqUni(a',b',c'))) x)))" proof- have c1: "(\l\set eq. evalUni (elimVarUni_atom L' (EqUni l)) x) = (\(a', b', c')\set eq. evalUni (elimVarUni_atom L' (EqUni (a', b', c'))) x)" by (metis (no_types, lifting) case_prodE case_prodI2) have c2: "(\l\set les. evalUni (elimVarUni_atom L' (LessUni l)) x) = (\(a', b', c')\set les. evalUni (elimVarUni_atom L' (LessUni (a', b', c'))) x)" by (metis (no_types, lifting) case_prodE case_prodI2) have c3: "(\l\set leq. evalUni (elimVarUni_atom L' (LeqUni l)) x) = (\(a', b', c')\set leq. evalUni (elimVarUni_atom L' (LeqUni (a', b', c'))) x)" by (metis (no_types, lifting) case_prodE case_prodI2) have c4: "(\l\set neq. evalUni (elimVarUni_atom L' (NeqUni l)) x) = (\(a', b', c')\set neq. evalUni (elimVarUni_atom L' (NeqUni (a', b', c'))) x)" by (metis (no_types, lifting) case_prodE case_prodI2) have h : "((\l\EqUni ` set eq. evalUni (elimVarUni_atom L' l) x) \ (\l\LessUni ` set les. evalUni (elimVarUni_atom L' l) x) \ (\l\LeqUni ` set leq. evalUni (elimVarUni_atom L' l) x) \ (\l\NeqUni ` set neq. evalUni (elimVarUni_atom L' l) x) ) = ((\l\set eq. evalUni (elimVarUni_atom L' (EqUni l)) x) \ (\l\set les. evalUni (elimVarUni_atom L' (LessUni l)) x) \ (\l\set leq. evalUni (elimVarUni_atom L' (LeqUni l)) x) \ (\l\set neq. evalUni (elimVarUni_atom L' (NeqUni l)) x) )" by auto then have "... = ((\(a', b', c')\set eq. evalUni (elimVarUni_atom L' (EqUni (a', b', c'))) x) \ (\(a', b', c')\set les. evalUni (elimVarUni_atom L' (LessUni (a', b', c'))) x) \ (\(a', b', c')\set leq. evalUni (elimVarUni_atom L' (LeqUni (a', b', c'))) x) \ (\(a', b', c')\set neq. evalUni (elimVarUni_atom L' (NeqUni (a', b', c'))) x))" using c1 c2 c3 c4 by auto then show ?thesis unfolding set_split[OF assms] set_append bex_Un image_set[symmetric] using case_prodE case_prodI2 by auto qed lemma split_elimvar : assumes "separateAtoms L = (eq,les,leq,neq)" shows "evalUni (elimVarUni_atom L At) x = evalUni (elimVarUni_atom ((map EqUni eq)@(map LessUni les) @ map LeqUni leq @ map NeqUni neq) At) x" proof(cases At) case (LessUni p) then show ?thesis apply(cases p) apply simp unfolding eval_list_conj_Uni set_split'[OF assms] by simp next case (EqUni p) then show ?thesis apply(cases p) apply simp unfolding eval_list_conj_Uni set_split'[OF assms] by simp next case (LeqUni p) then show ?thesis apply(cases p) apply simp unfolding eval_list_conj_Uni set_split'[OF assms] by simp next case (NeqUni p) then show ?thesis apply(cases p) apply simp unfolding eval_list_conj_Uni set_split'[OF assms] by simp qed lemma less : " ((a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. evalUni (substInfinitesimalLinearUni b' c' (EqUni (d, e, f))) x) \ (\(d, e, f)\set b. evalUni (substInfinitesimalLinearUni b' c' (LessUni (d, e, f))) x) \ (\(d, e, f)\set c. evalUni (substInfinitesimalLinearUni b' c' (LeqUni (d, e, f))) x) \ (\(d, e, f)\set d. evalUni (substInfinitesimalLinearUni b' c' (NeqUni (d, e, f))) x) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. evalUni (substInfinitesimalQuadraticUni (- b') 1 (b'\<^sup>2 - 4 * a' * c') (2 * a') (EqUni (d, e, f))) x) \ (\(d, e, f)\set b. evalUni (substInfinitesimalQuadraticUni (- b') 1 (b'\<^sup>2 - 4 * a' * c') (2 * a') (LessUni (d, e, f))) x) \ (\(d, e, f)\set c. evalUni (substInfinitesimalQuadraticUni (- b') 1 (b'\<^sup>2 - 4 * a' * c') (2 * a') (LeqUni (d, e, f))) x) \ (\(d, e, f)\set d. evalUni (substInfinitesimalQuadraticUni (- b') 1 (b'\<^sup>2 - 4 * a' * c') (2 * a') (NeqUni (d, e, f))) x) \ (\(d, e, f)\set a. evalUni (substInfinitesimalQuadraticUni (- b') (- 1) (b'\<^sup>2 - 4 * a' * c') (2 * a') (EqUni (d, e, f))) x) \ (\(d, e, f)\set b. evalUni (substInfinitesimalQuadraticUni (- b') (- 1) (b'\<^sup>2 - 4 * a' * c') (2 * a') (LessUni (d, e, f))) x) \ (\(d, e, f)\set c. evalUni (substInfinitesimalQuadraticUni (- b') (- 1) (b'\<^sup>2 - 4 * a' * c') (2 * a') (LeqUni (d, e, f))) x) \ (\(d, e, f)\set d. evalUni (substInfinitesimalQuadraticUni (- b') (- 1) (b'\<^sup>2 - 4 * a' * c') (2 * a') (NeqUni (d, e, f))) x))) = ((a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. (\y'::real>-c'/b'. \x::real \{-c'/b'<..y'}. aEvalUni (EqUni (d, e, f)) x)) \ (\(d, e, f)\set b. (\y'::real>-c'/b'. \x::real \{-c'/b'<..y'}. aEvalUni (LessUni (d, e, f)) x))\ (\(d, e, f)\set c. (\y'::real>-c'/b'. \x::real \{-c'/b'<..y'}. aEvalUni (LeqUni (d, e, f)) x)) \ (\(d, e, f)\set d. (\y'::real>-c'/b'. \x::real \{-c'/b'<..y'}. aEvalUni (NeqUni (d, e, f)) x)) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. (\y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (EqUni (d,e,f)) x)) \ (\(d, e, f)\set b. (\y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (LessUni (d,e,f)) x)) \ (\(d, e, f)\set c. (\y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (LeqUni (d,e,f)) x)) \ (\(d, e, f)\set d. (\y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (NeqUni (d,e,f)) x)) \ (\(d, e, f)\set a. (\y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (EqUni (d,e,f)) x)) \ (\(d, e, f)\set b. (\y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (LessUni (d,e,f)) x)) \ (\(d, e, f)\set c. (\y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (LeqUni (d,e,f)) x)) \ (\(d, e, f)\set d. (\y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. aEvalUni (NeqUni (d,e,f)) x))))" proof(cases "a'=0") case True then have a' : "a'=0" by auto then show ?thesis proof(cases "b'=0") case True then show ?thesis using a' by auto next case False then show ?thesis using True unfolding infinitesimal_linear'[of b' c' _ x, symmetric, OF False] by auto qed next case False then have a' : "a' \ 0" by auto then have d : "2 * a' \ 0" by auto show ?thesis proof(cases "0 \ b'\<^sup>2 - 4 * a' * c'") case True then show ?thesis using False unfolding infinitesimal_quad[OF d True, of "-b'", symmetric] by auto next case False then show ?thesis using a' by auto qed qed lemma eq_inf : "(\(a, b, c)\set (a::(real*real*real) list). \x. \y2 + b * y + c = 0) = (\(a, b, c)\set a. a=0\b=0\c=0)" using infinity_evalUni_EqUni[of _ x] by auto text "This is the main quantifier elimination lemma, in the simplified framework. We are located directly underneath the most internal existential quantifier so F is completely free in quantifier and consists only of conjunction and disjunction. The variable we are evaluating on, x, is removed in the generalVS\\_DNF converted formula as expanding out the evalUni function determines that x doesn't play a role in the computation of it. It would be okay to replace the x in the second half with any variable, but it is simplier this way This conversion is defined as a \"veritcal\" translation as we remain within the univariate framework in both sides of the expression" lemma eval_generalVS'' : "(\x. evalUni (list_conj_Uni (map AtomUni L)) x) = evalUni (generalVS_DNF L) x" proof(cases "separateAtoms L") case (fields a b c d) have a : "\ P. (\l\set (map EqUni a) \ (set (map LessUni b) \ (set (map LeqUni c) \ set (map NeqUni d))).P l) = ((\(d,e,f)\set a. P (EqUni (d,e,f))) \ (\(d,e,f)\set b. P (LessUni (d,e,f))) \ (\(d,e,f)\set c. P (LeqUni (d,e,f))) \ (\(d,e,f)\set d. P (NeqUni (d,e,f))))" by auto show ?thesis apply(simp add: eval_list_conj_Uni separate_aEval[OF fields] splitAtoms_negInfinity[OF fields] eval_list_disj_Uni del:elimVar.simps) unfolding eval_conj_atom generalVS_DNF.simps split_elimVar[OF fields ] split_elimvar[OF fields] unfolding elimVarUni_atom.simps evalUni.simps aEvalUni.simps Rings.mult_zero_class.mult_zero_left Groups.add_0 eval_list_conj_Uni Groups.group_add_class.minus_zero eval_map_all linearSubstitutionUni.simps quadraticSubUni.simps evalUni_if aEvalUni.simps set_append a less eq_inf using qe by auto qed lemma forallx_substNegInf : "(\evalUni (map_atomUni substNegInfinityUni F) x) = (\x. \ evalUni (map_atomUni substNegInfinityUni F) x)" proof(induction F) case TrueFUni then show ?case by simp next case FalseFUni then show ?case by simp next case (AtomUni At) then show ?case apply(cases At) by auto next case (AndUni F1 F2) then show ?case by auto next case (OrUni F1 F2) then show ?case by auto qed lemma linear_subst_map: "evalUni (map_atomUni (linearSubstitutionUni b c) F) x = evalUni F (-c/b)" apply(induction F)by auto lemma quadratic_subst_map : "evalUni (map_atomUni (quadraticSubUni a b c d) F) x = evalUni F ((a+b*sqrt(c))/d)" apply(induction F)by auto fun convert_atom_list :: "nat \ atom list \ real list \ (atomUni list) option" where "convert_atom_list var [] xs = Some []"| "convert_atom_list var (a#as) xs = ( case convert_atom var a xs of Some(a) \ (case convert_atom_list var as xs of Some(as) \ Some(a#as) | None \ None) | None \ None )" lemma convert_atom_list_change : assumes "length xs' = var" shows "convert_atom_list var L (xs' @ x # \) = convert_atom_list var L (xs' @ x' # \)" apply(induction L)using convert_atom_change[OF assms] apply simp_all by (metis) lemma negInf_convert : assumes "convert_atom_list var L (xs' @ x # xs) = Some L'" assumes "length xs' = var" shows "(\f\set L. eval (substNegInfinity var f) (xs' @ x # xs)) = (\f\set L'. evalUni (substNegInfinityUni f) x)" using assms proof(induction L arbitrary : L') case Nil then show ?case by auto next case (Cons a L) then show ?case proof(cases a) case (Less p) have h: "MPoly_Type.degree p var < 3 \ eval (substNegInfinity var (Less p)) (xs' @ x # xs) = evalUni (substNegInfinityUni (LessUni (insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0)), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)))) x" using convert_substNegInfinity[of var "Less p" xs' x xs, OF _ assms(2)] by simp show ?thesis using Cons(2)[symmetric] Cons(1) unfolding Less apply(cases " MPoly_Type.degree p var < 3") defer apply simp apply(cases "convert_atom_list var L (xs' @ x # xs)") apply (simp_all del: substNegInfinity.simps substNegInfinityUni.simps) unfolding h using assms(2) by presburger next case (Eq p) have h: "MPoly_Type.degree p var < 3 \ eval (substNegInfinity var (Eq p)) (xs' @ x # xs) = evalUni (substNegInfinityUni (EqUni (insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0)), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)))) x" using convert_substNegInfinity[of var "Eq p", OF _ assms(2)] by simp show ?thesis using Cons(2)[symmetric] Cons(1) unfolding Eq apply(cases " MPoly_Type.degree p var < 3") defer apply simp apply(cases "convert_atom_list var L (xs' @ x # xs)") apply (simp_all del: substNegInfinity.simps substNegInfinityUni.simps) unfolding h assms by auto next case (Leq p) have h: "MPoly_Type.degree p var < 3 \ eval (substNegInfinity var (Leq p)) (xs' @ x # xs) = evalUni (substNegInfinityUni (LeqUni (insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0)), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)))) x" using convert_substNegInfinity[of var "Leq p", OF _ assms(2)] by simp show ?thesis using Cons(2) unfolding Leq apply(cases " MPoly_Type.degree p var < 3") defer apply simp apply(cases "convert_atom_list var L (xs' @ x # xs)") apply (simp_all del: substNegInfinity.simps substNegInfinityUni.simps) unfolding h using Cons.IH assms by auto next case (Neq p) have h: "MPoly_Type.degree p var < 3 \ eval (substNegInfinity var (Neq p)) (xs' @ x # xs) = evalUni (substNegInfinityUni (NeqUni (insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0)), insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)))) x" using convert_substNegInfinity[of var "Neq p", OF _ assms(2)] by simp show ?thesis using Cons(2) unfolding Neq apply(cases " MPoly_Type.degree p var < 3") defer apply simp apply(cases "convert_atom_list var L (xs' @ x # xs)") apply (simp_all del: substNegInfinity.simps substNegInfinityUni.simps) unfolding h using Cons.IH assms by auto qed qed lemma elimVar_atom_single : assumes "convert_atom var A (xs' @ x # xs) = Some A'" assumes "convert_atom_list var L2 (xs' @ x # xs) = Some L2'" assumes "length xs' = var" shows "eval (elimVar var L2 [] A) (xs' @ x # xs) = evalUni (elimVarUni_atom L2' A') x" proof(cases A) case (Less p) define a where "a = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2)" have a_def' : "a = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 2)" unfolding a_def using insertion_isovarspars_free[of "(xs' @ x # xs)" var x p 2 0] assms by auto define b where "b = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0))" have b_def' : "b = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var (Suc 0))" unfolding b_def using insertion_isovarspars_free[of "(xs' @ x # xs)" var x p "(Suc 0)" 0] assms by auto define c where "c = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)" have c_def' : "c = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 0)" unfolding c_def using insertion_isovarspars_free[of "(xs' @ x # xs)" var x p 0 0] assms by auto have linear : "b\0 \ (\f\set L2. eval (substInfinitesimalLinear var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalLinearUni b c l) x)" using assms(2) proof(induction L2 arbitrary : L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(3) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(3) At' by (simp_all add: L2's) have h : "eval (substInfinitesimalLinear var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs) = evalUni (substInfinitesimalLinearUni b c At') x" proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by simp_all next case (Some a) have h1 : "var \ vars (isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar) have h2 : "var \ vars (isolate_variable_sparse p var 0)"by (simp add: not_in_isovarspar) have h : "evalUni (substInfinitesimalLinearUni b c a) x = evalUni (substInfinitesimalLinearUni b c At') x" proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq x3) then show ?thesis using At' Some by auto next case (Neq x4) then show ?thesis using At' Some by auto qed show ?thesis unfolding convert_substInfinitesimalLinear[OF Some b_def[symmetric] c_def[symmetric] Cons(2) h1 h2 assms(3)] using h . qed show ?case unfolding L2' using h Cons(1)[OF Cons(2) L2's] by auto qed have quadratic_1 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) 1 = 1" by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @ xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof (cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some aT) have h1 : "insertion (nth_default 0 (xs' @ x # xs)) (- isolate_variable_sparse p var (Suc 0)) = (-b)" unfolding b_def insertion_neg by auto have h2 : "insertion (nth_default 0 (xs' @ x # xs)) 1 = 1" by auto have h3 : "insertion (nth_default 0 (xs' @ x # xs)) (((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)) = (b\<^sup>2 - 4 * a * c)" unfolding insertion_mult insertion_pow insertion_four insertion_neg insertion_sub a_def b_def c_def by auto have h4 : "insertion (nth_default 0 (xs' @ x # xs)) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def by (metis insertion_add insertion_mult mult_2) have h5 : "2 * a \ 0" using Cons by auto have h6 : "0 \ b\<^sup>2 - 4 * a * c" using Cons by auto have h7 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h8 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h9 : "var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h10 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) aT) x = evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x"proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Leq x3) then show ?thesis using At' using Some by auto next case (Neq x4) then show ?thesis using At' using Some by auto qed show ?thesis unfolding convert_substInfinitesimalQuadratic[OF Some h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 assms(3)] using h . qed show ?case unfolding L2' apply(simp del : substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h by auto qed have quadratic_2 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) (- 1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalQuadraticUni (- b) (- 1) (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (-1) = (-1)" unfolding insertion_neg by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) using assms by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @ xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto using assms - by (metis (no_types, hide_lams) MPoly_Type.insertion_one add.inverse_inverse add_uminus_conv_diff arith_special(3) insertion_isovarspars_free insertion_neg insertion_sub list_update_id) + by (metis (no_types, opaque_lifting) MPoly_Type.insertion_one add.inverse_inverse add_uminus_conv_diff arith_special(3) insertion_isovarspars_free insertion_neg insertion_sub list_update_id) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(- 1::real mpoly)" by (metis h9 not_in_neg not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof (cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some aT) have h1 : "insertion (nth_default 0 (xs' @ x # xs)) (- isolate_variable_sparse p var (Suc 0)) = (-b)" unfolding b_def insertion_neg by auto have h2 : "insertion (nth_default 0 (xs' @ x # xs)) (-1) = -1" unfolding insertion_neg by auto have h3 : "insertion (nth_default 0 (xs' @ x # xs)) (((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)) = (b\<^sup>2 - 4 * a * c)" unfolding insertion_mult insertion_pow insertion_four insertion_neg insertion_sub a_def b_def c_def using assms by auto have h4 : "insertion (nth_default 0 (xs' @ x # xs)) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def by (metis insertion_add insertion_mult mult_2) have h5 : "2 * a \ 0" using Cons by auto have h6 : "0 \ b\<^sup>2 - 4 * a * c" using Cons by auto have h7 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h8 : "var\vars(- 1::real mpoly)" by (simp add: h10 not_in_neg) have h9 : "var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h10 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) aT) x = evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x"proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Leq x3) then show ?thesis using At' using Some option.inject by auto next case (Neq x4) then show ?thesis using At' using Some by auto qed show ?thesis unfolding convert_substInfinitesimalQuadratic[OF Some h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 assms(3)] using h . qed show ?case unfolding L2' apply(simp del : substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h by auto qed show ?thesis using assms(1)[symmetric] unfolding Less apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(simp del : substInfinitesimalLinear.simps substInfinitesimalLinearUni.simps substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps add: insertion_neg insertion_mult insertion_add insertion_pow insertion_sub insertion_four a_def[symmetric] b_def[symmetric] c_def[symmetric] a_def'[symmetric] b_def'[symmetric] c_def'[symmetric] eval_list_conj eval_list_conj_Uni ) using linear quadratic_1 quadratic_2 by smt next case (Eq p) define a where "a = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2)" have a_def' : "a = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 2)" unfolding a_def using insertion_isovarspars_free[of "xs' @x#xs" var x p 2 0] using assms by auto define b where "b = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0))" have b_def' : "b = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var (Suc 0))" unfolding b_def using insertion_isovarspars_free[of "xs' @x#xs" var x p "(Suc 0)" 0] using assms by auto define c where "c = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)" have c_def' : "c = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 0)" unfolding c_def using insertion_isovarspars_free[of "xs' @x#xs" var x p 0 0]using assms by auto have linear : "a=0 \ b\0 \ (\f\set L2. aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (linearSubstitutionUni b c l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var \ vars (isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar) have h2 : "var \ vars (isolate_variable_sparse p var 0)"by (simp add: not_in_isovarspar) have h : "aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs) = evalUni (linearSubstitutionUni b c At') x" proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some a) have h : "a=At'" using At' Some by auto show ?thesis unfolding convert_linearSubstitutionUni[OF Some b_def[symmetric] c_def[symmetric] Cons(3) h1 h2 assms(3)] unfolding h by auto qed have "(\f\set (At # L2). aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (aEval (linear_substitution var (-isolate_variable_sparse p var 0)(isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs)\ (\f\set (L2). aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)))" by auto also have "... = (evalUni (linearSubstitutionUni b c At') x \ (\l\set L2's. evalUni (linearSubstitutionUni b c l) x))" unfolding h Cons(1)[OF Cons(2) Cons(3) L2's] by auto finally show ?case unfolding L2' by auto qed have quadratic_1 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \(\f\set L2. eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (quadraticSubUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) 1 = 1" by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @ xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_add insertion_isovarspars_free insertion_mult list_update_length mult_2) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = aEval At (xs' @ (((- b + 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)) # xs))" using quadratic_sub[OF h1 h2 h3 h4 h5 h6 h7 h8, symmetric, of At] free_in_quad[OF h9 h10 h4 h11] by (metis assms(3) list_update_length var_not_in_eval3) have h2 : "aEval At (xs' @ (- b + 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) # xs) = evalUni (quadraticSubUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof(cases At) case (Less p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Less apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Eq p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Eq apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Leq x3) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Leq apply(cases "MPoly_Type.degree p var < 3") by auto qed next case (Neq x4) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Neq apply(cases "MPoly_Type.degree p var < 3") by auto qed qed show ?case unfolding L2' apply(simp del : quadratic_sub.simps quadraticSubUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h h2 by auto qed have quadratic_2 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) (- 1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (quadraticSubUni (- b) (- 1) (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" using assms by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (-1) = -1" unfolding insertion_neg by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(-1::real mpoly)" by (metis h9 not_in_neg not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = aEval At (xs' @ (((- b - 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)) # xs))" using quadratic_sub[OF h1 h2 h3 h4 h5 h6 h7 h8, symmetric, of At] var_not_in_eval3 free_in_quad[OF h9 h10 h4 h11] using assms(3) by fastforce have h2 : "aEval At (xs' @ (- b - 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) # xs) = evalUni (quadraticSubUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof(cases At) case (Less p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Less apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Eq p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Eq apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Leq x3) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Leq apply(cases "MPoly_Type.degree p var < 3") by auto qed next case (Neq x4) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Neq apply(cases "MPoly_Type.degree p var < 3") by auto qed qed show ?case unfolding L2' apply(simp del : quadratic_sub.simps quadraticSubUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h h2 by auto qed show ?thesis using assms(1)[symmetric] unfolding Eq apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(simp del : linearSubstitutionUni.simps quadraticSubUni.simps add: insertion_neg insertion_mult insertion_add insertion_pow insertion_sub insertion_four a_def[symmetric] b_def[symmetric] c_def[symmetric] a_def'[symmetric] b_def'[symmetric] c_def'[symmetric] eval_list_conj eval_list_conj_Uni )using linear using quadratic_1 quadratic_2 by smt next case (Leq p) define a where "a = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2)" have a_def' : "a = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 2)" unfolding a_def using insertion_isovarspars_free[of "xs'@ x#xs" var x p 2 0] using assms by auto define b where "b = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0))" have b_def' : "b = insertion (nth_default 0 (xs'@ 0 # xs)) (isolate_variable_sparse p var (Suc 0))" unfolding b_def using insertion_isovarspars_free[of "xs'@x#xs" var x p "(Suc 0)" 0] using assms by auto define c where "c = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)" have c_def' : "c = insertion (nth_default 0 (xs'@ 0 # xs)) (isolate_variable_sparse p var 0)" unfolding c_def using insertion_isovarspars_free[of "xs'@ x#xs" var x p 0 0] using assms by auto have linear : "a=0 \ b\0 \ (\f\set L2. aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (linearSubstitutionUni b c l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var \ vars (isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar) have h2 : "var \ vars (isolate_variable_sparse p var 0)"by (simp add: not_in_isovarspar) have h : "aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs) = evalUni (linearSubstitutionUni b c At') x" proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some a) have h : "a=At'" using At' Some by auto show ?thesis unfolding convert_linearSubstitutionUni[OF Some b_def[symmetric] c_def[symmetric] Cons(3) h1 h2 assms(3)] unfolding h by auto qed have "(\f\set (At # L2). aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs)\ (\f\set (L2). aEval (linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)))" by auto also have "... = (evalUni (linearSubstitutionUni b c At') x \ (\l\set L2's. evalUni (linearSubstitutionUni b c l) x))" unfolding h Cons(1)[OF Cons(2) Cons(3) L2's] by auto finally show ?case unfolding L2' by auto qed have quadratic_1 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \(\f\set L2. eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (quadraticSubUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) 1 = 1" by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = aEval At (xs' @ (((- b + 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)) # xs))" using quadratic_sub[OF h1 h2 h3 h4 h5 h6 h7 h8, symmetric, of At] var_not_in_eval3 free_in_quad[OF h9 h10 h4 h11] by (metis assms(3) list_update_length) have h2 : "aEval At (xs' @ (- b + 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) # xs) = evalUni (quadraticSubUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof(cases At) case (Less p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Less apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Eq p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Eq apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Leq x3) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Leq apply(cases "MPoly_Type.degree p var < 3") by auto qed next case (Neq x4) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Neq apply(cases "MPoly_Type.degree p var < 3") by auto qed qed show ?case unfolding L2' apply(simp del : quadratic_sub.simps quadraticSubUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h h2 by auto qed have quadratic_2 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) (- 1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (quadraticSubUni (- b) (- 1) (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (-1) = -1" unfolding insertion_neg by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @ xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(-1::real mpoly)" by (metis h9 not_in_neg not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = aEval At (xs' @(((- b - 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)) # xs))" using quadratic_sub[OF h1 h2 h3 h4 h5 h6 h7 h8, symmetric, of At] var_not_in_eval3 free_in_quad[OF h9 h10 h4 h11] using assms(3) by fastforce have h2 : "aEval At (xs' @(- b - 1 * sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) # xs) = evalUni (quadraticSubUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof(cases At) case (Less p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Less apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Eq p) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Eq apply(cases "MPoly_Type.degree p var < 3") by simp_all qed next case (Leq x3) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Leq apply(cases "MPoly_Type.degree p var < 3") by (auto) qed next case (Neq x4) then show ?thesis proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At'[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Some aT) then have Some : "\x. convert_atom var At (xs' @ x # xs) = Some aT" by (metis assms(3) convert_atom_change) show ?thesis unfolding aEval_aEvalUni[OF Some assms(3)] using At'[symmetric] Some[symmetric] unfolding Neq apply(cases "MPoly_Type.degree p var < 3") by auto qed qed show ?case unfolding L2' apply(simp del : quadratic_sub.simps quadraticSubUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h h2 by auto qed show ?thesis using assms(1)[symmetric] unfolding Leq apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(simp del : linearSubstitutionUni.simps quadraticSubUni.simps add: insertion_neg insertion_mult insertion_add insertion_pow insertion_sub insertion_four a_def[symmetric] b_def[symmetric] c_def[symmetric] a_def'[symmetric] b_def'[symmetric] c_def'[symmetric] eval_list_conj eval_list_conj_Uni ) using linear using quadratic_1 quadratic_2 by smt next case (Neq p) define a where "a = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2)" have a_def' : "a = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var 2)" unfolding a_def using insertion_isovarspars_free[of "xs' @x#xs" var x p 2 0] using assms by auto define b where "b = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0))" have b_def' : "b = insertion (nth_default 0 (xs' @ 0 # xs)) (isolate_variable_sparse p var (Suc 0))" unfolding b_def using insertion_isovarspars_free[of "xs'@x#xs" var x p "(Suc 0)" 0] using assms by auto define c where "c = insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0)" have c_def' : "c = insertion (nth_default 0 (xs'@0 # xs)) (isolate_variable_sparse p var 0)" unfolding c_def using insertion_isovarspars_free[of "xs'@x#xs" var x p 0 0] using assms by auto have linear : "b\0 \ (\f\set L2. eval (substInfinitesimalLinear var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalLinearUni b c l) x)" using assms(2) proof(induction L2 arbitrary : L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(3) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(3) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(3) At' by (simp_all add: L2's) have h : "eval (substInfinitesimalLinear var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) At) (xs' @ x # xs) = evalUni (substInfinitesimalLinearUni b c At') x" proof(cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by simp_all next case (Some a) have h1 : "var \ vars (isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar) have h2 : "var \ vars (isolate_variable_sparse p var 0)"by (simp add: not_in_isovarspar) have h : "evalUni (substInfinitesimalLinearUni b c a) x = evalUni (substInfinitesimalLinearUni b c At') x" proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq x3) then show ?thesis using At' Some by auto next case (Neq x4) then show ?thesis using At' Some by auto qed show ?thesis unfolding convert_substInfinitesimalLinear[OF Some b_def[symmetric] c_def[symmetric] Cons(2) h1 h2 assms(3)] using h . qed show ?case unfolding L2' using h Cons(1)[OF Cons(2) L2's] by auto qed have quadratic_1 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length (xs' @ x # xs)" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) 1 = 1" by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof (cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some aT) have h1 : "insertion (nth_default 0 (xs' @ x # xs)) (- isolate_variable_sparse p var (Suc 0)) = (-b)" unfolding b_def insertion_neg by auto have h2 : "insertion (nth_default 0 (xs' @ x # xs)) 1 = 1" by auto have h3 : "insertion (nth_default 0 (xs' @ x # xs)) (((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)) = (b\<^sup>2 - 4 * a * c)" unfolding insertion_mult insertion_pow insertion_four insertion_neg insertion_sub a_def b_def c_def by auto have h4 : "insertion (nth_default 0 (xs' @ x # xs)) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def by (metis insertion_add insertion_mult mult_2) have h5 : "2 * a \ 0" using Cons by auto have h6 : "0 \ b\<^sup>2 - 4 * a * c" using Cons by auto have h7 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h8 : "var\vars(1::real mpoly)" by (metis h9 not_in_pow power.simps(1)) have h9 : "var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h10 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) aT) x = evalUni (substInfinitesimalQuadraticUni (- b) 1 (b\<^sup>2 - 4 * a * c) (2 * a) At') x"proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Leq x3) then show ?thesis using At' using Some by auto next case (Neq x4) then show ?thesis using At' using Some by auto qed show ?thesis unfolding convert_substInfinitesimalQuadratic[OF Some h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 assms(3)] using h . qed show ?case unfolding L2' apply(simp del : substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h by auto qed have quadratic_2 : "(a \ 0) \ (4 * a * c \ b\<^sup>2) \ (\f\set L2. eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) (- 1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) f) (xs' @ x # xs)) = (\l\set L2'. evalUni (substInfinitesimalQuadraticUni (- b) (- 1) (b\<^sup>2 - 4 * a * c) (2 * a) l) x)" using assms(2) proof(induction L2 arbitrary: L2') case Nil then show ?case by auto next case (Cons At L2) have "\At'. convert_atom var At (xs' @ x # xs) = Some At'" proof(cases At) case (Less p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Eq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by simp_all next case (Leq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto next case (Neq p) then show ?thesis using Cons(4) apply simp apply(cases "MPoly_Type.degree p var < 3") by auto qed then obtain At' where At' : "convert_atom var At (xs' @ x # xs) = Some At'" by auto have "\L2's. convert_atom_list var L2 (xs' @ x # xs) = Some L2's" using Cons(4) At' apply(cases "convert_atom_list var L2 (xs' @ x # xs)") by auto then obtain L2's where L2's : "convert_atom_list var L2 (xs' @ x # xs) = Some L2's" by auto have L2' : "L2' = At' # L2's" using Cons(4) At' apply(cases At) apply auto by (simp_all add: L2's) have h1 : "var < length ((xs' @ x # xs))" using assms by auto have h2 : "2*a \0" using Cons by auto have h3 : "0\b^2-4*a*c" using Cons(3) by auto have h4 : "var\vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h5 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (- isolate_variable_sparse p var (Suc 0)) = -b" unfolding insertion_neg b_def by (metis insertion_isovarspars_free list_update_id) have h6 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (-1) = (-1)" unfolding insertion_neg by auto have h7 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) = b\<^sup>2 - 4 * a * c" apply(simp add: insertion_four insertion_mult insertion_sub insertion_pow b_def a_def c_def) by (metis insertion_isovarspars_free list_update_id) have "\xa. insertion (nth_default 0 (xs' @ xa # xs)) (2::real mpoly) = (2::real)" by (metis MPoly_Type.insertion_one insertion_add one_add_one) then have h8 : "\xa. insertion (nth_default 0 ((xs' @ x # xs)[var := xa])) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def apply auto by (metis assms(3) insertion_lowerPoly1 list_update_length not_in_isovarspar) have h9 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h10 : "var\vars(- 1::real mpoly)" by (metis h9 not_in_neg not_in_pow power.simps(1)) have h11 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "eval (substInfinitesimalQuadratic var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2) At) (xs' @ x # xs) = evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x" proof (cases "convert_atom var At (xs' @ x # xs)") case None then show ?thesis using At' apply(cases At) by auto next case (Some aT) have h1 : "insertion (nth_default 0 (xs' @ x # xs)) (- isolate_variable_sparse p var (Suc 0)) = (-b)" unfolding b_def insertion_neg by auto have h2 : "insertion (nth_default 0 (xs' @ x # xs)) (-1) = -1" unfolding insertion_neg by auto have h3 : "insertion (nth_default 0 (xs' @ x # xs)) (((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)) = (b\<^sup>2 - 4 * a * c)" unfolding insertion_mult insertion_pow insertion_four insertion_neg insertion_sub a_def b_def c_def by auto have h4 : "insertion (nth_default 0 (xs' @ x # xs)) (2 * isolate_variable_sparse p var 2) = 2 * a" unfolding insertion_mult a_def by (metis insertion_add insertion_mult mult_2) have h5 : "2 * a \ 0" using Cons by auto have h6 : "0 \ b\<^sup>2 - 4 * a * c" using Cons by auto have h7 : "var\vars(- isolate_variable_sparse p var (Suc 0))" by (simp add: not_in_isovarspar not_in_neg) have h8 : "var\vars(- 1::real mpoly)" by (simp add: h10 not_in_neg) have h9 : "var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" by (metis add_uminus_conv_diff not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow num_double numeral_times_numeral one_add_one power_0) have h10 : "var\vars(2 * isolate_variable_sparse p var 2)" by (metis isovarspar_sum mult_2 not_in_isovarspar) have h : "evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) aT) x = evalUni (substInfinitesimalQuadraticUni (- b) (-1) (b\<^sup>2 - 4 * a * c) (2 * a) At') x"proof(cases At) case (Less p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Eq p) then show ?thesis using At'[symmetric] Some[symmetric] apply(cases "MPoly_Type.degree p var < 3") by auto next case (Leq x3) then show ?thesis using At' using Some option.inject by auto next case (Neq x4) then show ?thesis using At' using Some by auto qed show ?thesis unfolding convert_substInfinitesimalQuadratic[OF Some h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 assms(3)] using h . qed show ?case unfolding L2' apply(simp del : substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps) unfolding Cons(1)[OF Cons(2) Cons(3) L2's] unfolding h by auto qed show ?thesis using assms(1)[symmetric] unfolding Neq apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(simp del : substInfinitesimalLinear.simps substInfinitesimalLinearUni.simps substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps add: insertion_neg insertion_mult insertion_add insertion_pow insertion_sub insertion_four a_def[symmetric] b_def[symmetric] c_def[symmetric] a_def'[symmetric] b_def'[symmetric] c_def'[symmetric] eval_list_conj eval_list_conj_Uni ) using linear quadratic_1 quadratic_2 by smt qed lemma convert_list : assumes "convert_atom_list var L (xs' @ x # xs) = Some L'" assumes "l\set(L)" shows "\l'\ set L'. convert_atom var l (xs' @ x # xs) = Some l'" using assms proof(induction L arbitrary : L') case Nil then show ?case by auto next case (Cons At L) then show ?case proof(cases At) case (Less p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Less apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all apply(cases "l = Less p") by simp_all next case (Eq p) show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Eq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all apply(cases "l = Eq p") by simp_all next case (Leq p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Leq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all apply(cases "l = Leq p") by simp_all next case (Neq p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Neq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all apply(cases "l = Neq p") by simp_all qed qed lemma convert_list2 : assumes "convert_atom_list var L (xs' @ x # xs) = Some L'" assumes "l'\set(L')" shows "\l\ set L. convert_atom var l (xs' @ x # xs) = Some l'" using assms proof(induction L arbitrary : L') case Nil then show ?case by auto next case (Cons At L) then show ?case proof(cases At) case (Less p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Less apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by blast next case (Eq p) show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Eq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by blast next case (Leq p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Leq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by blast next case (Neq p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Neq apply simp apply(cases "MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by blast qed qed lemma elimVar_atom_convert : assumes "convert_atom_list var L (xs' @ x # xs) = Some L'" assumes "convert_atom_list var L2 (xs' @ x # xs) = Some L2'" assumes "length xs' = var" shows "(\f\set L. eval (elimVar var L2 [] f) (xs' @ x # xs)) = (\f\set L'. evalUni (elimVarUni_atom L2' f) x)" proof safe fix f assume h : "f \ set L" "eval (elimVar var L2 [] f) (xs' @ x # xs)" have "\f'\set L'. convert_atom var f (xs' @ x # xs) = Some f'" using convert_list h assms by auto then obtain f' where f' : "f'\set L'" "convert_atom var f (xs' @ x # xs) = Some f'" by metis show "\f\set L'. evalUni (elimVarUni_atom L2' f) x" apply(rule bexI[where x=f']) using f' elimVar_atom_single[OF f'(2) assms(2) assms(3)] h by auto next fix f' assume h : "f' \ set L'" "evalUni (elimVarUni_atom L2' f') x" have "\f\set L. convert_atom var f (xs' @ x # xs) = Some f'" using convert_list2 h assms by auto then obtain f where f : "f\set L" "convert_atom var f (xs' @ x # xs) = Some f'" by metis show "\f\set L. eval (elimVar var L2 [] f) (xs' @ x # xs)" apply(rule bexI[where x=f]) using f elimVar_atom_single[OF f(2) assms(2) assms(3)] h by auto qed lemma eval_convert : assumes "convert_atom_list var L (xs' @ x # xs) = Some L'" assumes "length xs' = var" shows "(\f\set L. aEval f (xs' @ x # xs)) = (\f\set L'. aEvalUni f x)" using assms proof(induction L arbitrary : L') case Nil then show ?case by auto next case (Cons a L) then show ?case proof(cases a) case (Less p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Less apply(cases " MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by (simp add: poly_to_univar) next case (Eq p) then show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Eq apply(cases " MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by (simp add: poly_to_univar) next case (Leq p) show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Leq apply(cases " MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by (simp add: poly_to_univar) next case (Neq p) show ?thesis using Cons(2)[symmetric] Cons(1) Cons(3) unfolding Neq apply(cases " MPoly_Type.degree p var < 3") apply simp_all apply(cases "convert_atom_list var L (xs' @ x # xs)") apply simp_all by (simp add: poly_to_univar) qed qed lemma all_degree_2_convert : assumes "all_degree_2 var L" shows "\L'. convert_atom_list var L xs = Some L'" using assms proof(induction L) case Nil then show ?case by auto next case (Cons a L) then show ?case proof(cases a) case (Less p) show ?thesis using Cons unfolding Less all_degree_2.simps convert_atom_list.simps convert_atom.simps using degree_convert_eq[of var p xs] by auto next case (Eq p) then show ?thesis using Cons unfolding Eq all_degree_2.simps convert_atom_list.simps convert_atom.simps using degree_convert_eq[of var p xs] by auto next case (Leq x3) then show ?thesis using Cons by auto next case (Neq x4) then show ?thesis using Cons by auto qed qed lemma gen_qe_eval : assumes hlength : "length xs = var" shows "(\x. (eval (list_conj ((map Atom L) @ F)) (xs @ (x#\)))) = (\x.(eval (gen_qe var L F) (xs @ (x#\))))" proof(cases "luckyFind var L []") case None then have notLucky : "luckyFind var L [] = None" by auto then show ?thesis proof(cases F) case Nil then show ?thesis proof(cases "all_degree_2 var L") case True then have "\x.\L'. convert_atom_list var L (xs@x#\) = Some L'" using all_degree_2_convert[of var L "xs@_#\"] by auto then obtain L' where L' : "convert_atom_list var L (xs@x#\) = Some L'" by metis then have L' : "\x. convert_atom_list var L (xs@x#\) = Some L'" by (metis convert_atom_list_change hlength) show ?thesis unfolding Nil apply (simp add:eval_list_conj eval_list_disj True del:luckyFind.simps) unfolding notLucky apply (simp add:eval_list_conj eval_list_disj) using negInf_convert[OF L' assms] elimVar_atom_convert[OF L' L' assms] eval_convert[OF L' assms] using eval_generalVS''[of L'] unfolding eval_list_conj_Uni generalVS_DNF.simps eval_list_conj_Uni eval_list_disj_Uni eval_append eval_map eval_map_all evalUni.simps by auto next case False then show ?thesis using notLucky unfolding Nil False apply simp by (metis append_Nil2 hlength notLucky option.simps(4) qe_eq_repeat.simps qe_eq_repeat_eval) qed next case (Cons a list) show ?thesis apply(simp add:Cons del:qe_eq_repeat.simps) apply(rule qe_eq_repeat_eval[of xs var L "a # list" \]) using assms . qed next case (Some a) then show ?thesis using luckyFind_eval[OF Some assms] apply(cases F) apply simp apply(simp add:Cons del:qe_eq_repeat.simps) using qe_eq_repeat_eval[of xs var L _ \] using assms by auto qed lemma freeIn_elimVar : "freeIn var (elimVar var L F A)" proof(cases A) case (Less p) have two: "2 = Suc(Suc 0)" by auto have notIn4: "var \ vars (4::real mpoly)" by (metis isolate_var_one not_in_add not_in_isovarspar numeral_plus_numeral one_add_one semiring_norm(2) semiring_norm(6)) show ?thesis using Less apply auto using not_in_isovarspar apply force+ apply (rule freeIn_list_conj) apply auto defer defer using not_in_isovarspar apply force+ using not_in_sub[OF not_in_mult[of var 4, OF _ not_in_mult[of var "isolate_variable_sparse p var 2" "isolate_variable_sparse p var 0"]], of "(isolate_variable_sparse p var (Suc 0))\<^sup>2"] apply (simp add:not_in_isovarspar two) using not_in_mult[of var "isolate_variable_sparse p var (Suc 0)" "isolate_variable_sparse p var (Suc 0)"] apply (simp add:not_in_isovarspar notIn4) apply (simp add: ideal.scale_scale) apply(rule freeIn_list_conj) apply auto defer defer apply(rule freeIn_list_conj) apply auto apply(rule freeIn_substInfinitesimalQuadratic) apply auto using not_in_isovarspar not_in_neg apply blast apply (metis not_in_isovarspar not_in_neg not_in_pow power_0) using notIn4 not_in_isovarspar not_in_mult not_in_pow not_in_sub apply auto[1] apply (metis isovarspar_sum mult_2 not_in_isovarspar) using freeIn_substInfinitesimalQuadratic_fm[of var "(- isolate_variable_sparse p var (Suc 0))" "-1" "((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" "(2 * isolate_variable_sparse p var 2)"] apply auto[1] apply (metis (no_types, lifting) mult_2 notIn4 not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow not_in_sub power_0) apply(rule freeIn_substInfinitesimalLinear) apply (meson not_in_isovarspar not_in_neg) apply (simp add: not_in_isovarspar) using freeIn_substInfinitesimalLinear_fm using not_in_isovarspar not_in_neg apply force apply (metis (no_types, lifting) \\var \ vars 4; var \ vars (isolate_variable_sparse p var 2); var \ vars (isolate_variable_sparse p var 0); var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2)\ \ var \ vars (4 * (isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) - (isolate_variable_sparse p var (Suc 0))\<^sup>2)\ freeIn_substInfinitesimalQuadratic minus_diff_eq mult.assoc mult_2 notIn4 not_in_add not_in_isovarspar not_in_neg not_in_pow power_0) using freeIn_substInfinitesimalQuadratic_fm[of var "(- isolate_variable_sparse p var (Suc 0))" 1 "((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" "(2 * isolate_variable_sparse p var 2)"] apply auto by (metis (no_types, lifting) \\var \ vars 4; var \ vars (isolate_variable_sparse p var 2); var \ vars (isolate_variable_sparse p var 0); var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2)\ \ var \ vars (4 * (isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) - (isolate_variable_sparse p var (Suc 0))\<^sup>2)\ ideal.scale_scale minus_diff_eq mult_2 notIn4 not_in_add not_in_isovarspar not_in_neg not_in_pow power_0) next case (Eq p) then show ?thesis using freeIn_elimVar_eq by auto next case (Leq p) then show ?thesis using freeIn_elimVar_eq by auto next case (Neq p) have two: "2 = Suc(Suc 0)" by auto have notIn4: "var \ vars (4::real mpoly)" by (metis isolate_var_one not_in_add not_in_isovarspar numeral_plus_numeral one_add_one semiring_norm(2) semiring_norm(6)) show ?thesis using Neq apply auto using not_in_isovarspar apply force+ apply (rule freeIn_list_conj) apply auto defer defer using not_in_isovarspar apply force+ using not_in_sub[OF not_in_mult[of var 4, OF _ not_in_mult[of var "isolate_variable_sparse p var 2" "isolate_variable_sparse p var 0"]], of "(isolate_variable_sparse p var (Suc 0))\<^sup>2"] apply (simp add:not_in_isovarspar two) using not_in_mult[of var "isolate_variable_sparse p var (Suc 0)" "isolate_variable_sparse p var (Suc 0)"] apply (simp add:not_in_isovarspar notIn4) apply (simp add: ideal.scale_scale) apply(rule freeIn_list_conj) apply auto defer defer apply(rule freeIn_list_conj) apply auto apply(rule freeIn_substInfinitesimalQuadratic) apply auto using not_in_isovarspar not_in_neg apply blast apply (metis not_in_isovarspar not_in_neg not_in_pow power_0) using notIn4 not_in_isovarspar not_in_mult not_in_pow not_in_sub apply auto[1] apply (metis isovarspar_sum mult_2 not_in_isovarspar) using freeIn_substInfinitesimalQuadratic_fm[of var "(- isolate_variable_sparse p var (Suc 0))" "-1" "((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" "(2 * isolate_variable_sparse p var 2)"] apply auto[1] apply (metis (no_types, lifting) mult_2 notIn4 not_in_add not_in_isovarspar not_in_mult not_in_neg not_in_pow not_in_sub power_0) apply(rule freeIn_substInfinitesimalLinear) apply (meson not_in_isovarspar not_in_neg) apply (simp add: not_in_isovarspar) using freeIn_substInfinitesimalLinear_fm using not_in_isovarspar not_in_neg apply force apply (metis (no_types, lifting) \\var \ vars 4; var \ vars (isolate_variable_sparse p var 2); var \ vars (isolate_variable_sparse p var 0); var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2)\ \ var \ vars (4 * (isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) - (isolate_variable_sparse p var (Suc 0))\<^sup>2)\ freeIn_substInfinitesimalQuadratic minus_diff_eq mult.assoc mult_2 notIn4 not_in_add not_in_isovarspar not_in_neg not_in_pow power_0) using freeIn_substInfinitesimalQuadratic_fm[of var "(- isolate_variable_sparse p var (Suc 0))" 1 "((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0)" "(2 * isolate_variable_sparse p var 2)"] apply auto by (metis (no_types, lifting) \\var \ vars 4; var \ vars (isolate_variable_sparse p var 2); var \ vars (isolate_variable_sparse p var 0); var \ vars ((isolate_variable_sparse p var (Suc 0))\<^sup>2)\ \ var \ vars (4 * (isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) - (isolate_variable_sparse p var (Suc 0))\<^sup>2)\ ideal.scale_scale minus_diff_eq mult_2 notIn4 not_in_add not_in_isovarspar not_in_neg not_in_pow power_0) qed lemma freeInDisj: "freeIn var (list_disj (list_conj (map (substNegInfinity var) L) # map (elimVar var L []) L))" apply(rule freeIn_list_disj) apply(auto) apply(rule freeIn_list_conj) apply simp using freeIn_substNegInfinity[of var] apply simp using freeIn_elimVar by simp lemma gen_qe_eval' : assumes "all_degree_2 var L" assumes "length xs' = var" shows "(\x. (eval (list_conj (map Atom L)) (xs'@x#\))) = (\x.(eval (gen_qe var L []) (xs'@x#\)))" "freeIn var (gen_qe var L [])" proof- have h : "(\x. (eval (list_conj (map Atom L)) (xs'@x#\))) = (\x. eval (gen_qe var L []) (xs'@x # \))" using gen_qe_eval[OF assms(2), of L "[]" \] unfolding List.append.left_neutral by auto show "(\x. (eval (list_conj (map Atom L)) (xs'@x#\))) = (\x.(eval (gen_qe var L []) (xs'@x#\)))" unfolding h apply (simp add:assms) apply(cases "find_lucky_eq var L") apply simp using freeInDisj[of var L] using var_not_in_eval3[OF _ assms(2)] apply blast subgoal for a using freeIn_elimVar_eq[of var L "[]" a] apply(simp del:elimVar.simps) using var_not_in_eval3[OF _ assms(2)] by blast done next show "freeIn var (gen_qe var L []) " apply(simp add:assms) apply(cases "find_lucky_eq var L") apply (simp add:freeInDisj) subgoal for a using freeIn_elimVar_eq[of var L "[]" a] by(simp del:elimVar.simps) done qed lemma gen_qe_eval'' : assumes "all_degree_2 var L" assumes "length xs' = var" shows "(\x. (eval (list_conj (map Atom L)) (xs'@x#\))) = (\x.(eval (list_disj (list_conj (map (substNegInfinity var) L) # map (elimVar var L []) L)) (xs'@x#\)))" proof(cases "convert_atom_list var L (xs'@x#\)") case None then show ?thesis using all_degree_2_convert[OF assms(1), of "(xs' @ x # \)"] by auto next case (Some a) then have Some : "\x. convert_atom_list var L (xs'@x#\) = Some a" using convert_atom_list_change[OF assms(2), of L x \] by fastforce show ?thesis apply (simp add: eval_list_conj eval_list_disj) using negInf_convert[OF Some assms(2)] elimVar_atom_convert[OF Some Some assms(2)] eval_convert[OF Some assms(2)] using eval_generalVS''[of a] unfolding eval_list_conj_Uni generalVS_DNF.simps eval_list_conj_Uni eval_list_disj_Uni eval_append eval_map eval_map_all evalUni.simps by auto qed end \ No newline at end of file diff --git a/thys/Virtual_Substitution/InfinitesimalsUni.thy b/thys/Virtual_Substitution/InfinitesimalsUni.thy --- a/thys/Virtual_Substitution/InfinitesimalsUni.thy +++ b/thys/Virtual_Substitution/InfinitesimalsUni.thy @@ -1,1301 +1,1301 @@ theory InfinitesimalsUni imports Infinitesimals UniAtoms NegInfinityUni QE begin fun convertDerivativeUni :: "real * real * real \ atomUni fmUni" where "convertDerivativeUni (a,b,c) = OrUni(AtomUni(LessUni(a,b,c)))(AndUni(AtomUni(EqUni(a,b,c)))( OrUni(AtomUni(LessUni(0,2*a,b)))(AndUni(AtomUni(EqUni(0,2*a,b)))( (AtomUni(LessUni(0,0,2*a))) )) )) " lemma convert_convertDerivative : assumes "convert_poly var p (xs'@x#xs) = Some(a,b,c)" assumes "length xs' = var" shows "eval (convertDerivative var p) (xs'@x#xs) = evalUni (convertDerivativeUni (a,b,c)) x" proof(cases "MPoly_Type.degree p var = 0") case True then show ?thesis using assms apply (simp add: isovar_greater_degree eval_or eval_and insertion_mult insertion_const) using sum_over_zero[of p var] by auto next case False then have nonzero: "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True have h1 : "MPoly_Type.degree p var < 3" using True by auto have h2 : "get_coeffs var p = (isolate_variable_sparse p var 2, isolate_variable_sparse p var 1, isolate_variable_sparse p var 0)" by auto have h : "insertion (nth_default 0 (xs' @ x # xs)) p = b * x + c" using poly_to_univar[OF h1 h2 _ _ _ assms(2), of a x xs b c x] using assms(1) apply(cases "MPoly_Type.degree p var < 3") apply simp_all using isovar_greater_degree[of p var] unfolding True by simp have h3: "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1) var = 0" using degree_mult[of "isolate_variable_sparse p var (Suc 0)" "Const 1" var] using degree_isovarspar mult_one_right by presburger show ?thesis using assms True unfolding convertDerivative.simps[of _ p] convertDerivative.simps[of _ "(derivative var p)"] apply (simp add: derivative_def isovar_greater_degree eval_or eval_and insertion_add insertion_mult insertion_const HOL.arg_cong[OF sum_over_zero[of p var], of "insertion (nth_default var (xs'@x#xs))"] insertion_var_zero del:convertDerivative.simps) unfolding h h3 by(simp del:convertDerivative.simps) next case False then have deg2 : "MPoly_Type.degree p var = 2" by (metis One_nat_def assms(1) convert_poly.simps le_SucE less_Suc0 less_Suc_eq_le nonzero numeral_2_eq_2 numeral_3_eq_3 option.distinct(1)) then have first : "isolate_variable_sparse p var (Suc (Suc 0)) \ 0" by (metis MPoly_Type.degree_zero isolate_variable_sparse_degree_eq_zero_iff nat.distinct(1) numeral_2_eq_2) have second : "(isolate_variable_sparse p var (Suc (Suc 0)) * Var var)\0" by (metis MPoly_Type.degree_zero One_nat_def ExecutiblePolyProps.degree_one Zero_not_Suc first mult_eq_0_iff) have other : "Const (2::real)\0" by (simp add: nonzero_const_is_nonzero) have thing: "(Var var::real mpoly) \ 0" using second by auto have degree: "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var - Suc 0 = 0" apply simp apply(rule Nat.eq_imp_le) apply(rule degree_less_sum'[of _ _ 0]) apply (simp add: degree_isovarspar mult_one_right) apply auto unfolding degree_mult[OF second other, of var] degree_const apply simp unfolding degree_mult[OF first thing] degree_one using degree_isovarspar by force have insertion: "insertion (nth_default 0 (xs'@x#xs)) (\(i::nat)\2. isolate_variable_sparse p var i * Var var ^ i) = a * x^2 + b * x + c" proof(cases "MPoly_Type.degree p var = 1") case True then show ?thesis using False by blast next case False then have deg2 : "MPoly_Type.degree p var = 2" by (metis One_nat_def assms(1) convert_poly.simps le_SucE less_Suc0 less_Suc_eq_le nonzero numeral_2_eq_2 numeral_3_eq_3 option.distinct(1)) then have degless3 : "MPoly_Type.degree p var < 3" by auto have thing : "vari\2. isolate_variable_sparse p var i * Var var ^ i) = p" using deg2 by (metis sum_over_zero) have three: "(3::nat) = Suc(Suc(Suc(0)))" by auto have "(\i = 0..<3. insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var i) * x ^ i) = (insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) * x ^ 0) + (insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (1::nat)) * x ^ (1::nat)) + (insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (2::nat)) * x ^ (2::nat))" unfolding Set_Interval.comm_monoid_add_class.sum.atLeast0_lessThan_Suc three proof - have "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (MPoly_Type.degree p var)) * x ^ MPoly_Type.degree p var + (x * insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 1) + (insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) + (\n = 0..<0. insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var n) * x ^ n))) = insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) + x * insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 1) + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (MPoly_Type.degree p var)) * x ^ MPoly_Type.degree p var" by auto then show "(\n = 0..<0. insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var n) * x ^ n) + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) * x ^ 0 + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc 0)) * x ^ Suc 0 + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc (Suc 0))) * x ^ Suc (Suc 0) = insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) * x ^ 0 + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 1) * x ^ 1 + insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 2) * x\<^sup>2" by (metis (no_types) One_nat_def Suc_1 add.commute deg2 mult.commute mult.left_neutral power_0 power_one_right) qed also have "... = a * x\<^sup>2 + b * x + c" unfolding Power.power_class.power.power_0 Power.monoid_mult_class.power_one_right Groups.monoid_mult_class.mult_1_right using assms unfolding convert_poly.simps using degless3 by simp finally have h2 :"(\i = 0..<3. insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var i) * x ^ i) = a * x\<^sup>2 + b * x + c " . show ?thesis using sum_over_degree_insertion[OF thing deg2, of x] apply auto using h h2 using assms(2) by auto qed have insertionb: "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc 0)) = b" using assms apply(cases "MPoly_Type.degree p var < 3") by simp_all have insertiona : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc (Suc 0))) = a" using assms apply(cases "MPoly_Type.degree p var < 3") apply simp_all by (simp add: numeral_2_eq_2) have x : "insertion (nth_default 0 (xs' @ x # xs)) (Var var) = x" using insertion_var[of var "(xs' @ x # xs)" x] using assms(2) by auto have zero1 : "insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0)) var (Suc 0)) = 0" by (simp add: degree_isovarspar isovar_greater_degree) have zero2 : "insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc (Suc 0))) var 0) = a" using degree0isovarspar degree_isovarspar insertiona by presburger have zero3 : "insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse (Var var) var (Suc 0)) = 1" using isolate_var_one using MPoly_Type.insertion_one by fastforce have zero4 : "insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc (Suc 0))) var (Suc 0)) = 0" by (simp add: degree_isovarspar isovar_greater_degree) have insertion_deriv : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0)) = 2*a" unfolding isovarspar_sum isolate_variable_sparse_mult apply auto unfolding const_lookup_suc const_lookup_zero Rings.mult_zero_class.mult_zero_right Groups.group_add_class.add.group_left_neutral unfolding insertion_add insertion_mult insertion_const apply auto unfolding zero1 zero2 zero3 zero4 by simp have deg2 : "MPoly_Type.degree p var = 2" using degree_convert_eq[OF assms(1)] False nonzero by auto then show ?thesis using assms unfolding convertDerivative.simps[of _ p] convertDerivative.simps[of _ "(derivative var p)"] convertDerivative.simps[of _ "(derivative var (derivative var p))"] apply (simp add: x insertionb insertiona insertion_deriv insertion degree derivative_def isovar_greater_degree eval_or eval_and insertion_add insertion_mult insertion_const HOL.arg_cong[OF sum_over_zero[of p var], of "insertion (nth_default 0 (xs'@x#xs))"] insertion_var_zero del:convertDerivative.simps) by (smt (z3) One_nat_def degree_isovarspar distrib_right insertion_deriv isolate_variable_sparse_ne_zeroD mult_one_right neq0_conv not_one_le_zero zero1) qed qed fun linearSubstitutionUni :: "real \ real \ atomUni \ atomUni fmUni" where "linearSubstitutionUni b c a = (if aEvalUni a (-c/b) then TrueFUni else FalseFUni)" lemma convert_linearSubstitutionUni: assumes "convert_atom var a (xs'@x#xs) = Some(a')" assumes "insertion (nth_default 0 (xs'@x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@x#xs)) c = C" assumes "B \ 0" assumes "var\(vars b)" assumes "var\(vars c)" assumes "length xs' = var" shows "aEval (linear_substitution var (-c) b a) (xs'@x#xs) = evalUni (linearSubstitutionUni B C a') x" using assms proof - have "aEval a (xs'@(-C/B)#xs) = evalUni (linearSubstitutionUni B C a') x" using assms(1) proof(cases "a") case (Less p) then have "MPoly_Type.degree p var < 3" using assms(1)apply(cases "MPoly_Type.degree p var < 3") by auto then show ?thesis using Less assms apply simp using poly_to_univar by force next case (Eq p) then have "MPoly_Type.degree p var < 3" using assms(1)apply(cases "MPoly_Type.degree p var < 3") by auto then show ?thesis using Eq assms apply simp using poly_to_univar by force next case (Leq p) then have "MPoly_Type.degree p var < 3" using assms(1)apply(cases "MPoly_Type.degree p var < 3") by auto then show ?thesis using Leq assms apply simp using poly_to_univar by force next case (Neq p) then have "MPoly_Type.degree p var < 3" using assms(1)apply(cases "MPoly_Type.degree p var < 3") by auto then show ?thesis using Neq assms apply simp using poly_to_univar by force qed then have subst : "aEval a ((xs'@x#xs)[var := - C / B]) = evalUni (linearSubstitutionUni B C a') x" using assms by auto have hlength : "var< length (xs'@x#xs)" using assms by auto have inB : "insertion (nth_default 0 ((xs'@x#xs)[var := - C / B])) b = B" using assms apply auto apply(cases a) apply auto by (simp add: insertion_lowerPoly1)+ have inC : "insertion (nth_default 0 ((xs'@x#xs)[var := - C / B])) (-c) = -C" using assms apply auto apply(cases a) apply auto by (simp add: insertion_lowerPoly1 insertion_neg)+ have freenegc : "var\vars(-c)" using assms not_in_neg by auto show ?thesis using linear[OF hlength assms(4) freenegc assms(5) inC inB, of a ] subst using var_not_in_eval3[OF var_not_in_linear[OF freenegc assms(5)] assms(7)] by (metis assms(7) list_update_length) qed (* substInfinitesimalLinear var b c A substitutes -c/b+epsilon for variable var in atom A assumes b is nonzero defined in page 615 *) fun substInfinitesimalLinearUni :: "real \ real \ atomUni \ atomUni fmUni" where "substInfinitesimalLinearUni b c (EqUni p) = allZero' p"| "substInfinitesimalLinearUni b c (LessUni p) = map_atomUni (linearSubstitutionUni b c) (convertDerivativeUni p)"| "substInfinitesimalLinearUni b c (LeqUni p) = OrUni (allZero' p) (map_atomUni (linearSubstitutionUni b c) (convertDerivativeUni p))"| "substInfinitesimalLinearUni b c (NeqUni p) = negUni (allZero' p)" lemma convert_linear_subst_fm : assumes "convert_atom var a (xs'@x#xs) = Some a'" assumes "insertion (nth_default 0 (xs'@x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@x#xs)) c = C" assumes "B \ 0" assumes "var\(vars b)" assumes "var\(vars c)" assumes "length xs' = var" shows "aEval (linear_substitution (var + 0) (liftPoly 0 0 (-c)) (liftPoly 0 0 b) a) (xs'@x#xs) = evalUni (linearSubstitutionUni B C a') x" proof- have lb : "insertion (nth_default 0 (xs'@x#xs)) (liftPoly 0 0 b) = B" unfolding lift00 using assms(2) by auto have lc : "insertion (nth_default 0 (xs'@x#xs)) (liftPoly 0 0 c) = C" unfolding lift00 using assms(3) insertion_neg by auto have nb : "var \ vars (liftPoly 0 0 b)" using not_in_lift[OF assms(5), of 0] by auto have nc : "var \ vars (liftPoly 0 0 c)" using not_in_lift[OF assms(6)] not_in_neg using assms(6) lift00 by auto then show ?thesis using assms using lb lc convert_linearSubstitutionUni[OF assms(1) lb lc assms(4) nb nc] by (simp add: lift00) qed lemma evalUni_if : "evalUni (if cond then TrueFUni else FalseFUni) x = cond" by(cases cond)(auto) lemma degree_less_sum' : "MPoly_Type.degree (p::real mpoly) var = n \ MPoly_Type.degree (q::real mpoly) var = m \ n < m \ MPoly_Type.degree (p + q) var = m" using degree_less_sum[of q var m p n] by (simp add: add.commute) lemma convert_substInfinitesimalLinear_less : assumes "convert_poly var p (xs'@x#xs) = Some(p')" assumes "insertion (nth_default 0 (xs'@x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@x#xs)) c = C" assumes "B \ 0" assumes "var\(vars b)" assumes "var\(vars c)" assumes "length xs' = var" shows " eval (liftmap (\x. \A. Atom(linear_substitution (var+x) (liftPoly 0 x (-c)) (liftPoly 0 x b) A)) (convertDerivative var p) 0) (xs'@x#xs) = evalUni (map_atomUni (linearSubstitutionUni B C) (convertDerivativeUni p')) x" proof(cases p') case (fields a' b' c') have h : "convert_poly var p (xs'@x#xs) = Some (a', b', c')" using assms fields by auto have h2 : "\F'. convert_fm var (convertDerivative var p) xs = Some F'" unfolding convertDerivative.simps[of _ p] convertDerivative.simps[of _ "derivative var p"] convertDerivative.simps[of _ "derivative var (derivative var p)"] apply( auto simp del: convertDerivative.simps) using degree_convert_eq h apply blast using assms(1) degree_convert_eq apply blast apply (metis Suc_eq_plus1 degree_derivative gr_implies_not0 less_trans_Suc nat_neq_iff) using assms(1) degree_convert_eq apply blast apply (meson assms(1) degree_convert_eq) apply (metis One_nat_def Suc_eq_plus1 degree_derivative less_2_cases less_Suc_eq nat_neq_iff numeral_3_eq_3 one_add_one) using assms(1) degree_convert_eq apply blast using degree_derivative apply force using assms(1) degree_convert_eq apply blast apply (meson assms(1) degree_convert_eq) apply (metis degree_derivative eq_numeral_Suc less_add_one less_trans_Suc not_less_eq numeral_2_eq_2 pred_numeral_simps(3)) apply (meson assms(1) degree_convert_eq) using degree_derivative apply fastforce by (meson assms(1) degree_convert_eq) have c'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) = c'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto have b'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc 0)) = b'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto then have b'_insertion2 : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 1) = b'" by auto have a'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 2) = a'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto have liftb : "liftPoly 0 0 b = b" using lift00 by auto have liftc : "liftPoly 0 0 c = c" using lift00 by auto have b'_insertion' : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0)) var 0) = b'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") apply auto by (simp add: degree0isovarspar degree_isovarspar) have insertion_into_1 : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (Const 1) var 0) = 1" by (simp add: const_lookup_zero insertion_const) have twominusone : "((2-1)::nat) = 1" by auto show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True have simp: "(convertDerivative var p)=Atom(Less p)" using True by auto have azero : "a'=0" by (metis MPoly_Type.insertion_zero True a'_insertion isolate_variable_sparse_ne_zeroD nat.simps(3) not_less numeral_2_eq_2 zero_less_iff_neq_zero) have bzero : "b'=0" using True b'_insertion isovar_greater_degree by fastforce show ?thesis unfolding fields substInfinitesimalLinearUni.simps convertDerivativeUni.simps linearSubstitutionUni.simps map_atomUni.simps evalUni.simps evalUni_if aEvalUni.simps Rings.mult_zero_class.mult_zero_left Rings.mult_zero_class.mult_zero_right Groups.add_0 azero bzero substInfinitesimalLinear.simps convertDerivative.simps[of _ p] True simp liftmap.simps linear_substitution.simps apply (auto simp add:True) unfolding c'_insertion by auto next case False then have degnot0 : "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True then have simp : "convertDerivative var p = Or (fm.Atom (Less p)) (And (fm.Atom (Eq p)) (fm.Atom (Less (derivative var p))))" by (metis One_nat_def Suc_eq_plus1 add_right_imp_eq convertDerivative.simps degnot0 degree_derivative zero_less_one) have azero : "a'=0" by (metis MPoly_Type.insertion_zero One_nat_def True a'_insertion isovar_greater_degree lessI numeral_2_eq_2) have degderiv : "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1) var = 0" using degree_mult by (simp add: degree_isovarspar mult_one_right) show ?thesis unfolding fields substInfinitesimalLinearUni.simps convertDerivativeUni.simps linearSubstitutionUni.simps map_atomUni.simps evalUni.simps evalUni_if aEvalUni.simps Rings.mult_zero_class.mult_zero_left Rings.mult_zero_class.mult_zero_right Groups.add_0 azero substInfinitesimalLinear.simps True simp liftmap.simps linear_substitution.simps eval_Or eval_And liftb liftc apply auto unfolding derivative_def True insertion_sub insertion_mult c'_insertion b'_insertion assms lift00 apply auto unfolding insertion_sub insertion_mult c'_insertion b'_insertion assms lift00 apply (smt diff_divide_eq_iff divide_less_0_iff mult_less_0_iff) apply (smt mult_imp_less_div_pos neg_less_divide_eq zero_le_mult_iff) using assms(4) mult.commute nonzero_mult_div_cancel_left apply smt unfolding degderiv apply auto unfolding isolate_variable_sparse_mult apply auto unfolding insertion_mult defer apply (smt assms(4) diff_divide_eq_iff divide_less_0_iff mult_less_0_iff) defer using assms(4) apply blast unfolding b'_insertion' insertion_into_1 apply auto by (smt assms(4) less_divide_eq mult_pos_neg2 no_zero_divisors zero_less_mult_pos) next case False then have degreetwo : "MPoly_Type.degree p var = 2" using degnot0 by (metis One_nat_def degree_convert_eq h less_2_cases less_Suc_eq numeral_2_eq_2 numeral_3_eq_3) have two : "(2::nat) = Suc(Suc 0)" by auto have sum : "(\i = 0..<2. isolate_variable_sparse p var i * (- c) ^ i * b ^ (2 - i)) = isolate_variable_sparse p var 0 * (- c) ^ 0 * b ^ (2 - 0) + isolate_variable_sparse p var 1 * (- c) ^ 1 * b ^ (2 - 1) " unfolding Set_Interval.comm_monoid_add_class.sum.atLeast0_lessThan_Suc two by auto have a : "isolate_variable_sparse p var (Suc (Suc 0)) \ 0" by (metis degnot0 degree_isovarspar degreetwo isolate_variable_sparse_degree_eq_zero_iff numeral_2_eq_2) have b : "((Var var * Const 2) :: real mpoly) \ (0::real mpoly)" by (metis MPoly_Type.degree_zero ExecutiblePolyProps.degree_one mult_eq_0_iff nonzero_const_is_nonzero zero_neq_numeral zero_neq_one) have degreedeg1 : "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var = 1" apply(rule degree_less_sum'[where n ="0"]) apply (simp add: degree_isovarspar mult_one_right) defer apply simp using degree_mult[OF a b, of var] - by (metis (no_types, hide_lams) ExecutiblePolyProps.degree_one add.left_neutral b degree_const degree_isovarspar degree_mult mult.commute mult_zero_class.mult_zero_right) + by (metis (no_types, opaque_lifting) ExecutiblePolyProps.degree_one add.left_neutral b degree_const degree_isovarspar degree_mult mult.commute mult_zero_class.mult_zero_right) have simp : "(convertDerivative var p) = Or (fm.Atom (Less p)) (And (fm.Atom (Eq p)) (Or (fm.Atom (Less (derivative var p))) (And (fm.Atom (Eq (derivative var p))) (fm.Atom (Less (derivative var (derivative var p)))))))" using degreetwo by (metis One_nat_def Suc_1 Suc_eq_plus1 add_diff_cancel_right' convertDerivative.simps degree_derivative neq0_conv zero_less_Suc) have a : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var 0) = b'" unfolding isovarspar_sum isolate_variable_sparse_mult apply auto unfolding const_lookup_suc const_lookup_zero Rings.mult_zero_class.mult_zero_right Groups.group_add_class.add.group_left_neutral by (simp add: b'_insertion' isolate_var_0 mult_one_right) have b : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0)) = 2 * a'" unfolding isovarspar_sum isolate_variable_sparse_mult apply auto unfolding const_lookup_suc const_lookup_zero Rings.mult_zero_class.mult_zero_right Groups.group_add_class.add.group_left_neutral unfolding insertion_add insertion_mult insertion_const by (metis MPoly_Type.insertion_one MPoly_Type.insertion_zero One_nat_def a'_insertion add.commute add.right_neutral degree0isovarspar degree_isovarspar isolate_var_one isovar_greater_degree mult.commute mult.right_neutral mult_zero_class.mult_zero_right numeral_2_eq_2 zero_less_one) have simp_insertion_blob : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var 0 * b - isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0) * c) = b' * B - 2 * a' * C" unfolding insertion_sub insertion_mult assms a b by auto have a : "isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0) \ 0" by (metis MPoly_Type.degree_zero One_nat_def degreedeg1 isolate_variable_sparse_degree_eq_zero_iff zero_neq_one) have b' : "(Const 1::real mpoly) \ 0" by (simp add: nonzero_const_is_nonzero) have degreeblob : "MPoly_Type.degree (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0) * Const 1) var = 0" unfolding degree_mult[OF a b', of var] by (simp add: degree_isovarspar degree_eq_iff monomials_Const) have otherblob : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0) * Const 1) var 0) = 2 * a'" using b by (simp add: degree0isovarspar degree_isovarspar mult_one_right) have "(c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2 < 0) = ((c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2)/(B\<^sup>2) < 0)" by (simp add: assms(4) divide_less_0_iff) also have "... = (((c' * B\<^sup>2)/(B\<^sup>2) - (b' * C * B)/(B\<^sup>2) + (a' * C\<^sup>2)/(B\<^sup>2)) < 0)" by (metis (no_types, lifting) add_divide_distrib diff_divide_distrib ) also have "... = (a' * (C / B)\<^sup>2 - b' * C / B + c' < 0)" proof - { assume "c' + a' * (C / B)\<^sup>2 - b' * (C / B) < 0" then have ?thesis by (simp add: assms(4) power2_eq_square) } moreover { assume "\ c' + a' * (C / B)\<^sup>2 - b' * (C / B) < 0" then have ?thesis by (simp add: power2_eq_square) } ultimately show ?thesis by fastforce qed finally have h1: "(c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2 < 0) = (a' * (C / B)\<^sup>2 - b' * C / B + c' < 0)" . have "(c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2 = 0) = ((c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2)/(B\<^sup>2) = 0)" by (simp add: assms(4)) also have "... = (((c' * B\<^sup>2)/(B\<^sup>2) - (b' * C * B)/(B\<^sup>2) + (a' * C\<^sup>2)/(B\<^sup>2)) = 0)" by (metis (no_types, lifting) add_divide_distrib diff_divide_distrib ) also have "... = (a' * (C / B)\<^sup>2 - b' * C / B + c' = 0)" proof - { assume "c' + a' * (C * (C / (B * B))) - b' * (C / B) \ 0" then have ?thesis by (simp add: assms(4) power2_eq_square) } moreover { assume "c' + a' * (C * (C / (B * B))) - b' * (C / B) = 0" then have ?thesis by (simp add: power2_eq_square) } ultimately show ?thesis by fastforce qed finally have h2 : "(c' * B\<^sup>2 - b' * C * B + a' * C\<^sup>2 = 0) = (a' * (C / B)\<^sup>2 - b' * C / B + c' = 0)" . have h3 : "((b' * B - 2 * a' * C) * B < 0) = (b' < 2 * a' * C / B)" by (smt assms(4) less_divide_eq zero_le_mult_iff) have h4 : "(b' * B = 2 * a' * C) = (b' = 2 * a' * C / B)" by (simp add: assms(4) nonzero_eq_divide_eq) show ?thesis unfolding fields substInfinitesimalLinearUni.simps convertDerivativeUni.simps linearSubstitutionUni.simps map_atomUni.simps evalUni.simps evalUni_if aEvalUni.simps Rings.mult_zero_class.mult_zero_left Rings.mult_zero_class.mult_zero_right Groups.add_0 substInfinitesimalLinear.simps degreetwo simp liftmap.simps linear_substitution.simps eval_Or eval_And liftb liftc apply simp unfolding derivative_def degreetwo insertion_sub insertion_mult c'_insertion b'_insertion assms apply simp unfolding sum insertion_add insertion_mult insertion_pow insertion_neg assms unfolding b'_insertion2 c'_insertion a'_insertion unfolding Power.power_class.power.power_0 Groups.monoid_mult_class.mult_1_right Groups.cancel_comm_monoid_add_class.diff_zero Power.monoid_mult_class.power_one_right twominusone degreedeg1 apply simp unfolding insertion_mult assms simp_insertion_blob degreeblob unfolding insertion_mult insertion_sub assms otherblob apply simp unfolding otherblob h1 h2 h3 h4 unfolding lift00 insertion_neg assms insertion_isovarspars_free insertion_sum insertion_mult insertion_sub degree0isovarspar degree_isovarspar mult_one_right insertion_sum_var insertion_pow insertion_neg sum unfolding assms b'_insertion c'_insertion a'_insertion insertion_neg insertion_mult insertion_add insertion_pow apply simp by (smt assms(2) assms(3) b'_insertion h1 h2 h3 h4 insertion_mult insertion_sub mult_one_right simp_insertion_blob) qed qed qed lemma convert_substInfinitesimalLinear: assumes "convert_atom var a (xs'@x#xs) = Some(a')" assumes "insertion (nth_default 0 (xs'@x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@x#xs)) c = C" assumes "B \ 0" assumes "var\(vars b)" assumes "var\(vars c)" assumes "length xs' = var" shows "eval (substInfinitesimalLinear var (-c) b a) (xs'@x#xs) = evalUni (substInfinitesimalLinearUni B C a') x" using assms proof(cases a) case (Less p) have "\p'. convert_poly var p (xs'@x#xs) = Some p'" using Less assms(1) apply(cases "MPoly_Type.degree p var < 3") by auto then obtain p' where p'_def : "convert_poly var p (xs'@x#xs) = Some p'" by auto have A'_simp : "a' = LessUni p'" using assms Less using p'_def by auto have h1 : "eval (convertDerivative var p) (xs'@x#xs) = evalUni (convertDerivativeUni p') x" using convert_convertDerivative apply ( cases p') using A'_simp Less assms by auto show ?thesis unfolding A'_simp using convert_substInfinitesimalLinear_less[OF p'_def assms(2-7)] unfolding Less by auto next case (Eq p) define p' where "p' = (case convert_poly var p (xs'@x#xs) of Some p' \ p')" have A'_simp : "a' = EqUni p'" using assms Eq using p'_def by auto show ?thesis unfolding Eq A'_simp substInfinitesimalLinear.simps substInfinitesimalLinearUni.simps using convert_allZero A'_simp Eq assms by auto next case (Leq p) have "\p'. convert_poly var p (xs' @ x # xs) = Some p'" using assms(1) unfolding Leq apply auto apply(cases "MPoly_Type.degree p var < 3") by auto then obtain p' where p'_def : "convert_poly var p (xs' @ x # xs) = Some p'" by metis have A'_simp : "a' = LeqUni p'" using assms Leq using p'_def by auto have h1 : "eval (convertDerivative var p) (xs'@x#xs) = evalUni (convertDerivativeUni p') x" using convert_convertDerivative apply(cases p') using A'_simp Leq assms by auto show ?thesis unfolding A'_simp Leq substInfinitesimalLinear.simps eval_Or substInfinitesimalLinearUni.simps evalUni.simps using convert_substInfinitesimalLinear_less[OF p'_def assms(2-7)] convert_allZero[OF p'_def assms(7)] by simp next case (Neq p) have "\p'. convert_poly var p (xs' @ x # xs) = Some p'" using assms(1) unfolding Neq apply auto apply(cases "MPoly_Type.degree p var < 3") by auto then obtain p' where p'_def : "convert_poly var p (xs' @ x # xs) = Some p'" by metis have A'_simp : "a' = NeqUni p'" using assms Neq using p'_def by auto show ?thesis unfolding Neq A'_simp substInfinitesimalLinear.simps substInfinitesimalLinearUni.simps using convert_allZero[OF p'_def assms(7)] by (metis A'_simp Neq assms(1) assms(7) convert_substNegInfinity eval.simps(6) eval_neg substNegInfinityUni.simps(4) substNegInfinity.simps(4)) qed lemma either_or: fixes r :: "real" assumes a: "(\y'>r. \x\{r<..y'}. (aEvalUni (EqUni (a, b, c)) x) \ (aEvalUni (LessUni (a, b, c)) x))" shows "(\y'>r. \x\{r<..y'}. (aEvalUni (EqUni (a, b, c)) x)) \ (\y'>r. \x\{r<..y'}. (aEvalUni (LessUni (a, b, c)) x))" proof (cases "a = 0 \ b = 0 \ c= 0") case True then have "(\y'>r. \x\{r<..y'}. (aEvalUni (EqUni (a, b, c)) x))" using assms by auto then show ?thesis by blast next case False then have noz: "a\0 \ b\0 \ c\0" by auto obtain y1' where y1prop: "y1' > r \ (\x\{r<..y1'}. (aEvalUni (EqUni (a, b, c)) x) \ (aEvalUni (LessUni (a, b, c)) x))" using a by auto obtain y2' where y2prop: "y2' > r \ (\x\{r<..y2'}. a * x\<^sup>2 + b * x + c \ 0)" using noz nonzcoeffs[of a b c] by auto let ?y = "min y1' y2'" have ygt: "?y > r" using y1prop y2prop by auto have "\x\{r<..?y}. (aEvalUni (LessUni (a, b, c)) x)" using y1prop y2prop greaterThanAtMost_iff by force then show ?thesis using ygt by blast qed lemma infinitesimal_linear'_helper : assumes at_is: "At = LessUni p \ At = EqUni p" assumes "B \ 0" shows "((\y'::real>-C/B. \x::real \{-C/B<..y'}. aEvalUni At x) = evalUni (substInfinitesimalLinearUni B C At) x)" proof (cases "At = LessUni p") case True then have LessUni: "At = LessUni p" by auto then show ?thesis proof(cases p) case (fields a b c) then show ?thesis unfolding LessUni fields using one_root_a_lt0[where r="C/B", where a= "a", where b="b",where c= "c"] apply(auto) using continuity_lem_lt0_expanded[where a= "a", where b = "2 * a * C / B ", where c = "c"] apply (auto) using continuity_lem_gt0_expanded[where a = "a", where b = "2 * a * C / B",where c = "c", where r = "- (C / B)"] apply (auto) apply (meson less_eq_real_def linorder_not_less) using one_root_a_gt0[where r = "C/B", where a = "a", where b="b", where c="c"] apply (auto) using continuity_lem_lt0_expanded[where a= "a", where b = "2 * a * C / B", where c= "c"] apply (auto) using continuity_lem_gt0_expanded[where a = "a", where b = "2 * a * C / B",where c = "c", where r = "- (C / B)"] apply (auto) apply (meson less_eq_real_def linorder_not_less) using case_d1 apply (auto) using continuity_lem_lt0_expanded[where a= "a", where b = "b", where c= "c"] apply (auto) using continuity_lem_gt0_expanded[where a = "a", where b = "b",where c = "c", where r = "- (C / B)"] apply (auto) apply (meson less_eq_real_def linorder_not_less) using case_d4 apply (auto) using continuity_lem_lt0_expanded[where a= "a", where b = "b", where c= "c"] apply (auto) using continuity_lem_gt0_expanded[where a = "a", where b = "b",where c = "c", where r = "- (C / B)"] apply (auto) by (meson less_eq_real_def linorder_not_le) qed next case False then have EqUni: "At = EqUni p" using at_is by auto then show ?thesis proof(cases p) case (fields a b c) show ?thesis apply(auto simp add:EqUni fields) using continuity_lem_eq0[where r= "-(C/B)"] apply blast using continuity_lem_eq0[where r= "-(C/B)"] apply blast using continuity_lem_eq0[where r= "-(C/B)"] apply blast using linordered_field_no_ub by blast qed qed (* I assume substInfinitesimalLinearUni' was supposed to be substInfinitesimalLinearUni?*) lemma infinitesimal_linear' : assumes "B \ 0" shows "(\y'::real>-C/B. \x::real \{-C/B<..y'}. aEvalUni At x) = evalUni (substInfinitesimalLinearUni B C At) x" proof(cases At) case (LessUni p) then show ?thesis using infinitesimal_linear'_helper[of At p B C] assms by auto next case (EqUni p) then show ?thesis using infinitesimal_linear'_helper[of At p B C] assms by (auto) next case (LeqUni p) then show ?thesis proof(cases p) case (fields a b c) have same: "\x. aEvalUni (LeqUni p) x = (aEvalUni (EqUni p) x) \ (aEvalUni (LessUni p) x)" apply (simp add: fields) by force have "\a b c. At = LeqUni p \ p = (a, b, c) \ (\y'>- C / B. \x\{- C / B<..y'}. aEvalUni At x) = evalUni (substInfinitesimalLinearUni B C At) x " proof - fix a b c assume atis: "At = LeqUni p" assume p_is: " p = (a, b, c)" have s1: "(\y'>- C / B. \x\{- C / B<..y'}. aEvalUni At x) = (\y'>- C / B. \x\{- C / B<..y'}. (aEvalUni (EqUni p) x) \ (aEvalUni (LessUni p) x))" using atis same aEvalUni.simps(2) aEvalUni.simps(3) fields less_eq_real_def by blast have s2: "... = (\y'>- C / B. \x\{- C / B<..y'}. (aEvalUni (EqUni p) x)) \ (\y'>- C / B. \x\{- C / B<..y'}. (aEvalUni (LessUni p) x))" using either_or[where r = "-C/B"] p_is by blast have eq1: "(\y'>- C / B. \x\{- C / B<..y'}. (aEvalUni (EqUni p) x)) = (evalUni (substInfinitesimalLinearUni B C (EqUni p)) x)" using infinitesimal_linear'_helper[where At = "EqUni p", where p = "p", where B = "B", where C= "C"] assms by auto have eq2: "(\y'>- C / B. \x\{- C / B<..y'}. (aEvalUni (LessUni p) x)) = (evalUni (substInfinitesimalLinearUni B C (LessUni p)) x)" using infinitesimal_linear'_helper[where At = "LessUni p", where p = "p", where B = "B", where C= "C"] assms by auto have z1: "(\y'>- C / B. \x\{- C / B<..y'}. aEvalUni At x) = ((evalUni (substInfinitesimalLinearUni B C (EqUni p)) x) \ (evalUni (substInfinitesimalLinearUni B C (LessUni p)) x))" using s1 s2 eq1 eq2 by auto have z2: "(evalUni (substInfinitesimalLinearUni B C (EqUni p)) x) \ (evalUni (substInfinitesimalLinearUni B C (LessUni p)) x) = evalUni (substInfinitesimalLinearUni B C (LeqUni p)) x" by auto have z3: "(evalUni (substInfinitesimalLinearUni B C At) x) = evalUni (substInfinitesimalLinearUni B C (LeqUni p)) x" using LeqUni by auto then have z4: "(evalUni (substInfinitesimalLinearUni B C (EqUni p)) x) \ (evalUni (substInfinitesimalLinearUni B C (LessUni p)) x) = (evalUni (substInfinitesimalLinearUni B C At) x) " using z2 z3 by auto let ?a = "(evalUni (substInfinitesimalLinearUni B C (EqUni p)) x) \ (evalUni (substInfinitesimalLinearUni B C (LessUni p)) x)" let ?b = "(\y'>- C / B. \x\{- C / B<..y'}. aEvalUni At x)" let ?c = "(evalUni (substInfinitesimalLinearUni B C At) x)" have t1: "?b = ?a" using z1 by auto have t2: "?a = ?c" using z4 by (simp add: atis) then have "?b = ?c" using t1 t2 by auto then show "(\y'>- C / B. \x\{- C / B<..y'}. aEvalUni At x) = evalUni (substInfinitesimalLinearUni B C At) x" by auto qed then show ?thesis using LeqUni fields by blast qed next case (NeqUni p) then show ?thesis proof(cases p) case (fields a b c) then show ?thesis unfolding NeqUni fields using nonzcoeffs by (auto) qed qed fun quadraticSubUni :: "real \ real \ real \ real \ atomUni \ atomUni fmUni" where "quadraticSubUni a b c d A = (if aEvalUni A ((a+b*sqrt(c))/d) then TrueFUni else FalseFUni)" fun substInfinitesimalQuadraticUni :: "real \ real \ real \ real \ atomUni \ atomUni fmUni" where "substInfinitesimalQuadraticUni a b c d (EqUni p) = allZero' p"| "substInfinitesimalQuadraticUni a b c d (LessUni p) = map_atomUni (quadraticSubUni a b c d) (convertDerivativeUni p)"| "substInfinitesimalQuadraticUni a b c d (LeqUni p) = OrUni(map_atomUni (quadraticSubUni a b c d) (convertDerivativeUni p)) (allZero' p)"| "substInfinitesimalQuadraticUni a b c d (NeqUni p) = negUni (allZero' p)" lemma weird : fixes D::"real" assumes dneq: "D \ (0::real)" shows "((a'::real) * (((A::real) + (B::real) * sqrt (C::real)) / (D::real))\<^sup>2 + (b'::real) * (A + B * sqrt C) / D + c' < 0 \ a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + (c'::real) = 0 \ (b' + a' * (A + B * sqrt C) * 2 / D < 0 \ b' + a' * (A + B * sqrt C) * 2 / D = 0 \ 2 * a' < 0)) = (a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + c' < 0 \ a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + c' = 0 \ (2 * a' * (A + B * sqrt C) / D + b' < 0 \ 2 * a' * (A + B * sqrt C) / D + b' = 0 \ a' < 0))" proof (cases "(a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + c' < 0)") case True then show ?thesis by auto next case False have "a' * (A + B * sqrt C) * 2 = 2 * a' * (A + B * sqrt C)" by auto then have "a' * (A + B * sqrt C) * 2 / D =2 * a' * (A + B * sqrt C) / D " using dneq by simp then have "b' + a' * (A + B * sqrt C) * 2 / D = 2 * a' * (A + B * sqrt C) / D + b'" using add.commute by simp then have "(b' + a' * (A + B * sqrt C) * 2 / D < 0 \ b' + a' * (A + B * sqrt C) * 2 / D = 0 \ a' < 0) = (2 * a' * (A + B * sqrt C) / D + b' < 0 \ 2 * a' * (A + B * sqrt C) / D + b' = 0 \ a' < 0)" by (simp add: \b' + a' * (A + B * sqrt C) * 2 / D = 2 * a' * (A + B * sqrt C) / D + b'\) then have "(a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + c' = 0 \ (b' + a' * (A + B * sqrt C) * 2 / D < 0 \ b' + a' * (A + B * sqrt C) * 2 / D = 0 \ a' < 0)) = (a' * ((A + B * sqrt C) / D)\<^sup>2 + b' * (A + B * sqrt C) / D + c' = 0 \ (2 * a' * (A + B * sqrt C) / D + b' < 0 \ 2 * a' * (A + B * sqrt C) / D + b' = 0 \ a' < 0))" by blast then show ?thesis using False by simp qed lemma convert_substInfinitesimalQuadratic_less : assumes "convert_poly var p (xs'@x#xs) = Some p'" assumes "insertion (nth_default 0 (xs'@x#xs)) a = A" assumes "insertion (nth_default 0 (xs'@x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@x#xs)) c = C" assumes "insertion (nth_default 0 (xs'@x#xs)) d = D" assumes "D \ 0" assumes "0 \ C" assumes "var\(vars a)" assumes "var\(vars b)" assumes "var\(vars c)" assumes "var\(vars d)" assumes "length xs' = var" shows "eval (quadratic_sub_fm var a b c d (convertDerivative var p)) (xs'@x#xs) = evalUni (map_atomUni (quadraticSubUni A B C D) (convertDerivativeUni p')) x" proof(cases p') case (fields a' b' c') have h : "convert_poly var p (xs'@x#xs) = Some (a', b', c')" using assms fields by auto have h2 : "\F'. convert_fm var (convertDerivative var p) (xs'@x#xs) = Some F'" unfolding convertDerivative.simps[of _ p] convertDerivative.simps[of _ "derivative var p"] convertDerivative.simps[of _ "derivative var (derivative var p)"] apply (auto simp del: convertDerivative.simps) using degree_convert_eq h apply blast using assms(1) degree_convert_eq apply blast using degree_derivative apply fastforce apply (metis degree_convert_eq h numeral_3_eq_3 ) apply (metis (no_types, lifting) One_nat_def add.right_neutral add_Suc_right degree_derivative less_Suc_eq_0_disj less_Suc_eq_le neq0_conv numeral_3_eq_3) apply (metis One_nat_def Suc_eq_plus1 degree_derivative less_2_cases less_Suc_eq nat_neq_iff numeral_3_eq_3 one_add_one) apply (meson assms(1) degree_convert_eq) using degree_derivative apply fastforce using assms(1) degree_convert_eq apply blast apply (meson assms(1) degree_convert_eq) apply (metis degree_derivative less_Suc_eq less_add_one not_less_eq numeral_3_eq_3) apply (meson assms(1) degree_convert_eq) - apply (metis (no_types, hide_lams) Suc_1 Suc_eq_plus1 degree_derivative less_2_cases less_Suc_eq numeral_3_eq_3) + apply (metis (no_types, opaque_lifting) Suc_1 Suc_eq_plus1 degree_derivative less_2_cases less_Suc_eq numeral_3_eq_3) using assms(1) degree_convert_eq by blast have c'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 0) = c'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto then have c'_insertion'' : "\x. insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 0) = c'" using assms(12) not_in_isovarspar[of p var 0 "isolate_variable_sparse p var 0", OF HOL.refl] by (metis list_update_length not_contains_insertion) have b'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var (Suc 0)) = b'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto then have b'_insertion'' : "\x. insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var (Suc 0)) = b'" using assms(12) not_in_isovarspar[of p var "Suc 0" "isolate_variable_sparse p var (Suc 0)", OF HOL.refl] by (metis list_update_length not_contains_insertion) then have b'_insertion2 : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 1) = b'" by auto have a'_insertion : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse p var 2) = a'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") by auto then have a'_insertion'' : "\x. insertion (nth_default 0 (xs' @ x # xs)) (isolate_variable_sparse p var 2) = a'" using assms(12) not_in_isovarspar[of p var 2 "isolate_variable_sparse p var 2", OF HOL.refl] by (metis list_update_length not_contains_insertion) have liftb : "liftPoly 0 0 b = b" using lift00 by auto have liftc : "liftPoly 0 0 c = c" using lift00 by auto have b'_insertion' : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0)) var 0) = b'" using assms fields unfolding convert_poly.simps apply(cases "MPoly_Type.degree p var < 3") apply auto using degree0isovarspar degree_isovarspar by auto have insertion_into_1 : "insertion (nth_default 0 (xs'@x#xs)) (isolate_variable_sparse (Const 1) var 0) = 1" by (simp add: const_lookup_zero insertion_const) have twominusone : "((2-1)::nat) = 1" by auto have length0 : "var < length (xs'@x#xs)" using assms by auto have altinserta : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) a = A" using assms by (metis list_update_length not_contains_insertion) have altinserta' : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) a = A" using assms by (metis list_update_length not_contains_insertion) have altinsertb : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) b = B" using assms by (metis list_update_length not_contains_insertion) have altinsertb' : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) b = B" using assms by (metis list_update_length not_contains_insertion) have altinsertc : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) c = C" using assms by (metis list_update_length not_contains_insertion) have altinsertc' : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) c = C" using assms by (metis list_update_length not_contains_insertion) have altinsertd : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) d = D" using assms by (metis list_update_length not_contains_insertion) have altinsertd' : "\xa. insertion (nth_default 0 ((xs'@x#xs)[var := xa])) d = D" using assms by (metis list_update_length not_contains_insertion) have freeInQuadraticSub : "\At. eval (quadratic_sub var a b c d At) ((xs'@x#xs)[var := sqrt C]) = eval (quadratic_sub var a b c d At) ((xs'@x#xs))" by (metis assms(10) assms(11) assms(8) assms(9) free_in_quad list_update_id var_not_in_eval) have quad : "\At. (eval (quadratic_sub var a b c d At) (xs'@x#xs) = aEval At ((xs'@x#xs)[var := (A + B * sqrt C) / D]))" using quadratic_sub[OF length0 assms(6-7) assms(10) altinserta altinsertb altinsertc altinsertd, symmetric] using freeInQuadraticSub by auto show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True then have simp: "(convertDerivative var p)=Atom(Less p)" by auto have azero : "a'=0" by (metis MPoly_Type.insertion_zero True a'_insertion isolate_variable_sparse_ne_zeroD nat.simps(3) not_less numeral_2_eq_2 zero_less_iff_neq_zero) have bzero : "b'=0" using True b'_insertion isovar_greater_degree by fastforce define p1 where "p1 = isolate_variable_sparse p var 0" have degree_p1: "MPoly_Type.degree p1 var = 0" unfolding p1_def by (simp add: degree_isovarspar) define p2 where "p2 = isolate_variable_sparse p1 var 0 * Const 0 * Var var + isolate_variable_sparse p1 var 0 * Const 1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var (Suc 0)" show ?thesis unfolding substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps fields convertDerivativeUni.simps map_atomUni.simps quadraticSubUni.simps aEvalUni.simps evalUni.simps evalUni_if Rings.mult_zero_class.mult_zero_left Groups.add_0 Rings.mult_zero_class.mult_zero_right True simp azero bzero quadratic_sub_fm.simps quadratic_sub_fm_helper.simps liftmap.simps lift00 quad aEval.simps apply (simp add:True c'_insertion p1_def[symmetric] degree_p1 p2_def[symmetric] A_def[symmetric] B_def[symmetric]) unfolding A_def B_def p2_def p1_def degree0isovarspar[OF True] isovarspar_sum mult_one_right mult_zero_right mult_zero_left const_lookup_zero const_lookup_suc apply simp unfolding insertion_add insertion_sub insertion_mult insertion_pow insertion_const c'_insertion apply simp using \isolate_variable_sparse p var 0 = p\ b'_insertion2 bzero c'_insertion by force next case False then have degreenonzero : "MPoly_Type.degree p var \0" by auto show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True then have simp : "convertDerivative var p = Or (fm.Atom (Less p)) (And (fm.Atom (Eq p)) (fm.Atom (Less (derivative var p))))" by (metis One_nat_def Suc_eq_plus1 add_right_imp_eq convertDerivative.simps degree_derivative degreenonzero less_numeral_extra(1)) have azero : "a'=0" by (metis MPoly_Type.insertion_zero One_nat_def True a'_insertion isovar_greater_degree lessI numeral_2_eq_2) have degderiv : "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1) var = 0" using degree_mult by (simp add: degree_isovarspar mult_one_right) have thing : "var2 + b' * (A + B * sqrt C) / D + c'" using sum_over_degree_insertion[OF insertionp degree2, of "(A + B * sqrt C) / D", symmetric] unfolding a'_insertion[symmetric] b'_insertion[symmetric] c'_insertion[symmetric] insertion_isovarspars_free[of _ _ "(A + B * sqrt C) / D" _ _ x] unfolding two apply simp using assms(12) by force have insertion_simp : "insertion (nth_default 0 ((xs' @ x # xs)[var := (A + B * sqrt C) / D])) = insertion (nth_default 0 ((xs' @ ((A + B * sqrt C) / D) # xs)))" using assms by (metis list_update_length) have degreeone : "MPoly_Type.degree (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var = 1" apply(rule degree_less_sum'[where n=0]) apply (simp add: degree_isovarspar mult_one_right) apply (smt One_nat_def ExecutiblePolyProps.degree_one degree2 degree_const degree_isovarspar degree_mult degreenonzero isolate_variable_sparse_degree_eq_zero_iff mult.commute nonzero_const_is_nonzero numeral_2_eq_2 plus_1_eq_Suc) by simp have zero1 : " insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0)) var (Suc 0)) = 0" by (simp add: degree_isovarspar isovar_greater_degree) have zero2 : "insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc (Suc 0))) var 0) = a'" using a'_insertion'' degree0isovarspar degree_isovarspar numeral_2_eq_2 by force have zero3 : "insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse (Var var) var (Suc 0)) = 1" using isolate_var_one by fastforce have zero4 : "insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc (Suc 0))) var (Suc 0)) = 0" by (simp add: degree_isovarspar isovar_greater_degree) have insertiona' : " insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse (isolate_variable_sparse p var (Suc 0) * Const 1 + isolate_variable_sparse p var (Suc (Suc 0)) * Var var * Const 2) var (Suc 0) * Const 1) = 2 * a'" unfolding isovarspar_sum isolate_variable_sparse_mult apply auto unfolding const_lookup_suc const_lookup_zero Rings.mult_zero_class.mult_zero_right Groups.group_add_class.add.group_left_neutral unfolding insertion_add insertion_mult insertion_const b'_insertion' apply auto unfolding zero1 zero2 zero3 zero4 by auto have a' : "insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (isolate_variable_sparse p var (Suc (Suc 0))) = a'" unfolding two[symmetric] unfolding a'_insertion'' by auto have var: "insertion (nth_default 0 (xs' @ (A + B * sqrt C) / D # xs)) (Var var) = (A + B * sqrt C) / D" using assms by (metis insertion_simp insertion_var length0) show ?thesis unfolding substInfinitesimalQuadratic.simps substInfinitesimalQuadraticUni.simps fields convertDerivativeUni.simps map_atomUni.simps quadraticSubUni.simps aEvalUni.simps evalUni.simps evalUni_if Rings.mult_zero_class.mult_zero_left Groups.add_0 Rings.mult_zero_class.mult_zero_right degree2 simp quadratic_sub_fm.simps quadratic_sub_fm_helper.simps liftmap.simps lift00 Groups.monoid_add_class.add_0_right quad aEval.simps eval.simps derivative_def apply (simp add:insertion_sum insertion_add insertion_mult insertion_const insertion_var_zero) unfolding insertionp unfolding insertion_simp unfolding b'_insertion'' a'_insertion'' unfolding degreeone apply simp unfolding a' var unfolding insertiona' using weird[OF assms(6)] by auto qed qed qed lemma convert_substInfinitesimalQuadratic: assumes "convert_atom var At (xs'@ x#xs) = Some(At')" assumes "insertion (nth_default 0 (xs'@ x#xs)) a = A" assumes "insertion (nth_default 0 (xs'@ x#xs)) b = B" assumes "insertion (nth_default 0 (xs'@ x#xs)) c = C" assumes "insertion (nth_default 0 (xs'@ x#xs)) d = D" assumes "D \ 0" assumes "0 \ C" assumes "var\(vars a)" assumes "var\(vars b)" assumes "var\(vars c)" assumes "var\(vars d)" assumes "length xs' = var" shows "eval (substInfinitesimalQuadratic var a b c d At) (xs'@ x#xs) = evalUni (substInfinitesimalQuadraticUni A B C D At') x" using assms proof(cases At) case (Less p) define p' where "p' = (case convert_poly var p (xs'@ x#xs) of Some p' \ p')" have At'_simp : "At' = LessUni p'" using assms Less using p'_def by auto show ?thesis using convert_substInfinitesimalQuadratic_less[OF _ assms(2-12)] by (metis At'_simp Less None_eq_map_option_iff assms(1) convert_atom.simps(1) option.distinct(1) option.exhaust_sel option.the_def p'_def substInfinitesimalQuadraticUni.simps(2) substInfinitesimalQuadratic.simps(2)) next case (Eq p) define p' where "p' = (case convert_poly var p (xs'@ x#xs) of Some p' \ p')" have At'_simp : "At' = EqUni p'" using assms Eq using p'_def by auto show ?thesis unfolding At'_simp Eq substInfinitesimalQuadraticUni.simps substInfinitesimalQuadratic.simps using At'_simp Eq assms(1) convert_substNegInfinity assms(12) by fastforce next case (Leq p) define p' where "p' = (case convert_poly var p (xs'@ x#xs) of Some p' \ p')" have At'_simp : "At' = LeqUni p'" using assms Leq using p'_def by auto have allzero : "eval (allZero p var) (xs'@ x#xs) = evalUni (allZero' p') x" using Leq assms(1) convert_allZero p'_def assms(12) by force have less : "eval (quadratic_sub_fm var a b c d (convertDerivative var p)) (xs'@ x#xs) = evalUni (map_atomUni (quadraticSubUni A B C D) (convertDerivativeUni p')) x" using convert_substInfinitesimalQuadratic_less[OF _ assms(2-12)] by (metis Leq assms(1) convert_atom.simps(3) option.distinct(1) option.exhaust_sel option.map(1) option.the_def p'_def) show ?thesis unfolding At'_simp Leq substInfinitesimalQuadraticUni.simps substInfinitesimalQuadratic.simps eval.simps evalUni.simps using allzero less by auto next case (Neq p) define p' where "p' = (case convert_poly var p (xs'@ x#xs) of Some p' \ p')" have At'_simp : "At' = NeqUni p'" using assms Neq using p'_def by auto show ?thesis unfolding At'_simp Neq substInfinitesimalQuadraticUni.simps substInfinitesimalQuadratic.simps by (metis assms(12) At'_simp Neq assms(1) convert_substNegInfinity eval.simps(6) eval_neg substNegInfinityUni.simps(4) substNegInfinity.simps(4)) qed lemma infinitesimal_quad_helper: fixes A B C D:: "real" assumes at_is: "At = LessUni p \ At = EqUni p" assumes "D\0" assumes "C\0" shows "(\y'::real>((A+B * sqrt(C))/(D)). \x::real \{((A+B * sqrt(C))/(D))<..y'}. aEvalUni At x) = (evalUni (substInfinitesimalQuadraticUni A B C D At) x)" proof(cases "At = LessUni p") case True then have LessUni: "At = LessUni p" by auto then show ?thesis proof(cases p) case (fields a b c) show ?thesis proof(cases "2 * (a::real) * (A + B * sqrt C) / D + b = 0") case True then have True1 : "2 * a * (A + B * sqrt C) / D + b = 0" by auto show ?thesis proof(cases "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0") case True then have True2 : "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0" by auto then show ?thesis proof(cases "a<0") case True then show ?thesis unfolding LessUni fields apply (simp add:True1 True2 True) using True1 True2 True proof - assume beq: "2 * a * (A + B * sqrt C) / D + b = 0" assume root: "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0" assume alt: "a < 0 " let ?r = "-((A + B * sqrt C) / D)" have beq_var: "b = 2 * a * ?r" using beq by auto have root_var: " a * ?r^2 - 2*a*?r*?r + c = 0" using root by (simp add: beq_var) have "\y'>- ?r. \x\{- ?r<..y'}. a * x\<^sup>2 + 2 * a *?r * x + c < 0" using beq_var root_var alt one_root_a_lt0[where a = "a", where b="b", where c="c", where r="?r"] by auto then show "\y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" using beq_var by auto qed next case False then show ?thesis unfolding LessUni fields apply (simp add:True1 True2 False) using True1 True2 False proof clarsimp fix y' assume beq: " 2 * a * (A + B * sqrt C) / D + b = 0" assume root: " a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0" assume agteq: "\ a < 0 " assume y_prop: "(A + B * sqrt C) / D < y'" have beq_var: "b = 2 * a * (- A - B * sqrt C) / D" using beq - by (metis (no_types, hide_lams) ab_group_add_class.ab_diff_conv_add_uminus add.left_neutral add_diff_cancel_left' divide_inverse mult.commute mult_minus_right) + by (metis (no_types, opaque_lifting) ab_group_add_class.ab_diff_conv_add_uminus add.left_neutral add_diff_cancel_left' divide_inverse mult.commute mult_minus_right) have root_var: " a * ((- A - B * sqrt C) / D)\<^sup>2 - 2 * a * (- A - B * sqrt C) * (- A - B * sqrt C) / (D * D) + c = 0" using root proof - have f1: "\r ra. - ((r::real) + ra) = - r - ra" by auto have f2: "\r ra. r * (a * 2 * (- A - B * sqrt C)) / (ra * D) = r / (ra / b)" by (simp add: beq_var) have f3: "c - 0 + a * ((A + B * sqrt C) / D)\<^sup>2 = - (b * (A + B * sqrt C) / D)" using root by force have f4: "\r ra rb. ((- (r::real) - ra) / rb)\<^sup>2 = ((r + ra) / rb)\<^sup>2" using f1 by (metis (no_types) divide_minus_left power2_minus) have "\r ra rb rc. - ((r::real) * (ra + rb) / rc) = r * (- ra - rb) / rc" using f1 by (metis divide_divide_eq_right divide_minus_left mult.commute) then show ?thesis using f4 f3 f2 by (simp add: mult.commute) qed have y_prop_var: "- ((- A - B * sqrt C) / D) < y'" using y_prop by (metis add.commute diff_minus_eq_add divide_minus_left minus_diff_eq) have "\x\{- (- (A + B * sqrt C) / D)<..y'}. \ a * x\<^sup>2 + 2 * a * (- (A + B * sqrt C) / D) * x + c < 0" using y_prop_var beq_var root_var agteq one_root_a_gt0[where a = "a", where b ="b", where c = "c", where r= "-(A + B * sqrt C) / D"] by auto then show " \x\{(A + B * sqrt C) / D<..y'}. \ a * x\<^sup>2 + b * x + c < 0" proof - have f1: "2 * a * (A + B * sqrt C) * inverse D + b = 0" by (metis True1 divide_inverse) obtain rr :: real where f2: "rr \ {- (- (A + B * sqrt C) / D)<..y'} \ \ a * rr\<^sup>2 + 2 * a * (- (A + B * sqrt C) / D) * rr + c < 0" using \\x\{- (- (A + B * sqrt C) / D)<..y'}. \ a * x\<^sup>2 + 2 * a * (- (A + B * sqrt C) / D) * x + c < 0\ by blast have f3: "a * ((A + B * sqrt C) * (inverse D * 2)) = - b" using f1 by linarith have f4: "\r. - (- (r::real)) = r" by simp have f5: "\r ra. (ra::real) * - r = r * - ra" by simp have f6: "a * ((A + B * sqrt C) * (inverse D * - 2)) = b" using f3 by simp have f7: "\r ra rb. (rb::real) * (ra * r) = r * (rb * ra)" by auto have f8: "\r ra. - (ra::real) * r = ra * - r" by linarith then have f9: "a * (inverse D * ((A + B * sqrt C) * - 2)) = b" using f7 f6 f5 by presburger have f10: "rr \ {inverse D * (A + B * sqrt C)<..y'}" using f4 f2 by (metis (no_types) divide_inverse mult.commute mult_minus_right) have "\ c + (rr * b + a * rr\<^sup>2) < 0" using f9 f8 f7 f2 by (metis (no_types) add.commute divide_inverse mult.commute mult_minus_right) then show ?thesis using f10 by (metis (no_types) add.commute divide_inverse mult.commute) qed qed qed next case False then have False1 : "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c \ 0" by auto show ?thesis proof(cases "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c < 0") case True show ?thesis unfolding LessUni fields apply (simp add: True1 True) using True1 True proof - let ?r = "(A + B * sqrt C) / D" assume " 2 * a * (A + B * sqrt C) / D + b = 0" assume "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c < 0 " then have " \y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. poly [:c, b, a:] x < 0" using continuity_lem_lt0[where r= "(A + B * sqrt C) / D", where c = "c", where b = "b", where a="a"] quadratic_poly_eval[of c b a ?r] by auto then show "\y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" using quadratic_poly_eval[of c b a] by fastforce qed next case False then have False' : "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c > 0" using False1 by auto show ?thesis unfolding LessUni fields apply(simp add: True1 False False1) using True1 False' continuity_lem_gt0_expanded[where a = "a", where b = "b",where c = "c", where r = "((A + B * sqrt C) / D)"] by (metis mult_less_0_iff not_square_less_zero times_divide_eq_right) qed qed next case False then have False1 : "2 * a * (A + B * sqrt C) / D + b \ 0" by auto have c1: "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0 \ 2 * a * (A + B * sqrt C) / D + b < 0 \ \y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume root: "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0" assume blt: " 2 * a * (A + B * sqrt C) / D + b < 0" let ?r = "-(A + B * sqrt C) / D" have bltvar: "b < 2 * a * ?r" using blt divide_minus_left mult_2 mult_minus_right real_add_less_0_iff by (metis times_divide_eq_right) have rootvar: "a * ?r^2 - b * ?r + c = 0" using root proof - have f1: "\r ra. - (ra::real) * r = ra * - r" by simp have f2: "\r ra. ((ra::real) * - r)\<^sup>2 = (ra * r)\<^sup>2" by simp have f3: "a * (inverse D * (A - B * - sqrt C))\<^sup>2 - inverse D * (b * - (A - B * - sqrt C)) - - c = 0" by (metis (no_types) diff_minus_eq_add divide_inverse mult.commute mult_minus_left root) have f4: "\r ra rb. (rb::real) * (ra * r) = ra * (r * rb)" by simp have "\r ra. (ra::real) * - r = r * - ra" by simp then have "a * (inverse D * (A - B * - sqrt C))\<^sup>2 - b * (inverse D * - (A - B * - sqrt C)) - - c = 0" using f4 f3 f1 by (metis (no_types)) then have "a * (inverse D * - (A - B * - sqrt C))\<^sup>2 - b * (inverse D * - (A - B * - sqrt C)) - - c = 0" using f2 by presburger then show ?thesis by (simp add: divide_inverse mult.commute) qed have "\y'> ((A + B * sqrt C) / D). \x\{((A + B * sqrt C) / D)<..y'}. a * x\<^sup>2 + b * x + c < 0" using rootvar bltvar case_d1[where a= "a", where b = "b", where c = "c", where r = ?r] by (metis add.inverse_inverse divide_inverse mult_minus_left) then show ?thesis by blast qed have c2: " \y'. a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0 \ \ 2 * a * (A + B * sqrt C) / D + b < 0 \ (A + B * sqrt C) / D < y' \ \x\{(A + B * sqrt C) / D<..y'}. \ a * x\<^sup>2 + b * x + c < 0" proof - let ?r = "(A + B * sqrt C) / D" fix y' assume h1: "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c = 0" assume h2: "\ 2 * a * (A + B * sqrt C) / D + b < 0" assume h3: " (A + B * sqrt C) / D < y'" have eq: "2 * a * (A + B * sqrt C) / D + b = 0 \ \x\{(A + B * sqrt C) / D..y'}. \ a * x\<^sup>2 + b * x + c < 0" using False1 by blast have "2 * a * (A + B * sqrt C) / D + b > 0 \ \x\{?r<..y'}. \ a * x\<^sup>2 + b * x + c < 0" using case_d4[where a = "a", where b = "b", where c= "c", where r = "-?r"] h1 h2 h3 by auto then show "\x\{(A + B * sqrt C) / D<..y'}. \ a * x\<^sup>2 + b * x + c < 0" using h2 eq using False1 by linarith qed have c3: "((a::real) * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c < 0) \ (\y'>((A + B * sqrt C) / D). \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0)" proof clarsimp assume assump: "a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c < 0 " have "a * ((A + B * sqrt C) / D)\<^sup>2 + b * ((A + B * sqrt C) / D) + c < 0 \ \y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" using continuity_lem_lt0_expanded[where a= "a", where b = "b", where c = "c", where r = "((A + B * sqrt C) / D)::real"] by auto then have "\y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" using assump by auto then obtain y where y_prop: "y >(A + B * sqrt C) / D \ (\x\{(A + B * sqrt C) / D<..y}. a * x\<^sup>2 + b * x + c < 0)" by auto then have h: "\ k. k >(A + B * sqrt C) / D \ k < y" using dense by blast then obtain k where k_prop: "k >(A + B * sqrt C) / D \ k < y" by auto then have "\x\{(A + B * sqrt C) / D..k}. a * x\<^sup>2 + b * x + c < 0" using y_prop using assump by force then show "\y'>((A + B * sqrt C) / D::real). \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c < 0" using k_prop by auto qed have c4: "\y'. a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c \ 0 \ \ a * ((A + B * sqrt C) / D)\<^sup>2 + b * (A + B * sqrt C) / D + c < 0 \ (A + B * sqrt C) / D < y' \ \x\{(A + B * sqrt C) / D<..y'}. \ a * x\<^sup>2 + b * x + c < 0" using continuity_lem_gt0_expanded[where a= "a", where b = "b", where c = "c", where r= "(A + B * sqrt C) / D"] by (metis Groups.mult_ac(1) divide_inverse less_eq_real_def linorder_not_le) show ?thesis unfolding LessUni fields apply(simp add: False1) using c1 c2 c3 c4 by auto qed qed next case False then have EqUni: "At = EqUni p" using at_is by auto then show ?thesis proof(cases p) case (fields a b c) have " \y'. (A + B * sqrt C) / D < y' \ \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c = 0 \ (a = 0 \ b = 0 \ c = 0)" proof - fix y' assume "(A + B * sqrt C) / D < y'" then show " \x\{(A + B * sqrt C) / D<..y'}. a * x\<^sup>2 + b * x + c = 0 \ (a = 0 \ b = 0 \ c = 0)" using assms continuity_lem_eq0[where r = "(A + B * sqrt C) / D", where p = "y'", where a= "a", where b ="b", where c="c"] by auto qed then show ?thesis apply (auto simp add:EqUni fields ) using linordered_field_no_ub by blast qed qed lemma infinitesimal_quad: fixes A B C D:: "real" assumes "D\0" assumes "C\0" shows "(\y'::real>((A+B * sqrt(C))/(D)). \x::real \{((A+B * sqrt(C))/(D))<..y'}. aEvalUni At x) = (evalUni (substInfinitesimalQuadraticUni A B C D At) x)" proof(cases At) case (LessUni p) then show ?thesis using infinitesimal_quad_helper assms by blast next case (EqUni p) then show ?thesis using infinitesimal_quad_helper assms by blast next case (LeqUni p) then show ?thesis proof (cases p) case (fields a b c) have same: "\x. aEvalUni (LeqUni p) x = (aEvalUni (EqUni p) x) \ (aEvalUni (LessUni p) x)" apply (simp add: fields) by force let ?r = "(A + B * sqrt C) / D" have "\a b c. At = LeqUni p \ p = (a, b, c) \ (\y'>(A + B * sqrt C) / D. \x\{(A + B * sqrt C) / D<..y'}. aEvalUni At x) = evalUni (substInfinitesimalQuadraticUni A B C D At) x" proof - fix a b c assume atis: "At = LeqUni p" assume p_is: " p = (a, b, c)" have s1: "(\y'>?r. \x\{?r<..y'}. aEvalUni At x) = (\y'>?r. \x\{?r<..y'}. (aEvalUni (EqUni p) x) \ (aEvalUni (LessUni p) x))" using atis same aEvalUni.simps(2) aEvalUni.simps(3) fields less_eq_real_def by blast have s2: "... = (\y'>?r. \x\{?r<..y'}. (aEvalUni (EqUni p) x)) \ (\y'>?r. \x\{?r<..y'}. (aEvalUni (LessUni p) x))" using either_or[where r = "?r"] p_is by blast have eq1: "(\y'>?r. \x\{?r<..y'}. (aEvalUni (EqUni p) x)) = (evalUni (substInfinitesimalQuadraticUni A B C D (EqUni p)) x)" using infinitesimal_quad_helper[where At = "EqUni p", where p = "p", where B = "B", where C= "C", where A= "A", where D="D"] assms by auto have eq2: "(\y'>?r. \x\{?r<..y'}. (aEvalUni (LessUni p) x)) = (evalUni (substInfinitesimalQuadraticUni A B C D (LessUni p)) x)" using infinitesimal_quad_helper[where At = "LessUni p", where p = "p", where B = "B", where C= "C", where A= "A", where D="D"] assms by auto have z1: "(\y'>?r. \x\{?r<..y'}. aEvalUni At x) = ((evalUni (substInfinitesimalQuadraticUni A B C D (EqUni p)) x) \ (evalUni (substInfinitesimalQuadraticUni A B C D (LessUni p)) x))" using s1 s2 eq1 eq2 by auto have z2: "(evalUni (substInfinitesimalQuadraticUni A B C D (EqUni p)) x) \ (evalUni (substInfinitesimalQuadraticUni A B C D (LessUni p)) x) = evalUni (substInfinitesimalQuadraticUni A B C D (LeqUni p)) x" by auto have z3: "(evalUni (substInfinitesimalQuadraticUni A B C D At) x) = evalUni (substInfinitesimalQuadraticUni A B C D (LeqUni p)) x" using LeqUni by auto then have z4: "(evalUni (substInfinitesimalQuadraticUni A B C D (EqUni p)) x) \ (evalUni (substInfinitesimalQuadraticUni A B C D (LessUni p)) x) = (evalUni (substInfinitesimalQuadraticUni A B C D At) x) " using z2 z3 by auto let ?a = "(evalUni (substInfinitesimalQuadraticUni A B C D (EqUni p)) x) \ (evalUni (substInfinitesimalQuadraticUni A B C D (LessUni p)) x)" let ?b = "(\y'>?r. \x\{?r<..y'}. aEvalUni At x)" let ?c = "(evalUni (substInfinitesimalQuadraticUni A B C D At) x)" have t1: "?b = ?a" using z1 by auto have t2: "?a = ?c" using z4 using atis by auto then have "?b = ?c" using t1 t2 by auto then show "(\y'>?r. \x\{?r<..y'}. aEvalUni At x) = evalUni (substInfinitesimalQuadraticUni A B C D At) x" by auto qed then show ?thesis using LeqUni fields by blast qed next case (NeqUni p) then show ?thesis proof (cases p) case (fields a b c) then show ?thesis unfolding NeqUni fields using nonzcoeffs by auto qed qed end \ No newline at end of file diff --git a/thys/Virtual_Substitution/MPolyExtension.thy b/thys/Virtual_Substitution/MPolyExtension.thy --- a/thys/Virtual_Substitution/MPolyExtension.thy +++ b/thys/Virtual_Substitution/MPolyExtension.thy @@ -1,192 +1,192 @@ section "Multivariate Polynomials Extension" theory MPolyExtension imports Polynomials.Polynomials (*MPoly_Type_Efficient_Code*) Polynomials.MPoly_Type_Class_FMap begin subsection "Definition Lifting" lift_definition coeff_code::"'a::zero mpoly \ (nat \\<^sub>0 nat) \ 'a" is "lookup" . lemma coeff_code[code]: "coeff = coeff_code" unfolding coeff_def apply(transfer) by auto lemma coeff_transfer[transfer_rule]:\ \TODO: coeff should be defined via lifting, this gives us the illusion\ "rel_fun cr_mpoly (=) lookup coeff" using coeff_code.transfer[folded coeff_code] . lemmas coeff_add = coeff_add[symmetric] lemma plus_monom_zero[simp]: "p + MPoly_Type.monom m 0 = p" by transfer auto lift_definition monomials::"'a::zero mpoly \ (nat \\<^sub>0 nat) set" is "Poly_Mapping.keys::((nat\\<^sub>0nat) \\<^sub>0 'a) \ _ set" . lemma mpoly_induct [case_names monom sum]:\ \TODO: overwrites @{thm mpoly_induct}\ assumes monom:"\m a. P (MPoly_Type.monom m a)" and sum:"(\p1 p2 m a. P p1 \ P p2 \ p2 = (MPoly_Type.monom m a) \ m \ monomials p1 \ a \ 0 \ P (p1+p2))" shows "P p" using assms proof (induction p rule: mpoly_induct) case (sum p1 p2 m a) then show ?case by (cases "a = 0") (auto simp: monomials.rep_eq) qed simp value "monomials ((Var 0 + Const (3::int) * Var 1)^2)" lemma Sum_any_lookup_times_eq: "(\k. ((lookup (x::'a\\<^sub>0('b::comm_semiring_1)) (k::'a)) * ((f:: 'a\('b::comm_semiring_1)) k))) = (\k\keys x. (lookup x (k::'a)) * (f k))" by (subst Sum_any.conditionalize) (auto simp: in_keys_iff intro!: Sum_any.cong) lemma Prod_any_power_lookup_eq: "(\k::'a. f k ^ lookup (x::'a\\<^sub>0nat) k) = (\k\keys x. f k ^ lookup x k)" by (subst Prod_any.conditionalize) (auto simp: in_keys_iff intro!: Prod_any.cong) lemma insertion_monom: "insertion i (monom m a) = a * (\k\keys m. i k ^ lookup m k)" by transfer (auto simp: insertion_aux_def insertion_fun_def Sum_any_lookup_times_eq Prod_any_power_lookup_eq) lemma monomials_monom[simp]: "monomials (monom m a) = (if a = 0 then {} else {m})" by transfer auto lemma finite_monomials[simp]: "finite (monomials m)" by transfer auto lemma monomials_add_disjoint: "monomials (a + b) = monomials a \ monomials b" if "monomials a \ monomials b = {}" using that by transfer (auto simp add: keys_plus_eqI) lemma coeff_monom[simp]: "coeff (monom m a) m = a" by transfer simp lemma coeff_not_in_monomials[simp]: "coeff m x = 0" if "x \ monomials m" using that by transfer (simp add: in_keys_iff) code_thms insertion lemma insertion_code[code]: "insertion i mp = (\m\monomials mp. (coeff mp m) * (\k\keys m. i k ^ lookup m k))" proof (induction mp rule: mpoly_induct) case (monom m a) show ?case by (simp add: insertion_monom) next case (sum p1 p2 m a) have monomials_add: "monomials (p1 + p2) = insert m (monomials p1)" using sum.hyps by (auto simp: monomials_add_disjoint) have *: "coeff (monom m a) x = 0" if "x \ monomials p1" for x using sum.hyps that by (subst coeff_not_in_monomials) auto show ?case unfolding insertion_add monomials_add sum.IH using sum.hyps by (auto simp: coeff_add * algebra_simps) qed (* insertion f p takes in a mapping from natural numbers to the type of the polynomial and substitutes in the constant (f var) for each var variable in polynomial p *) code_thms insertion value "insertion (nth [1, 2.3]) ((Var 0 + Const (3::int) * Var 1)^2)" (* isolate_variable_sparse p var degree returns the coefficient of the term a*var^degree in polynomial p *) lift_definition isolate_variable_sparse::"'a::comm_monoid_add mpoly \ nat \ nat \ 'a mpoly" is "\(mp::'a mpoly) x d. sum (\m. monomial (coeff mp m) (Poly_Mapping.update x 0 m)) {m \ monomials mp. lookup m x = d}" . lemma Poly_Mapping_update_code[code]: "Poly_Mapping.update a b (Pm_fmap fm) = Pm_fmap (fmupd a b fm)" by (auto intro!: poly_mapping_eqI simp: update.rep_eq fmlookup_default_def) lemma monom_zero [simp] : "monom m 0 = 0" by (simp add: coeff_all_0) lemma remove_key_code[code]: "remove_key v (Pm_fmap fm) = Pm_fmap (fmdrop v fm)" by (auto simp: remove_key_lookup fmlookup_default_def intro!: poly_mapping_eqI) lemma extract_var_code[code]: "extract_var p v = (\m\monomials p. monom (remove_key v m) (monom (Poly_Mapping.single v (lookup m v)) (coeff p m)))" apply (rule extract_var_finite_set[where S="monomials p"]) using coeff_not_in_monomials by auto value "extract_var ((Var 0 + Const (3::real) * Var 1)^2) 0" (* degree p var takes in polynomial p and a variable var and finds the degree of that variable in the polynomial missing code theorems? still manages to evaluate *) code_thms degree value "degree ((Var 0 + Const (3::real) * Var 1)^2) 0" (* this function gives a set of all the variables in the polynomial *) lemma vars_code[code]: "vars p = \ (keys ` monomials p)" unfolding monomials.rep_eq vars_def .. value "vars ((Var 0 + Const (3::real) * Var 1)^2)" (* return true if the polynomial contains no variables *) fun is_constant :: "'a::zero mpoly \ bool" where "is_constant p = Set.is_empty (vars p)" value "is_constant (Const (0::int))" (* if the polynomial is constant, returns the real value associated with the polynomial, otherwise returns none *) fun get_if_const :: "real mpoly \ real option" where "get_if_const p = (if is_constant p then Some (coeff p 0) else None)" term "coeff p 0" -lemma insertionNegative : "insertion f p = - insertion f (-p)" - by (metis (no_types, hide_lams) add_eq_0_iff cancel_comm_monoid_add_class.diff_cancel insertion_add insertion_zero uminus_add_conv_diff) +lemma insertionNegative : "insertion f p = - insertion f (-p)" try + by (metis add.right_inverse eq_neg_iff_add_eq_0 insertion_add insertion_zero) definition derivative :: "nat \ real mpoly \ real mpoly" where "derivative x p = (\i\{0..degree p x-1}. isolate_variable_sparse p x (i+1) * (Var x)^i * (Const (i+1)))" text "get\\_coeffs $x$ $p$ gets the tuple of coefficients $a$ $b$ $c$ of the term $a*x^2+b*x+c$ in polynomial $p$" fun get_coeffs :: "nat \ real mpoly \ real mpoly * real mpoly * real mpoly" where "get_coeffs var x = ( isolate_variable_sparse x var 2, isolate_variable_sparse x var 1, isolate_variable_sparse x var 0) " end diff --git a/thys/Virtual_Substitution/NegInfinityUni.thy b/thys/Virtual_Substitution/NegInfinityUni.thy --- a/thys/Virtual_Substitution/NegInfinityUni.thy +++ b/thys/Virtual_Substitution/NegInfinityUni.thy @@ -1,632 +1,632 @@ theory NegInfinityUni imports UniAtoms NegInfinity QE begin fun allZero' :: "real * real * real \ atomUni fmUni" where "allZero' (a,b,c) = AndUni(AndUni(AtomUni(EqUni(0,0,a))) (AtomUni(EqUni(0,0,b)))) (AtomUni(EqUni(0,0,c)))" lemma convert_allZero : assumes "convert_poly var p (xs'@x#xs) = Some p'" assumes "length xs' = var" shows "eval (allZero p var) (xs'@x#xs) = evalUni (allZero' p') x" proof(cases p') case (fields a b c) then show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True then show ?thesis using assms fields by(simp add:eval_list_conj isovar_greater_degree) next case False then have nonzero : "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True then show ?thesis using assms fields apply(simp add:eval_list_conj isovar_greater_degree) by (metis) next case False then have degree2 : "MPoly_Type.degree p var = 2" using degree_convert_eq[OF assms(1)] nonzero by auto then show ?thesis using assms apply(simp add:eval_list_conj isovar_greater_degree) using insertion_isovarspars_free list_update_code(2) apply auto by (metis One_nat_def Suc_1 less_2_cases less_Suc_eq numeral_3_eq_3) qed qed qed fun alternateNegInfinity' :: "real * real * real \ atomUni fmUni" where "alternateNegInfinity' (a,b,c) = OrUni(AtomUni(LessUni(0,0,a)))( AndUni(AtomUni(EqUni(0,0,a))) ( OrUni(AtomUni(LessUni(0,0,-b)))( AndUni(AtomUni(EqUni(0,0,b)))( AtomUni(LessUni(0,0,c)) )) )) " lemma convert_alternateNegInfinity : assumes "convert_poly var p (xs'@x#xs) = Some X" assumes "length xs' = var" shows "eval (alternateNegInfinity p var) (xs'@x#xs) = evalUni (alternateNegInfinity' X) x" proof(cases X) case (fields a b c) then show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True then show ?thesis using assms apply (simp add: isovar_greater_degree) apply auto apply (metis aEval.simps(2) eval.simps(1) eval_and eval_false eval_or mult_one_left) by (metis aEval.simps(2) eval.simps(1) eval_or mult_one_left) next case False then have nonzero : "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True have letbind: "eval (let a_n = isolate_variable_sparse p var (Suc 0) in or (fm.Atom (Less (Const (- 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var 0 in or (fm.Atom (Less (Const 1 * a_n))) (and (fm.Atom (Eq a_n)) FalseF)))) (xs'@x#xs) = eval (or (fm.Atom (Less (Const (- 1) * (isolate_variable_sparse p var (Suc 0))))) (and (fm.Atom (Eq (isolate_variable_sparse p var (Suc 0)))) (or (fm.Atom (Less (Const 1 * (isolate_variable_sparse p var 0)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 0))) FalseF)))) (xs'@x#xs)" by meson show ?thesis using assms True unfolding fields by (simp add: isovar_greater_degree letbind eval_or eval_and insertion_mult insertion_const) next case False then have degree2 : "MPoly_Type.degree p var = 2" using degree_convert_eq[OF assms(1)] nonzero by auto have "[0..<3] = [0,1,2]" by (simp add: upt_rec) then have unfold : " (foldl (\F i. let a_n = isolate_variable_sparse p var i in or (fm.Atom (Less ((if i mod 2 = 0 then Const 1 else Const (- 1)) * a_n))) (and (fm.Atom (Eq a_n)) F)) FalseF [0..<3]) = (let a_n = isolate_variable_sparse p var 2 in or (fm.Atom (Less ((Const 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var (Suc 0) in or (fm.Atom (Less (Const (- 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var 0 in or (fm.Atom (Less (Const 1 * a_n))) (and (fm.Atom (Eq a_n)) FalseF))))))" by auto have letbind : "eval (foldl (\F i. let a_n = isolate_variable_sparse p var i in or (fm.Atom (Less ((if i mod 2 = 0 then Const 1 else Const (- 1)) * a_n))) (and (fm.Atom (Eq a_n)) F)) FalseF [0..<3]) (xs'@x#xs) = eval (or (fm.Atom (Less ( Const 1 * (isolate_variable_sparse p var 2)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 2))) (or (fm.Atom (Less (Const (- 1) * (isolate_variable_sparse p var (Suc 0))))) (and (fm.Atom (Eq (isolate_variable_sparse p var (Suc 0)))) (or (fm.Atom (Less (Const 1 * (isolate_variable_sparse p var 0)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 0))) FalseF)))))) (xs'@x#xs)" apply (simp add:unfold) by metis show ?thesis using degree2 assms unfolding fields by (simp add: isovar_greater_degree eval_or eval_and letbind insertion_mult insertion_const) qed qed qed fun substNegInfinityUni :: "atomUni \ atomUni fmUni" where "substNegInfinityUni (EqUni p) = allZero' p " | "substNegInfinityUni (LessUni p) = alternateNegInfinity' p"| "substNegInfinityUni (LeqUni p) = OrUni (alternateNegInfinity' p) (allZero' p)"| "substNegInfinityUni (NeqUni p) = negUni (allZero' p)" lemma convert_substNegInfinity : assumes "convert_atom var A (xs'@x#xs) = Some(A')" assumes "length xs' = var" shows "eval (substNegInfinity var A) (xs'@x#xs) = evalUni (substNegInfinityUni A') x" using assms proof(cases A) case (Less p) have "\X. convert_poly var p (xs' @ x # xs) = Some X" using assms Less apply(cases "MPoly_Type.degree p var < 3") by (simp_all) then obtain X where X_def: "convert_poly var p (xs' @ x # xs) = Some X" by auto then have A' : "A' = LessUni X" using assms Less apply(cases "MPoly_Type.degree p var < 3") by (simp_all) show ?thesis unfolding Less substNegInfinity.simps unfolding convert_alternateNegInfinity[OF X_def assms(2)] A' apply(cases X) by simp next case (Eq p) then show ?thesis using assms convert_allZero by auto next case (Leq p) define p' where "p' = (case convert_poly var p (xs'@x#xs) of Some p' \ p')" have A'_simp : "A' = LeqUni p'" using assms Leq using p'_def by auto have allZ : "eval (allZero p var) (xs'@x#xs) = evalUni (allZero' p') x" using convert_allZero using Leq assms p'_def by auto have altNeg : "eval (alternateNegInfinity p var) (xs'@x#xs) = evalUni (alternateNegInfinity' p') x" using convert_alternateNegInfinity using Leq assms p'_def by auto show ?thesis unfolding Leq substNegInfinity.simps eval_Or A'_simp substNegInfinityUni.simps evalUni.simps using allZ altNeg by auto next case (Neq p) then show ?thesis using assms convert_allZero convert_neg by auto qed lemma change_eval_eq: fixes x y:: "real" assumes "((aEvalUni (EqUni(a,b,c)) x \ aEvalUni (EqUni(a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms by auto lemma change_eval_lt: fixes x y:: "real" assumes "((aEvalUni (LessUni (a,b,c)) x \ aEvalUni (LessUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - let ?p = "[:c, b, a:]" have "sign ?p x \ sign ?p y" using assms unfolding sign_def apply (simp add: distrib_left mult.commute mult.left_commute power2_eq_square) by linarith then have "(\w \ (root_list ?p). x \ w \ w \ y)" using changes_sign assms by auto then obtain w where w_prop: "w \ (root_list ?p) \ x \ w \ w \ y" by auto then have "a*w^2 + b*w + c = 0" unfolding root_list_def using add.commute distrib_right mult.assoc mult.commute power2_eq_square using quadratic_poly_eval by force then show ?thesis using w_prop by auto qed lemma no_change_eval_lt: fixes x y:: "real" assumes "x < y" assumes "\(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" shows "((aEvalUni (LessUni (a,b,c)) x = aEvalUni (LessUni (a,b,c)) y))" using change_eval_lt using assms(1) assms(2) by blast lemma change_eval_neq: fixes x y:: "real" assumes "((aEvalUni (NeqUni (a,b,c)) x \ aEvalUni (NeqUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms by auto lemma change_eval_leq: fixes x y:: "real" assumes "((aEvalUni (LeqUni (a,b,c)) x \ aEvalUni (LeqUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - let ?p = "[:c, b, a:]" have "sign ?p x \ sign ?p y" using assms unfolding sign_def apply (simp add: distrib_left mult.commute mult.left_commute power2_eq_square) by linarith then have "(\w \ (root_list ?p). x \ w \ w \ y)" using changes_sign assms by auto then obtain w where w_prop: "w \ (root_list ?p) \ x \ w \ w \ y" by auto then have "a*w^2 + b*w + c = 0" unfolding root_list_def using add.commute distrib_right mult.assoc mult.commute power2_eq_square using quadratic_poly_eval by force then show ?thesis using w_prop by auto qed lemma change_eval: fixes x y:: "real" fixes At:: "atomUni" assumes xlt: "x < y" assumes noteq: "((aEvalUni At) x \ aEvalUni (At) y)" assumes "getPoly At = (a, b, c)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - have four_types: "At = (EqUni (a,b,c)) \ At = (LessUni (a,b,c)) \ At = (LeqUni (a,b,c)) \ At = (NeqUni (a,b,c))" using getPoly.elims assms(3) by force have eq_h: "At = (EqUni (a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms(1) assms(2) change_eval_eq by blast have less_h: "At = (LessUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_lt assms by blast have leq_h: "At = (LeqUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_leq assms by blast have neq_h: "At = (NeqUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_neq assms by blast show ?thesis using four_types eq_h less_h leq_h neq_h by blast qed lemma no_change_eval: fixes x y:: "real" assumes "getPoly At = (a, b, c)" assumes "x < y" assumes "\(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" shows "((aEvalUni At) x = aEvalUni (At) y \ x < y)" using change_eval using assms(1) assms(2) assms(3) by blast lemma same_eval'' : assumes "getPoly At = (a, b, c)" assumes nonz: "a \ 0 \ b \ 0 \ c \ 0" shows "\x. \yy. poly ?p y = a*y^2 + b*y + c" by (simp add: distrib_left power2_eq_square) have "?p \ 0" using nonz by auto then have "finite {y. poly ?p y = 0}" using poly_roots_finite by blast then have "finite {y. c + (a * y\<^sup>2 + b * y) = 0} \ \y. y * (b + y * a) = a * y\<^sup>2 + b * y \ finite {y. a * y\<^sup>2 + b * y + c = 0}" proof - assume a1: "finite {y. c + (a * y\<^sup>2 + b * y) = 0}" have "\x0. c + (a * x0\<^sup>2 + b * x0) = a * x0\<^sup>2 + b * x0 + c" by simp then show ?thesis using a1 by presburger qed then have finset: "finite {y. a*y^2 + b*y + c = 0}" using poly_eval by (metis \finite {y. poly [:c, b, a:] y = 0}\ poly_roots_set_same) then have "\x. (\y. a*y^2 + b*y + c = 0 \ x < y)" proof - let ?l = "sorted_list_of_set {y. a*y^2 + b*y + c = 0}" have "\x. x < ?l ! 0" using infzeros nonz by blast then obtain x where x_prop: "x < ?l! 0" by auto then have "\ y. List.member ?l y \ x < y" proof clarsimp fix y assume "List.member ?l y" then have "\n. n < length ?l \ ?l ! n = y" by (meson in_set_conv_nth in_set_member) then obtain n where n_prop: "n < length ?l \ ?l ! n = y" by auto have h: "\n < length ?l. ?l ! n \ ?l !0" using sorted_iff_nth_mono using sorted_sorted_list_of_set by blast then have h: "y \ ?l ! 0" using n_prop by auto then show "x < y" using x_prop by auto qed then show ?thesis using finset set_sorted_list_of_set in_set_member by (metis (mono_tags, lifting) mem_Collect_eq) qed then obtain x where x_prop: "(\y. a*y^2 + b*y + c = 0 \ x < y)" by auto then have same_as: "\yx0. (x0 < x) = (\ 0 \ x0 + - 1 * x)" by auto have f2: "(0 \ - 1 * x + v0_0) = (x + - 1 * v0_0 \ 0)" by auto have f3: "v0_0 + - 1 * x = - 1 * x + v0_0" by auto have f4: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" by auto have f5: "\x3 x4 x5. (aEvalUni x3 x5 \ aEvalUni x3 x4) = ((\ aEvalUni x3 x5) = aEvalUni x3 x4)" by fastforce have f6: "\x0 x1 x2 x3 x4 x5. (x5 < x4 \ (\ aEvalUni x3 x5) = aEvalUni x3 x4 \ getPoly x3 = (x2, x1, x0) \ (\v6\x5. v6 \ x4 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0)) = ((\ x5 < x4 \ (\ aEvalUni x3 x5) \ aEvalUni x3 x4 \ getPoly x3 \ (x2, x1, x0)) \ (\v6\x5. v6 \ x4 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0))" by fastforce have f7: "\x0 x5. ((x0::real) \ x5) = (x0 + - 1 * x5 \ 0)" by auto have f8: "\x0 x6. ((x6::real) \ x0) = (0 \ x0 + - 1 * x6)" by auto have "\x4 x5. ((x5::real) < x4) = (\ x4 + - 1 * x5 \ 0)" by auto then have "(\r ra a rb rc rd. r < ra \ aEvalUni a r \ aEvalUni a ra \ getPoly a = (rb, rc, rd) \ (\re\r. re \ ra \ rb * re\<^sup>2 + rc * re + rd = 0)) = (\r ra a rb rc rd. (ra + - 1 * r \ 0 \ (\ aEvalUni a r) \ aEvalUni a ra \ getPoly a \ (rb, rc, rd)) \ (\re. 0 \ re + - 1 * r \ re + - 1 * ra \ 0 \ rd + rb * re\<^sup>2 + rc * re = 0))" using f8 f7 f6 f5 f4 by presburger then have f9: "\r ra a rb rc rd. (ra + - 1 * r \ 0 \ (\ aEvalUni a r) \ aEvalUni a ra \ getPoly a \ (rb, rc, rd)) \ (\re. 0 \ re + - 1 * r \ re + - 1 * ra \ 0 \ rd + rb * re\<^sup>2 + rc * re = 0)" by (meson change_eval) obtain rr :: "real \ real \ real \ real \ real \ real" where "\x0 x1 x2 x4 x5. (\v6. 0 \ v6 + - 1 * x5 \ v6 + - 1 * x4 \ 0 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0) = (0 \ rr x0 x1 x2 x4 x5 + - 1 * x5 \ rr x0 x1 x2 x4 x5 + - 1 * x4 \ 0 \ x0 + x2 * (rr x0 x1 x2 x4 x5)\<^sup>2 + x1 * rr x0 x1 x2 x4 x5 = 0)" by moura then have f10: "\r ra a rb rc rd. ra + - 1 * r \ 0 \ aEvalUni a r = aEvalUni a ra \ getPoly a \ (rb, rc, rd) \ r + - 1 * rr rd rc rb ra r \ 0 \ 0 \ ra + - 1 * rr rd rc rb ra r \ rd + rb * (rr rd rc rb ra r)\<^sup>2 + rc * rr rd rc rb ra r = 0" using f9 by simp have f11: "(rr c b a x v0_0 + - 1 * x \ 0) = (0 \ x + - 1 * rr c b a x v0_0)" by force have "\x0. (x < x0) = (\ x0 + - 1 * x \ 0)" by auto then have f12: "\r. c + a * r\<^sup>2 + b * r \ 0 \ \ r + - 1 * x \ 0" using x_prop by auto obtain rra :: real where "(\v0. \ 0 \ v0 + - 1 * x \ aEvalUni At v0 \ aEvalUni At x) = (\ 0 \ rra + - 1 * x \ aEvalUni At rra \ aEvalUni At x)" by moura then show ?thesis using f12 f11 f10 f3 f2 f1 proof - have f1: "\x0. (x0 < x) = (\ 0 \ x0 + - 1 * x)" by auto have f2: "(0 \ v0_0a + - 1 * x) = (x + - 1 * v0_0a \ 0)" by auto have f3: "(rr c b a x v0_0a + - 1 * x \ 0) = (0 \ x + - 1 * rr c b a x v0_0a)" by auto obtain rrb :: real where "(\v0. \ 0 \ v0 + - 1 * x \ aEvalUni At v0 \ aEvalUni At x) = (\ 0 \ rrb + - 1 * x \ aEvalUni At rrb \ aEvalUni At x)" by blast then show ?thesis using f3 f2 f1 assms(1) f10 f12 by smt qed qed then show ?thesis by auto qed lemma inequality_case : "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) < 0) = (a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0))" proof- let ?At = "(LessUni (a, b, c))" have firsth : "\x. (\y2 + b * y + c < 0 \ a\0)" proof - fix x assume lt: "\y2 + b * y + c < 0" have "\w. \y < w. y^2 > (-b/a)*y - c/a" using ysq_dom_y_plus_coeff[where b = "-b/a", where c = "-c/a"] by auto then have gtdomhelp: "a > 0 \ \w. \y < w. a*y^2 > a*((-b/a)*y - c/a)" by auto have "\y. (a > 0 \ a*((-b/a)*y - c/a) = -b*y - c)" by (simp add: right_diff_distrib') then have gtdom: "a > 0 \ \w.\y < w. a*y^2 > -b*y - c" using gtdomhelp by simp then have " a > 0 \ False" proof - assume "a > 0" then have "\w.\y < w. a*y^2 > -b*y - c" using gtdom by auto then obtain w where w_prop: "\y < w. a*y^2 + b*y + c > 0" by fastforce let ?mn = "min w x - 1" have gtz: "a*?mn^2 + b*?mn + c > 0" using w_prop by auto have ltz: "a*?mn^2 + b*?mn + c < 0" using lt by auto then show "False" using gtz ltz by auto qed then show "a \ 0" by fastforce qed have bleq0 : "\x. (\y b\0)" proof - fix x assume lt: "\y \w.\y < w. b*y > - c" by (metis mult.commute neg_less_divide_eq) then have "b < 0 \ False" proof - assume "b < 0" then have "\w.\y < w. b*y > - c" using gtdom by auto then obtain w where w_prop: "\y < w .b*y + c > 0" by fastforce let ?mn = "min w x - 1" have gtz: "b*?mn + c > 0" using w_prop by auto have ltz: "b*?mn + c < 0" using lt by auto then show "False" using gtz ltz by auto qed then show "b \ 0" by fastforce qed have secondh: "\x. (\y2 + b * y + c < 0 \ \ a < 0 \ \ 0 < b \ b = 0)" using firsth bleq0 by (metis add.commute add.right_neutral less_eq_real_def mult_zero_class.mult_zero_left) have thirdh : "\x. \y2 + b * y + c < 0 \ \ a < 0 \ \ 0 < b \ c < 0" using firsth secondh add.commute add.right_neutral infzeros mult_zero_class.mult_zero_left not_numeral_le_zero order_refl by (metis less_eq_real_def) have fourthh : "a < 0 \ \x. \y2 + b * y + c < 0" proof - assume aleq: "a < 0" have "\(w::real). \(y::real). (y < w \ y^2 > (-b/a)*y + (-c/a))" using ysq_dom_y_plus_coeff[where b = "-b/a", where c = "-c/a"] by blast then have hyp:"\(w::real). \(y::real). (y < w \ a*y^2 \ a*(-b/a)*y + a*(-c/a))" - by (metis (no_types, hide_lams) \a < 0\ distrib_left less_eq_real_def linorder_not_le mult.assoc mult_less_cancel_left) + by (metis (no_types, opaque_lifting) \a < 0\ distrib_left less_eq_real_def linorder_not_le mult.assoc mult_less_cancel_left) have "\y. a*(-b/a)*y + a*(-c/a) = -b*y -c" using \a < 0\ by auto then have "\(w::real). \(y::real). (y < w \ a*y^2 \ -b*y - c)" using hyp by auto then have "\(w::real). \(y::real). (y < w \ a*y^2 + b*y + c \ 0)" by (metis add.commute add_uminus_conv_diff le_diff_eq mult_minus_left real_add_le_0_iff) then obtain w where w_prop: "\(y::real). (y < w \ a*y^2 + b*y + c \ 0)" by auto have "\x. \y < x. aEvalUni ?At x = aEvalUni ?At y" using same_eval'' proof - have f1: "\x0 x1. ((x0::real) < x1) = (\ 0 \ x0 + - 1 * x1)" by linarith have "a \ 0" using \a < 0\ by force then obtain rr :: "atomUni \ real" where "\r. 0 \ r + - 1 * rr (LessUni (a, b, c)) \ aEvalUni (LessUni (a, b, c)) r = aEvalUni (LessUni (a, b, c)) (rr (LessUni (a, b, c)))" using f1 by (metis getPoly.simps(4) same_eval'') then show ?thesis using f1 by meson qed then obtain x where x_prop: "\y < x. aEvalUni ?At x = aEvalUni ?At y" by auto let ?mn = "min x w - 1" have "\y < ?mn. aEvalUni ?At y = True \ aEvalUni ?At y = False" using x_prop by auto have "\ y < ?mn. aEvalUni ?At y = False \ a*y^2 + b*y + c \ 0" by auto then have "\y. \y2 + b * y + c \ 0 \ y < min x w - 1 \ \ a * y\<^sup>2 + b * y + c < 0 \ a * y\<^sup>2 + b * y + c = 0" proof - fix y :: real assume a1: "y < min x w - 1" assume a2: "\ a * y\<^sup>2 + b * y + c < 0" assume a3: "\y2 + b * y + c \ 0" have "y < w" using a1 by linarith then show "a * y\<^sup>2 + b * y + c = 0" using a3 a2 less_eq_real_def by blast qed then have "\ y < ?mn. aEvalUni ?At y = False \ a*y^2 + b*y + c = 0" using w_prop by auto then have "\ y < ?mn. aEvalUni ?At y = False \ False" using infzeros aleq by (metis power_zero_numeral zero_less_power2) then have "\ y < ?mn. aEvalUni ?At y = True" proof - { fix rr :: real have "\r ra. (ra::real) < r \ \ ra < r + - 1" by linarith then have "\ rr < min x w - 1 \ aEvalUni (LessUni (a, b, c)) rr" by (metis (no_types) \\y False\ ab_group_add_class.ab_diff_conv_add_uminus less_eq_real_def min_less_iff_disj not_le x_prop) } then show ?thesis by blast qed then show ?thesis by auto qed have fifthh : "b > 0 \ \x. \y 0" show "\x. \y(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0) = (a > 0 \ a = 0 \ (0 > b \ b = 0 \ c > 0))" proof - have s1: "\y. - 1 * a * y\<^sup>2 + - 1 * b * y + - 1 * c < 0 \ a * y\<^sup>2 + b * y + c > 0" by auto have s2: "(- 1 * a < 0 \ - 1 * a = 0 \ (0 < - 1 * b \ - 1 * b = 0 \ - 1 * c < 0)) \ (a > 0 \ a = 0 \ (0 > b \ b = 0 \ c > 0)) " by auto have "(\x. \y2 + - 1 * b * y + - 1 * c < 0) = (- 1 * a < 0 \ - 1 * a = 0 \ (0 < - 1 * b \ - 1 * b = 0 \ - 1 * c < 0))" using inequality_case[where a = "-1*a", where b = "-1*b", where c= "-1*c"] by auto then show ?thesis using s1 s2 by auto qed lemma infinity_evalUni_LessUni : "(\x. \yx. \yx. \yx. \yx. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0)) \ (\x. \y2 + b * y + c \ 0)" using less_eq_real_def by auto have h2: "(\x. \y2 + b * y + c \ 0) \ ((\x. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0))" proof - assume a1: "(\x. \y2 + b * y + c \ 0)" have "\(\x. \y2 + b * y + c = 0) \ (\x. \y2 + b * y + c < 0) " proof - assume a2: "\(\x. \y2 + b * y + c = 0)" then have "a \ 0 \ b \ 0 \ c \ 0" by auto then have "(a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0))" proof - have x1: "a > 0 \ False" proof - assume "a > 0" then have "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0)" using inequality_case_geq by auto then show ?thesis using a1 by (meson a2 linorder_not_le min_less_iff_conj) qed then have x2: "a = 0 \ 0 > b \ False" proof - assume "a = 0 \ 0 > b" then have "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0)" using inequality_case_geq by blast then show ?thesis using a1 by (meson a2 linorder_not_le min_less_iff_conj) qed then have x3: "a = 0 \ b = 0 \ c > 0 \ False " using a1 a2 by auto show ?thesis using x1 x2 x3 by (meson \a \ 0 \ b \ 0 \ c \ 0\ linorder_neqE_linordered_idom) qed then show "(\x. \y2 + b * y + c < 0)" using inequality_case by auto qed then show ?thesis by auto qed have "(\x. \y2 + b * y + c \ 0) = (\x. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0)" using h1 h2 by auto then show "(\x. \y2 + b * y + c \ 0) = (a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0) \ a = 0 \ b = 0 \ c = 0)" using inequality_case[of a b c] infzeros[of _ a b c] by auto qed qed text "This is the vertical translation for substNegInfinityUni where we represent the virtual substution of negative infinity in the univariate case" lemma infinity_evalUni : shows "(\x. \y real \ int" where "sign p x \ (if poly p x = 0 then 0 else (if poly p x > 0 then 1 else -1))" definition sign_num:: "real \ int" where "sign_num x \ (if x = 0 then 0 else (if x > 0 then 1 else -1))" definition root_list:: "real Polynomial.poly \ real set" where "root_list p \ ({(x::real). poly p x = 0}::real set)" definition root_set:: "(real \ real \ real) set \ real set" where "root_set les \ ({(x::real). (\ (a, b, c) \ les. a*x^2 + b*x + c = 0)}::real set)" definition sorted_root_list_set:: "(real \ real \ real) set \ real list" where "sorted_root_list_set p \ sorted_list_of_set (root_set p)" definition nonzero_root_set:: "(real \ real \ real) set \ real set" where "nonzero_root_set les \ ({(x::real). (\ (a, b, c) \ les. (a \ 0 \ b \ 0) \ a*x^2 + b*x + c = 0)}::real set)" definition sorted_nonzero_root_list_set:: "(real \ real \ real) set \ real list" where "sorted_nonzero_root_list_set p \ sorted_list_of_set (nonzero_root_set p)" (* Important property of sorted lists *) lemma sorted_list_prop: fixes l::"real list" fixes x::"real" assumes sorted: "sorted l" assumes lengt: "length l > 0" assumes xgt: "x > l ! 0" assumes xlt: "x \ l ! (length l - 1)" shows "\n. (n+1) < (length l) \ x \ l !n \ x \ l ! (n + 1)" proof - have "\(\n. (n+1) < (length l) \ x \ l !n \ x \ l ! (n + 1)) \ False" proof clarsimp fix n assume alln: "\n. l ! n \ x \ Suc n < length l \ \ x \ l ! Suc n" have "\k. (k < length l \ x > l ! k)" proof clarsimp fix k show "k < length l \ l ! k < x" proof (induct k) case 0 then show ?case using xgt by auto next case (Suc k) then show ?case using alln using less_eq_real_def by auto qed qed then show "False" using xlt diff_Suc_less lengt not_less by (metis One_nat_def) qed then show ?thesis by auto qed subsection "Quadratic polynomial properties" lemma quadratic_poly_eval: fixes a b c::"real" fixes x::"real" shows "poly [:c, b, a:] x = a*x^2 + b*x + c" proof - have "x * (b + x * a) = a * x\<^sup>2 + b * x" by (metis add.commute distrib_right mult.assoc mult.commute power2_eq_square) then show ?thesis by auto qed lemma poly_roots_set_same: fixes a b c:: "real" shows "{(x::real). a * x\<^sup>2 + b * x + c = 0} = {x. poly [:c, b, a:] x = 0}" proof - have "\x. a*x^2 + b*x + c = poly [:c, b, a:] x" proof clarsimp fix x show "a * x\<^sup>2 + b * x = x * (b + x * a)" using quadratic_poly_eval[of c b a x] by auto qed then show ?thesis by auto qed lemma root_set_finite: assumes fin: "finite les" assumes nex: "\(\ (a, b, c) \ les. a = 0 \ b = 0 \ c = 0 )" shows "finite (root_set les)" proof - have "\(a, b, c) \ les. finite {(x::real). a*x^2 + b*x + c = 0}" proof clarsimp fix a b c assume "(a, b, c) \ les" then have "[:c, b, a:] \ 0" using nex by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show "finite {x. a * x\<^sup>2 + b * x + c = 0}" using poly_roots_set_same by auto qed then show ?thesis using fin unfolding root_set_def by auto qed lemma nonzero_root_set_finite: assumes fin: "finite les" shows "finite (nonzero_root_set les)" proof - have "\(a, b, c) \ les. (a \ 0 \ b \ 0) \ finite {(x::real). a*x^2 + b*x + c = 0}" proof clarsimp fix a b c assume ins: "(a, b, c) \ les" assume "a = 0 \ b \ 0" then have "[:c, b, a:] \ 0" using ins by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show "finite {x. a * x\<^sup>2 + b * x + c = 0}" using poly_roots_set_same by auto qed then show ?thesis using fin unfolding nonzero_root_set_def by auto qed lemma discriminant_lemma: fixes a b c r::"real" assumes aneq: "a \ 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2 * a * r*r + c = 0" shows "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" proof - have "c = a*r^2" using root by (simp add: power2_eq_square) then have same: "b^2 - 4*a*c = (2 * a * r)^2 - 4*a*(a*r^2)" using beq by blast have "(2 * a * r)^2 = 4*a^2*r^2" by (simp add: mult.commute power2_eq_square) then have "(2 * a * r)^2 - 4*a*(a*(r)^2) = 0" using power2_eq_square by auto then have "b^2 - 4*a*c = 0" using same by simp then have "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -b / (2 * a)" using discriminant_zero aneq unfolding discrim_def by auto then show ?thesis using beq by (simp add: aneq) qed (* Show a polynomial only changes sign when it passes through a root *) lemma changes_sign: fixes p:: "real Polynomial.poly" shows "\x::real. \ y::real. ((sign p x \ sign p y \ x < y) \ (\c \ (root_list p). x \ c \ c \ y))" proof clarsimp fix x y assume "sign p x \ sign p y" assume "x < y" then show "\c\root_list p. x \ c \ c \ y" using poly_IVT_pos[of x y p] poly_IVT_neg[of x y p] by (metis (mono_tags) \sign p x \ sign p y\ less_eq_real_def linorder_neqE_linordered_idom mem_Collect_eq root_list_def sign_def) qed (* Show a polynomial only changes sign if it passes through a root *) lemma changes_sign_var: fixes a b c x y:: "real" shows "((sign_num (a*x^2 + b*x + c) \ sign_num (a*y^2 + b*y + c) \ x < y) \ (\q. (a*q^2 + b*q + c = 0 \ x \ q \ q \ y)))" proof clarsimp assume sn: "sign_num (a * x\<^sup>2 + b * x + c) \ sign_num (a * y\<^sup>2 + b * y + c)" assume xy: " x < y" let ?p = "[:c, b, a:]" have cs: "((sign ?p x \ sign ?p y \ x < y) \ (\c \ (root_list ?p). x \ c \ c \ y))" using changes_sign[of ?p] by auto have "(sign ?p x \ sign ?p y \ x < y)" using sn xy unfolding sign_def sign_num_def using quadratic_poly_eval by presburger then have "(\c \ (root_list ?p). x \ c \ c \ y)" using cs by auto then obtain q where "q \ root_list ?p \ x \ q \ q \ y" by auto then have "a*q^2 + b*q + c = 0 \ x \ q \ q \ y" unfolding root_list_def using quadratic_poly_eval[of c b a q] by auto then show "\q. a * q\<^sup>2 + b * q + c = 0 \ x \ q \ q \ y" by auto qed subsection "Continuity Properties" lemma continuity_lem_eq0: fixes p::"real" shows "r < p \ \x\{r <..p}. a * x\<^sup>2 + b * x + c = 0 \ (a = 0 \ b = 0 \ c = 0)" proof - assume r_lt: "r < p" assume inf_zer: "\x\{r <..p}. a * x\<^sup>2 + b * x + c = 0" have nf: "\finite {r..(a = 0 \ b = 0 \ c = 0) \ False" proof - assume "\(a = 0 \ b = 0 \ c = 0)" then have "[:c, b, a:] \ 0" by auto then have fin: "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto have "{x. a*x^2 + b*x + c = 0} = {x. poly [:c, b, a:] x = 0}" using quadratic_poly_eval by auto then have finset: "finite {x. a*x^2 + b*x + c = 0}" using fin by auto have "{r <..p} \ {x. a*x^2 + b*x + c = 0}" using inf_zer by blast then show "False" using finset nf using finite_subset by (metis (no_types, lifting) infinite_Ioc_iff r_lt) qed then show "(a = 0 \ b = 0 \ c = 0)" by auto qed lemma continuity_lem_lt0: fixes r:: "real" fixes a b c:: "real" shows "poly [:c, b, a:] r < 0 \ \y'> r. \x\{r<..y'}. poly [:c, b, a:] x < 0" proof - let ?f = "poly [:c,b,a:]" assume r_ltz: "poly [:c, b, a:] r < 0" then have "[:c, b, a:] \ 0" by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then have fin: "finite {x. x > r \ poly [:c, b, a:] x = 0}" using finite_Collect_conjI by blast let ?l = "sorted_list_of_set {x. x > r \ poly [:c, b, a:] x = 0}" show ?thesis proof (cases "length ?l = 0") case True then have no_zer: "\(\x>r. poly [:c, b, a:] x = 0)" using sorted_list_of_set_eq_Nil_iff fin by auto then have "\y. y > r \ y < (r + 1) \ poly [:c, b, a:] y < 0 " proof - fix y assume "y > r \ y < r + 1" then show "poly [:c, b, a:] y < 0" using r_ltz no_zer poly_IVT_pos[where a = "r", where p = "[:c, b, a:]", where b = "y"] by (meson linorder_neqE_linordered_idom) qed then show ?thesis - by (metis (no_types, hide_lams) \\ (\x>r. poly [:c, b, a:] x = 0)\ \poly [:c, b, a:] r < 0\ greaterThanAtMost_iff linorder_neqE_linordered_idom linordered_field_no_ub poly_IVT_pos) + by (metis greaterThanAtMost_iff less_add_one less_eq_real_def linorder_not_le no_zer poly_IVT_pos r_ltz) next case False then have len_nonz: "length (sorted_list_of_set {x. r < x \ poly [:c, b, a:] x = 0}) \ 0" by blast then have "\n \ {x. x > r \ poly [:c, b, a:] x = 0}. (nth_default 0 ?l 0) \ n" using fin set_sorted_list_of_set sorted_sorted_list_of_set using in_set_conv_nth leI not_less0 sorted_nth_mono by (smt not_less_iff_gr_or_eq nth_default_def) then have no_zer: "\(\x>r. (x < (nth_default 0 ?l 0) \ poly [:c, b, a:] x = 0))" using sorted_sorted_list_of_set by auto then have fa: "\y. y > r \ y < (nth_default 0 ?l 0) \ poly [:c, b, a:] y < 0 " proof - fix y assume "y > r \ y < (nth_default 0 ?l 0)" then show "poly [:c, b, a:] y < 0" using r_ltz no_zer poly_IVT_pos[where a = "r", where p = "[:c, b, a:]", where b = "y"] by (meson less_imp_le less_le_trans linorder_neqE_linordered_idom) qed have "nth_default 0 ?l 0 > r" using fin set_sorted_list_of_set using len_nonz length_0_conv length_greater_0_conv mem_Collect_eq nth_mem by (metis (no_types, lifting) nth_default_def) then have "\(y'::real). r < y' \ y' < (nth_default 0 ?l 0)" using dense by blast then obtain y' where y_prop:"r < y' \y' < (nth_default 0 ?l 0)" by auto then have "\x\{r<..y'}. poly [:c, b, a:] x < 0" using fa by auto then show ?thesis using y_prop by blast qed qed lemma continuity_lem_gt0: fixes r:: "real" fixes a b c:: "real" shows "poly [:c, b, a:] r > 0 \ \y'> r. \x\{r<..y'}. poly [:c, b, a:] x > 0" proof - assume r_gtz: "poly [:c, b, a:] r > 0 " let ?p = "[:-c, -b, -a:]" - have revpoly: "\x. -1*(poly [:c, b, a:] x) = poly [:-c, -b, -a:] x" - by (metis (no_types, hide_lams) Polynomial.poly_minus add.inverse_neutral minus_pCons mult_minus1) + have revpoly: "\x. -1*(poly [:c, b, a:] x) = poly [:-c, -b, -a:] x" + by (metis (no_types, opaque_lifting) Polynomial.poly_minus minus_pCons mult_minus1 neg_equal_0_iff_equal) then have "poly ?p r < 0" using r_gtz by (metis mult_minus1 neg_less_0_iff_less) then have "\y'> r. \x\{r<..y'}. poly ?p x < 0" using continuity_lem_lt0 by blast then obtain y' where y_prop: "y' > r \ (\x\{r<..y'}. poly ?p x < 0)" by auto then have "\x\{r<..y'}. poly [:c, b, a:] x > 0 " using revpoly using neg_less_0_iff_less by fastforce then show ?thesis using y_prop by blast qed lemma continuity_lem_lt0_expanded: fixes r:: "real" fixes a b c:: "real" shows "a*r^2 + b*r + c < 0 \ \y'> r. \x\{r<..y'}. a*x^2 + b*x + c < 0" using quadratic_poly_eval continuity_lem_lt0 by (simp add: add.commute) lemma continuity_lem_gt0_expanded: fixes r:: "real" fixes a b c:: "real" fixes k::"real" assumes kgt: "k > r" shows "a*r^2 + b*r + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c > 0" proof - assume "a*r^2 + b*r + c > 0" then have "\y'> r. \x\{r<..y'}. poly [:c, b, a:] x > 0" using continuity_lem_gt0 quadratic_poly_eval[of c b a r] by auto then obtain y' where y_prop: "y' > r \ (\x\{r<..y'}. poly [:c, b, a:] x > 0)" by auto then have "\q. q > r \ q < min k y'" using kgt dense by (metis min_less_iff_conj) then obtain q where q_prop: "q > r \q < min k y'" by auto then have "a*q^2 + b*q + c > 0" using y_prop quadratic_poly_eval[of c b a q] by (metis greaterThanAtMost_iff less_eq_real_def min_less_iff_conj) then show ?thesis using q_prop by auto qed subsection "Negative Infinity (Limit) Properties" lemma ysq_dom_y: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - have c1: "b \ 0 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b \ 0" then have p1: "\(y:: real). (y < -1 \ y*b \ 0)" by (simp add: mult_nonneg_nonpos2) have p2: "\(y:: real). (y < -1 \ y^2 > 0)" by auto then have h1: "\(y:: real). (y < -1 \ y^2 > b*y)" using p1 p2 by (metis less_eq_real_def less_le_trans mult.commute) then show ?thesis by auto qed have c2: "b < 0 \ b > -1 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b < 0 \ b > -1 " then have h1: "\(y:: real). (y < -1 \ y^2 > b*y)" by (simp add: power2_eq_square) then show ?thesis by auto qed have c3: "b \ -1 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b \ -1 " then have h1: "\(y:: real). (y < b \ y^2 > b*y)" by (metis le_minus_one_simps(3) less_irrefl less_le_trans mult.commute mult_less_cancel_left power2_eq_square) then show ?thesis by auto qed then show ?thesis using c1 c2 c3 by (metis less_trans linorder_not_less) qed lemma ysq_dom_y_plus_coeff: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y::real). (y < w \ y^2 > b*y + c)" proof - have "\(w::real). \(y:: real). (y < w \ y^2 > b*y)" using ysq_dom_y by auto then obtain w where w_prop: "\(y:: real). (y < w \ y^2 > b*y)" by auto have "c \ 0 \ \(y::real). (y < w \ y^2 > b*y + c)" using w_prop by auto then have c1: "c \ 0 \ \(w::real). \(y::real). (y < w \ y^2 > b*y + c)" by auto have "\(w::real). \(y:: real). (y < w \ y^2 > (b-c)*y)" using ysq_dom_y by auto then obtain k where k_prop: "\(y:: real). (y < k \ y^2 > (b-c)*y)" by auto let ?mn = "min k (-1)" have "(c> 0 \ (\ y < -1. -c*y > c))" proof - assume cgt: " c> 0" show "\(y::real) < -1. -c*y > c" proof clarsimp fix y::"real" assume "y < -1" then have "-y > 1" by auto then have "c < c*(-y)" using cgt by (metis \1 < - y\ mult.right_neutral mult_less_cancel_left_pos) then show " c < - (c * y) " by auto qed qed then have "(c> 0 \ (\ y < ?mn. (b-c)*y > b*y + c))" by (simp add: left_diff_distrib) then have c2: "c > 0 \ \(y::real). (y < ?mn \ y^2 > b*y + c)" using k_prop by force then have c2: "c > 0 \ \(w::real). \(y::real). (y < w \ y^2 > b*y + c)" by blast show ?thesis using c1 c2 by fastforce qed lemma ysq_dom_y_plus_coeff_2: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y::real). (y > w \ y^2 > b*y + c)" proof - have "\(w::real). \(y::real). (y < w \ y^2 > -b*y + c)" using ysq_dom_y_plus_coeff[where b = "-b", where c = "c"] by auto then obtain w where w_prop: "\(y::real). (y < w \ y^2 > -b*y + c)" by auto let ?mn = "min w (-1)" have "\(y::real). (y < ?mn \ y^2 > -b*y + c)" using w_prop by auto then have "\(y::real). (y > (-1*?mn) \ y^2 > b*y + c)" - by (metis (no_types, hide_lams) add.inverse_inverse minus_less_iff mult_minus1 mult_minus_left mult_minus_right power2_eq_square) + by (metis (no_types, opaque_lifting) minus_less_iff minus_mult_commute mult_1 power2_eq_iff) then show ?thesis by auto qed lemma neg_lc_dom_quad: fixes a:: "real" fixes b:: "real" fixes c:: "real" assumes alt: "a < 0" shows "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c < 0)" proof - have "\(w::real). \(y::real). (y > w \ y^2 > (-b/a)*y + (-c/a))" using ysq_dom_y_plus_coeff_2[where b = "-b/a", where c = "-c/a"] by auto then have keyh: "\(w::real). \(y::real). (y > w \ a*y^2 < a*((-b/a)*y + (-c/a)))" using alt by auto have simp1: "\y. a*((-b/a)*y + (-c/a)) = a*(-b/a)*y + a*(-c/a)" using diff_divide_eq_iff by fastforce have simp2: "\y. a*(-b/a)*y + a*(-c/a) = -b*y + a*(-c/a)" using assms by auto have simp3: "\y. -b*y + a*(-c/a) = -b*y - c" using assms by auto then have "\y. a*((-b/a)*y + (-c/a)) = -b*y - c" using simp1 simp2 simp3 by auto then have keyh2: "\(w::real). \(y::real). (y > w \ a*y^2 < -b*y-c)" using keyh by auto have "\y. a*y^2 < -b*y-c \ a*y^2 + b*y + c < 0" by auto then show ?thesis using keyh2 by auto qed lemma pos_lc_dom_quad: fixes a:: "real" fixes b:: "real" fixes c:: "real" assumes alt: "a > 0" shows "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" proof - have "-a < 0" using alt by simp then have "\(w::real). \(y::real). (y > w \ -a*y^2 - b*y - c < 0)" using neg_lc_dom_quad[where a = "-a", where b = "-b", where c = "-c"] by auto then obtain w where w_prop: "\(y::real). (y > w \ -a*y^2 - b*y - c < 0)" by auto then have "\(y::real). (y > w \ a*y^2 + b*y + c > 0)" by auto then show ?thesis by auto qed (* lemma interval_infinite: fixes r p::"real" assumes "r < p" shows "infinite {r<..(d, e, f)\set les. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\y'>q. (\(d, e, f)\set les. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof (induct les) case Nil then show ?case using gt_ex by auto next case (Cons z les) have "\a\set les. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.prems by auto then have " \y'>q. \a\set les. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set les. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f < 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f < 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#les). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f < 0)" using y1_prop y2_prop by force then show ?case by blast qed lemma have_inbetween_point_les: fixes r::"real" assumes "(\(d, e, f)\set les. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - have "(\(d, e, f)\set les. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\y'>r. (\(d, e, f)\set les. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" using les_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set les. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set les. \x\{r<..y}. d * x\<^sup>2 + e * x + f < 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set les. d*q^2 + e*q + f < 0)" using y_prop by auto then show ?thesis by auto qed lemma one_root_a_gt0: fixes a b c r:: "real" shows "\y'. b = 2 * a * r \ \ a < 0 \ a * r^2 - 2 * a *r*r + c = 0 \ - r < y' \ \x\{-r<..y'}. \ a * x\<^sup>2 + 2 * a * r*x + c < 0" proof - fix y' assume beq: "b = 2 * a * r" assume aprop: " \ a < 0" assume root: " a * r\<^sup>2 - 2 * a *r*r + c = 0" assume rootlt: "- r < y'" show " \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a* r*x+ c < 0" proof - have h: "a = 0 \ (b = 0 \ c = 0)" using beq root by auto then have aeq: "a = 0 \ \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a*r*x + c < 0" using rootlt by (metis add.left_neutral continuity_lem_eq0 less_numeral_extra(3) mult_zero_left mult_zero_right) then have alt: "a > 0 \ \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a *r*x + c < 0" proof - assume agt: "a > 0" then have "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" using pos_lc_dom_quad by auto then obtain w where w_prop: "\y::real. (y > w \ a*y^2 + b*y + c > 0)" by auto have isroot: "a*(-r)^2 + b*(-r) + c = 0" using root beq by auto then have wgteq: "w \ -(r)" proof - have "w < -r \ False" using w_prop isroot by auto then show ?thesis using not_less by blast qed then have w1: "w + 1 > -r" by auto have w2: "a*(w + 1)^2 + b*(w+1) + c > 0" using w_prop by auto have rootiff: "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" using discriminant_lemma[where a = "a", where b = "b", where c= "c", where r = "r"] isroot agt beq by auto have allgt: "\x > -r. a*x^2 + b*x + c > 0" proof clarsimp fix x assume "x > -r" have xgtw: "x > w + 1 \ a*x^2 + b*x + c > 0 " using w1 w2 rootiff poly_IVT_neg[where a = "w+1", where b = "x", where p = "[:c,b,a:]"] quadratic_poly_eval by (metis less_eq_real_def linorder_not_less) have xltw: "x < w + 1 \ a*x^2 + b*x + c > 0" using w1 w2 rootiff poly_IVT_pos[where a= "x", where b = "w + 1", where p = "[:c,b,a:]"] quadratic_poly_eval less_eq_real_def linorder_not_less by (metis \- r < x\) then show "a*x^2 + b*x + c > 0" using w2 xgtw xltw by fastforce qed have "\z. z > -r \ z < y'" using rootlt dense[where x = "-r", where y = "y'"] by auto then obtain z where z_prop: " z > -r \ z < y'" by auto then have "a*z^2 + b*z + c > 0" using allgt by auto then show ?thesis using z_prop using beq greaterThanAtMost_iff by force qed then show ?thesis using aeq alt aprop by linarith qed qed lemma leq_qe_inf_helper: fixes q:: "real" shows"(\(d, e, f)\set leq. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>q. (\(d, e, f)\set leq. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof (induct leq) case Nil then show ?case using gt_ex by auto next case (Cons z leq) have "\a\set leq. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then have " \y'>q. \a\set leq. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set leq. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#leq). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f \ 0)" using y1_prop y2_prop by force then show ?case by blast qed lemma neq_qe_inf_helper: fixes q:: "real" shows"(\(d, e, f)\set neq. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>q. (\(d, e, f)\set neq. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof (induct neq) case Nil then show ?case using gt_ex by auto next case (Cons z neq) have "\a\set neq. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then have " \y'>q. \a\set neq. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set neq. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#neq). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f \ 0)" using y1_prop y2_prop by force then show ?case by blast qed subsection "Some Casework" lemma quadratic_shape1a: fixes a b c x y::"real" assumes agt: "a > 0" assumes xyroots: "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > x \ z < y \ a*z^2 + b*z + c < 0)" proof clarsimp fix z assume zgt: "z > x" assume zlt: "z < y" have frac_gtz: "(1/(2*a)) > 0" using agt by simp have xy_prop:"(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using xyroots agt discriminant_iff unfolding discrim_def by auto have "b^2 - 4*a*c \ 0" using xyroots discriminant_iff using assms(1) discrim_def by auto then have pos_discrim: "b^2 - 4*a*c > 0" using xyroots discriminant_zero using \0 \ b\<^sup>2 - 4 * a * c\ assms(1) discrim_def less_eq_real_def linorder_not_less by metis then have sqrt_gt: "sqrt(b^2 - 4*a*c) > 0" using real_sqrt_gt_0_iff by blast then have "(- b - sqrt(b^2 - 4*a*c)) < (- b + sqrt(b^2 - 4*a*c))" by auto then have "(- b - sqrt(b^2 - 4*a*c))*(1/(2*a)) < (- b + sqrt(b^2 - 4*a*c))*(1/(2*a)) " using frac_gtz by (simp add: divide_strict_right_mono) then have "(- b - sqrt(b^2 - 4*a*c))/(2*a) < (- b + sqrt(b^2 - 4*a*c))/(2*a)" by auto then have xandy: "x = (- b - sqrt(b^2 - 4*a*c))/(2*a) \ y = (- b + sqrt(b^2 - 4*a*c))/(2*a)" using xy_prop xyroots by auto let ?mdpt = "-b/(2*a)" have xlt: "x < ?mdpt" using xandy sqrt_gt using frac_gtz divide_minus_left divide_strict_right_mono sqrt_gt by (smt (verit) agt) have ylt: "?mdpt < y" using xandy sqrt_gt frac_gtz by (smt (verit, del_insts) divide_strict_right_mono zero_less_divide_1_iff) have mdpt_val: "a*?mdpt^2 + b*?mdpt + c < 0" proof - have firsteq: "a*?mdpt^2 + b*?mdpt + c = (a*b^2)/(4*a^2) - (b^2)/(2*a) + c" by (simp add: power2_eq_square) have h1: "(a*b^2)/(4*a^2) = (b^2)/(4*a)" by (simp add: power2_eq_square) have h2: "(b^2)/(2*a) = (2*b^2)/(4*a)" by linarith have h3: "c = (4*a*c)/(4*a)" using agt by auto have "a*?mdpt^2 + b*?mdpt + c = (b^2)/(4*a) - (2*b^2)/(4*a) + (4*a*c)/(4*a) " using firsteq h1 h2 h3 by linarith then have "a*?mdpt^2 + b*?mdpt + c = (b^2 - 2*b^2 + 4*a*c)/(4*a)" by (simp add: diff_divide_distrib) then have eq2: "a*?mdpt^2 + b*?mdpt + c = (4*a*c - b^2)/(4*a)" by simp have h: "4*a*c - b^2 < 0" using pos_discrim by auto have "1/(4*a) > 0" using agt by auto then have "(4*a*c - b^2)*(1/(4*a)) < 0" using h using mult_less_0_iff by blast then show ?thesis using eq2 by auto qed have nex: "\ (\k> x. k < y \ poly [:c, b, a:] k = 0)" using discriminant_iff agt - by (metis (no_types, hide_lams) discrim_def order_less_irrefl quadratic_poly_eval xandy) + by (metis discrim_def less_irrefl quadratic_poly_eval xandy) have nor2: "\ (\x>z. x < - b / (2 * a) \ poly [:c, b, a:] x = 0)" using nex xlt ylt zgt zlt by auto have nor: "\ (\x>- b / (2 * a). x < z \ poly [:c, b, a:] x = 0)" using nex xlt ylt zgt zlt discriminant_iff agt by auto then have mdpt_lt: "?mdpt < z \ a*z^2 + b*z + c < 0 " using mdpt_val zgt zlt xlt ylt poly_IVT_pos[where p = "[:c, b, a:]", where a= "?mdpt", where b = "z"] quadratic_poly_eval[of c b a] by (metis \\ (\k>x. k < y \ poly [:c, b, a:] k = 0)\ linorder_neqE_linordered_idom) have mdpt_gt: "?mdpt > z \ a*z^2 + b*z + c < 0 " using zgt zlt mdpt_val xlt ylt nor2 poly_IVT_neg[where p = "[:c, b, a:]", where a = "z", where b = "?mdpt"] quadratic_poly_eval[of c b a] by (metis linorder_neqE_linordered_idom nex) then show "a*z^2 + b*z + c < 0" using mdpt_lt mdpt_gt mdpt_val by fastforce qed lemma quadratic_shape1b: fixes a b c x y::"real" assumes agt: "a > 0" assumes xy_roots: "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > y \ a*z^2 + b*z + c > 0)" proof - fix z assume z_gt :"z > y" have nogt: "\(\w. w > y \ a*w^2 + b*w + c = 0)" using xy_roots discriminant_iff by (metis agt less_eq_real_def linorder_not_less) have "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" using agt pos_lc_dom_quad by auto then have "\k > y. a*k^2 + b*k + c > 0" by (metis add.commute agt less_add_same_cancel1 linorder_neqE_linordered_idom pos_add_strict) then obtain k where k_prop: "k > y \ a*k^2 + b*k + c > 0" by auto have kgt: "k > z \ a*z^2 + b*z + c > 0" proof - assume kgt: "k > z" then have zneq: "a*z^2 + b*z + c = 0 \ False" using nogt using z_gt by blast have znlt: "a*z^2 + b*z + c < 0 \ False" using kgt k_prop quadratic_poly_eval[of c b a] z_gt nogt poly_IVT_pos[where a= "z", where b = "k", where p = "[:c, b, a:]"] by (metis less_eq_real_def less_le_trans) then show "a*z^2 + b*z + c > 0" using zneq znlt using linorder_neqE_linordered_idom by blast qed have klt: "k < z \ a*z^2 + b*z + c > 0" proof - assume klt: "k < z" then have zneq: "a*z^2 + b*z + c = 0 \ False" using nogt using z_gt by blast have znlt: "a*z^2 + b*z + c < 0 \ False" using klt k_prop quadratic_poly_eval[of c b a] z_gt nogt poly_IVT_neg[where a= "k", where b = "z", where p = "[:c, b, a:]"] by (metis add.commute add_less_cancel_left add_mono_thms_linordered_field(3) less_eq_real_def) then show "a*z^2 + b*z + c > 0" using zneq znlt using linorder_neqE_linordered_idom by blast qed then show "a*z^2 + b*z + c > 0" using k_prop kgt klt by fastforce qed lemma quadratic_shape2a: fixes a b c x y::"real" assumes "a < 0" assumes "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > x \ z < y \ a*z^2 + b*z + c > 0)" using quadratic_shape1a[where a= "-a", where b = "-b", where c = "-c", where x = "x", where y = "y"] using assms(1) assms(2) by fastforce lemma quadratic_shape2b: fixes a b c x y::"real" assumes "a < 0" assumes "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > y \ a*z^2 + b*z + c < 0)" using quadratic_shape1b[where a= "-a", where b = "-b", where c = "-c", where x = "x", where y = "y"] using assms(1) assms(2) by force lemma case_d1: fixes a b c r::"real" shows "b < 2 * a * r \ a * r^2 - b*r + c = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume b_lt: "b < 2*a*r" assume root: "a*r^2 - b*r + c = 0" then have "c = b*r - a*r^2" by auto have aeq: "a = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume azer: "a = 0" then have bltz: "b < 0" using b_lt by auto then have "c = b*r" using azer root by auto then have eval: "\x. a*x^2 + b*x + c = b*(x + r)" using azer by (simp add: distrib_left) have "\x > -r. b*(x + r) < 0" proof clarsimp fix x assume xgt: "- r < x" then have "x + r > 0" by linarith then show "b * (x + r) < 0" using bltz using mult_less_0_iff by blast qed then show ?thesis using eval using less_add_same_cancel1 zero_less_one by (metis greaterThanAtMost_iff) qed have aneq: "a \ 0 \\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume aneq: "(a::real) \ 0" have "b^2 - 4*a*c < 0 \ a * r\<^sup>2 + b * r + c \ 0" using root discriminant_negative[of a b c r] unfolding discrim_def using aneq by auto then have " a * r\<^sup>2 + b * r + c \ 0 \ a * r\<^sup>2 - b * r + c = 0 \ b\<^sup>2 < 4 * a * c \ False" proof - assume a1: "a * r\<^sup>2 - b * r + c = 0" assume a2: "b\<^sup>2 < 4 * a * c" have f3: "(0 \ - 1 * (4 * a * c) + (- 1 * b)\<^sup>2) = (4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0)" by simp have f4: "(- 1 * b)\<^sup>2 + - 1 * (4 * a * c) = - 1 * (4 * a * c) + (- 1 * b)\<^sup>2" by auto have f5: "c + a * r\<^sup>2 + - 1 * b * r = a * r\<^sup>2 + c + - 1 * b * r" by auto have f6: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" by simp have f7: "\x1 x2 x3. (discrim x3 x2 x1 < 0) = (\ 0 \ discrim x3 x2 x1)" by auto have f8: "\r ra rb. discrim r ra rb = ra\<^sup>2 + - 1 * (4 * r * rb)" using discrim_def by auto have "\ 4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0" using a2 by simp then have "a * r\<^sup>2 + c + - 1 * b * r \ 0" using f8 f7 f6 f5 f4 f3 by (metis (no_types) aneq discriminant_negative) then show False using a1 by linarith qed then have "\(b^2 - 4*a*c < 0)" using root using \b\<^sup>2 - 4 * a * c < 0 \ a * r\<^sup>2 + b * r + c \ 0\ by auto then have discrim: "b\<^sup>2 \ 4 * a * c " by auto then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using aneq root discriminant_iff[where a="a", where b ="-b", where c="c", where x="r"] unfolding discrim_def by auto then have "r = (b - sqrt(b^2 - 4*a*c))/(2*a) \ b > 2*a*r" proof - assume req: "r = (b - sqrt(b^2 - 4*a*c))/(2*a)" then have h1: "2*a*r = 2*a*((b - sqrt(b^2 - 4*a*c))/(2*a))" by auto then have h2: "2*a*((b - sqrt(b^2 - 4*a*c))/(2*a)) = b - sqrt(b^2 - 4*a*c)" using aneq by auto have h3: "sqrt(b^2 - 4*a*c) \ 0" using discrim by auto then have "b - sqrt(b^2 - 4*a*c) < b" using b_lt h1 h2 by linarith then show ?thesis using req h2 by auto qed then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a)" using req b_lt by auto then have discrim2: "b^2 - 4 *a*c > 0" using aneq b_lt by auto then have "\x y. x \ y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" using aneq discriminant_pos_ex[of a b c] unfolding discrim_def by auto then obtain x y where xy_prop: "x < y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" by (meson linorder_neqE_linordered_idom) then have "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using aneq discriminant_iff unfolding discrim_def by auto then have xy_prop2: "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = -r) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = -r)" using req by (simp add: \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ minus_divide_left) (* When a < 0, -r is the bigger root *) have alt: "a < 0 \ \k > -r. a * k^2 + b * k + c < 0" proof clarsimp fix k assume alt: " a < 0" assume "- r < k" have alt2: " (1/(2*a)::real) < 0" using alt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) > (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 using mult_less_cancel_left_neg by fastforce then have rgtroot: "-r > (-b + sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(y = -r \ x = (-b + sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto then show "a * k^2 + b * k + c < 0" using xy_prop \- r < k\ alt quadratic_shape2b xy_prop by blast qed (* When a > 0, -r is the smaller root *) have agt: "a > 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume agt: "a> 0" have alt2: " (1/(2*a)::real) > 0" using agt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) < (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 proof - have f1: "- b - sqrt (b\<^sup>2 - c * (4 * a)) < - b + sqrt (b\<^sup>2 - c * (4 * a))" by (metis \- b - sqrt (b\<^sup>2 - 4 * a * c) < - b + sqrt (b\<^sup>2 - 4 * a * c)\ mult.commute) have "0 < a * 2" using \0 < 1 / (2 * a)\ by auto then show ?thesis using f1 by (simp add: divide_strict_right_mono mult.commute) qed then have rlltroot: "-r < (-b + sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(x = -r \ y = (-b + sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto have "\k. x < k \ k < y" using xy_prop dense by auto then obtain k where k_prop: "x < k \ k < y" by auto then have "\x\{-r<..k}. a * x\<^sup>2 + b * x + c < 0" using agt quadratic_shape1a[where a= "a", where b = "b", where c= "c", where x = "x", where y = "y"] using \x = - r \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ greaterThanAtMost_iff xy_prop by auto then show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" using k_prop using \x = - r \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ by blast qed show ?thesis using alt agt by (metis aneq greaterThanAtMost_iff less_add_same_cancel1 linorder_neqE_linordered_idom zero_less_one) qed show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" using aeq aneq by blast qed lemma case_d4: fixes a b c r::"real" shows "\y'. b \ 2 * a * r \ \ b < 2 * a * r \ a *r^2 - b * r + c = 0 \ -r < y' \ \x\{-r<..y'}. \ a * x\<^sup>2 + b * x + c < 0" proof - fix y' assume bneq: "b \ 2 * a * r" assume bnotless: "\ b < 2 * a * r" assume root: "a *r^2 - b * r + c = 0" assume y_prop: "-r < y'" have b_gt: "b > 2*a*r" using bneq bnotless by auto have aeq: "a = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume azer: "a = 0" then have bgt: "b > 0" using b_gt by auto then have "c = b*r" using azer root by auto then have eval: "\x. a*x^2 + b*x + c = b*(x + r)" using azer by (simp add: distrib_left) have "\x > -r. b*(x + r) > 0" proof clarsimp fix x assume xgt: "- r < x" then have "x + r > 0" by linarith then show "b * (x + r) > 0" using bgt by auto qed then show ?thesis using eval using less_add_same_cancel1 zero_less_one by (metis greaterThanAtMost_iff) qed have aneq: "a \ 0 \\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume aneq: "a\0" { assume a1: "a * r\<^sup>2 - b * r + c = 0" assume a2: "b\<^sup>2 < 4 * a * c" have f3: "(0 \ - 1 * (4 * a * c) + (- 1 * b)\<^sup>2) = (4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0)" by simp have f4: "(- 1 * b)\<^sup>2 + - 1 * (4 * a * c) = - 1 * (4 * a * c) + (- 1 * b)\<^sup>2" by auto have f5: "c + a * r\<^sup>2 + - 1 * b * r = a * r\<^sup>2 + c + - 1 * b * r" by auto have f6: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" by simp have f7: "\x1 x2 x3. (discrim x3 x2 x1 < 0) = (\ 0 \ discrim x3 x2 x1)" by auto have f8: "\r ra rb. discrim r ra rb = ra\<^sup>2 + - 1 * (4 * r * rb)" using discrim_def by auto have "\ 4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0" using a2 by simp then have "a * r\<^sup>2 + c + - 1 * b * r \ 0" using f8 f7 f6 f5 f4 f3 by (metis (no_types) aneq discriminant_negative) then have False using a1 by linarith } note * = this have "b^2 - 4*a*c < 0 \ a * r\<^sup>2 + b * r + c \ 0" using root discriminant_negative[of a b c r] unfolding discrim_def using aneq by auto then have "\(b^2 - 4*a*c < 0)" using root * by auto then have discrim: "b\<^sup>2 \ 4 * a * c " by auto then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using aneq root discriminant_iff[where a="a", where b ="-b", where c="c", where x="r"] unfolding discrim_def by auto then have "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ b < 2*a*r" proof - assume req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a)" then have h1: "2*a*r = 2*a*((b + sqrt(b^2 - 4*a*c))/(2*a))" by auto then have h2: "2*a*((b + sqrt(b^2 - 4*a*c))/(2*a)) = b + sqrt(b^2 - 4*a*c)" using aneq by auto have h3: "sqrt(b^2 - 4*a*c) \ 0" using discrim by auto then have "b + sqrt(b^2 - 4*a*c) > b" using b_gt h1 h2 by linarith then show ?thesis using req h2 by auto qed then have req: "r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using req b_gt using aneq discrim by auto then have discrim2: "b^2 - 4 *a*c > 0" using aneq b_gt by auto then have "\x y. x \ y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" using aneq discriminant_pos_ex[of a b c] unfolding discrim_def by auto then obtain x y where xy_prop: "x < y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" by (meson linorder_neqE_linordered_idom) then have "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using aneq discriminant_iff unfolding discrim_def by auto then have xy_prop2: "(x = (-b - sqrt(b^2 - 4*a*c))/(2*a) \ y = -r) \ (y = (-b - sqrt(b^2 - 4*a*c))/(2*a) \ x = -r)" using req divide_inverse minus_diff_eq mult.commute mult_minus_right by (smt (verit, ccfv_SIG) uminus_add_conv_diff) (* When a > 0, -r is the greater root *) have agt: "a > 0 \ \k > -r. a * k^2 + b * k + c > 0" proof clarsimp fix k assume agt: " a > 0" assume "- r < k" have agt2: " (1/(2*a)::real) > 0" using agt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) < (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using agt2 by (simp add: divide_strict_right_mono) then have rgtroot: "-r > (-b - sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(x = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ y = -r" using xy_prop xy_prop2 by auto then show "a * k^2 + b * k + c > 0" using \- r < k\ xy_prop agt quadratic_shape1b[where a= "a", where b ="b", where c="c", where x = "x", where y = "-r", where z = "k"] by blast qed (* When a < 0, -r is the smaller root *) have agt2: "a < 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume alt: "a<0" have alt2: " (1/(2*a)::real) < 0" using alt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) > (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 using mult_less_cancel_left_neg by fastforce then have rlltroot: "-r < (-b - sqrt(b^2 - 4*a*c))/(2*a)" using req using \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have h: "(x = -r \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto have "\k. x < k \ k < y" using xy_prop dense by auto then obtain k where k_prop: "x < k \ k < y" by auto then have "\x\{-r<..k}. a * x\<^sup>2 + b * x + c > 0" using alt quadratic_shape2a[where a= "a", where b = "b", where c= "c", where x = "x", where y = "y"] xy_prop h greaterThanAtMost_iff by auto then show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" using k_prop using h by blast qed show ?thesis using aneq agt agt2 by (meson greaterThanAtMost_iff linorder_neqE_linordered_idom y_prop) qed show "\x\{-r<..y'}. \ a * x\<^sup>2 + b * x + c < 0" using aneq aeq by (metis greaterThanAtMost_iff less_eq_real_def linorder_not_less y_prop) qed lemma one_root_a_lt0: fixes a b c r y'::"real" assumes alt: "a < 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2*a*r*r + c = 0" shows "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2*a*r*x + c < 0" proof - have root_iff: "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" using alt root discriminant_lemma[where r = "r"] beq by auto have "a < 0 \ (\y. \x > y. a*x^2 + b*x + c < 0)" using neg_lc_dom_quad by auto then obtain k where k_prop: "\x > k. a*x^2 + b*x + c < 0" using alt by auto let ?mx = "max (k+1) (-r + 1)" have "a*?mx^2 + b*?mx + c < 0" using k_prop by auto then have "\y > -r. a*y^2 + b*y + c < 0" by force then obtain z where z_prop: "z > -r \ a*z^2 + b*z + c < 0" by auto have poly_eval_prop: "\(x::real). poly [:c, b, a :] x = a*x^2 + b*x + c" using quadratic_poly_eval by auto then have nozer: "\(\x. (x > -r \ poly [:c, b, a :] x = 0))" using root_iff by (simp add: add.commute) have poly_z: "poly [:c, b, a:] z < 0" using z_prop poly_eval_prop by auto have "\y > -r. a*y^2 + b*y + c < 0" proof clarsimp fix y assume ygt: "- r < y" have h1: "y = z \ a * y\<^sup>2 + b * y + c < 0" using z_prop by auto have h2: "y < z \ a * y\<^sup>2 + b * y + c < 0" proof - assume ylt: "y < z" have notz: "a*y^2 + b*y + c \ 0" using ygt nozer poly_eval_prop by auto have h1: "a *y^2 + b*y + c > 0 \ poly [:c, b, a:] y > 0" using poly_eval_prop by auto have ivtprop: "poly [:c, b, a:] y > 0 \ (\x. y < x \ x < z \ poly [:c, b, a:] x = 0)" using ylt poly_z poly_IVT_neg[where a = "y", where b = "z", where p = "[:c, b, a:]"] by auto then have "a*y^2 + b*y + c > 0 \ False" using h1 ivtprop ygt nozer by auto then show "a*y^2 + b*y + c < 0" using notz using linorder_neqE_linordered_idom by blast qed have h3: "y > z \ a * y\<^sup>2 + b * y + c < 0" proof - assume ygtz: "y > z" have notz: "a*y^2 + b*y + c \ 0" using ygt nozer poly_eval_prop by auto have h1: "a *y^2 + b*y + c > 0 \ poly [:c, b, a:] y > 0" using poly_eval_prop by auto have ivtprop: "poly [:c, b, a:] y > 0 \ (\x. z < x \ x < y \ poly [:c, b, a:] x = 0)" using ygtz poly_z using poly_IVT_pos by blast then have "a*y^2 + b*y + c > 0 \ False" using h1 ivtprop z_prop nozer by auto then show "a*y^2 + b*y + c < 0" using notz using linorder_neqE_linordered_idom by blast qed show "a * y\<^sup>2 + b * y + c < 0" using h1 h2 h3 using linorder_neqE_linordered_idom by blast qed then show ?thesis using \\y>- r. a * y\<^sup>2 + b * y + c < 0\ beq by auto qed lemma one_root_a_lt0_var: fixes a b c r y'::"real" assumes alt: "a < 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2*a*r*r + c = 0" shows "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2*a*r*x + c \ 0" proof - have h1: "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2 * a * r * x + c < 0 \ \y'>-r. \x\{- r<..y'}. a * x\<^sup>2 + 2 * a *r * x + c \ 0" using less_eq_real_def by blast then show ?thesis using one_root_a_lt0[of a b r] assms by auto qed subsection "More Continuity Properties" lemma continuity_lem_gt0_expanded_var: fixes r:: "real" fixes a b c:: "real" fixes k::"real" assumes kgt: "k > r" shows "a*r^2 + b*r + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c \ 0" proof - assume a: "a*r^2 + b*r + c > 0 " have h: "\x\{r<..k}. a*x^2 + b*x + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c \ 0" using less_eq_real_def by blast have "\x\{r<..k}. a*x^2 + b*x + c > 0" using a continuity_lem_gt0_expanded[of r k a b c] assms by auto then show "\x\{r<..k}. a*x^2 + b*x + c \ 0" using h by auto qed lemma continuity_lem_lt0_expanded_var: fixes r:: "real" fixes a b c:: "real" shows "a*r^2 + b*r + c < 0 \ \y'> r. \x\{r<..y'}. a*x^2 + b*x + c \ 0" proof - assume "a*r^2 + b*r + c < 0 " then have " \y'> r. \x\{r<..y'}. a*x^2 + b*x + c < 0" using continuity_lem_lt0_expanded by auto then show " \y'> r. \x\{r<..y'}. a*x^2 + b*x + c \ 0" using less_eq_real_def by auto qed lemma nonzcoeffs: fixes a b c r::"real" shows "a\0 \ b\0 \ c\0 \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0 " proof - assume "a\0 \ b\0 \ c\0" then have fin: "finite {x. a*x^2 + b*x + c = 0}" by (metis pCons_eq_0_iff poly_roots_finite poly_roots_set_same) (* then have fin2: "finite {x. a*x^2 + b*x + c = 0 \ x > r}" using finite_Collect_conjI by blast *) let ?s = "{x. a*x^2 + b*x + c = 0}" have imp: "(\q \ ?s. q > r) \ (\q \ ?s. (q > r \ (\x \ ?s. x > r \ x \ q)))" proof - assume asm: "(\q \ ?s. q > r)" then have none: "{q. q \ ?s \ q > r} \ {}" by blast have fin2: "finite {q. q \ ?s \ q > r}" using fin by simp have "\k. k = Min {q. q \ ?s \ q > r}" using fin2 none by blast then obtain k where k_prop: "k = Min {q. q \ ?s \ q > r}" by auto then have kp1: "k \ ?s \ k > r" using Min_in fin2 none by blast then have kp2: "\x \ ?s. x > r \ x \ k" using Min_le fin2 using k_prop by blast show "(\q \ ?s. (q > r \ (\x \ ?s. x > r \ x \ q)))" using kp1 kp2 by blast qed have h2: "(\q \ ?s. q > r) \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" proof - assume "(\q \ ?s. q > r)" then obtain q where q_prop: "q \ ?s \ (q > r \ (\x \ ?s. x > r \ x \ q))" using imp by blast then have "\w. w > r \ w < q" using dense by blast then obtain w where w_prop: "w > r \ w < q" by auto then have "\(\x\{r<..w}. x \ ?s)" using w_prop q_prop by auto then have "\x\{r<..w}. a * x\<^sup>2 + b * x + c \ 0" by blast then show "\y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" using w_prop by blast qed have h1: "\(\q \ ?s. q > r) \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" proof - assume "\(\q \ ?s. q > r)" then have "\x\{r<..(r+1)}. a * x\<^sup>2 + b * x + c \ 0" using greaterThanAtMost_iff by blast then show ?thesis using less_add_same_cancel1 less_numeral_extra(1) by blast qed then show "\y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" using h2 by blast qed (* Show if there are infinitely many values of x where a*x^2 + b*x + c is 0, then the a*x^2 + b*x + c is the zero polynomial *) lemma infzeros : fixes y:: "real" assumes "\x::real < (y::real). a * x\<^sup>2 + b * x + c = 0" shows "a = 0 \ b=0 \ c=0" proof - let ?A = "{(x::real). x < y}" have "\ (n::nat) f. ?A = f ` {i. i < n} \ inj_on f {i. i < n} \ False" proof clarsimp fix n:: "nat" fix f assume xlt: "{x. x < y} = f ` {i. i < n}" assume injh: "inj_on f {i. i < n}" have "?A \ {}" by (simp add: linordered_field_no_lb) then have ngtz: "n > 0" using xlt injh using gr_implies_not_zero by auto have cardisn: "card ?A = n" using xlt injh by (simp add: card_image) have "\k::nat. ((y - (k::nat) - 1) \ ?A)" by auto then have subs: "{k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n} \ ?A" by auto have seteq: "(\x. y - real x - 1) ` {0..n} ={k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n}" by auto have injf: "inj_on (\x. y - real x - 1) {0..n}" unfolding inj_on_def by auto have "card {k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n} = n + 1" using injf seteq card_atMost inj_on_iff_eq_card[where A = "{0..n}", where f = "\x. y - x - 1"] by auto then have if_fin: "finite ?A \ card ?A \ n + 1" using subs card_mono by (metis (lifting) card_mono) then have if_inf: "infinite ?A \ card ?A = 0" by (meson card.infinite) then show "False" using if_fin if_inf cardisn ngtz by auto qed then have nfin: "\ finite {(x::real). x < y}" using finite_imp_nat_seg_image_inj_on by blast have "{(x::real). x < y} \ {x. a*x^2 + b*x + c = 0}" using assms by auto then have nfin2: "\ finite {x. a*x^2 + b*x + c = 0}" using nfin finite_subset by blast { fix x assume "a * x\<^sup>2 + b * x + c = 0" then have f1: "a * (x * x) + x * b + c = 0" by (simp add: Groups.mult_ac(2) power2_eq_square) have f2: "\r. c + (r + (c + - c)) = r + c" by simp have f3: "\r ra rb. (rb::real) * ra + ra * r = (rb + r) * ra" by (metis Groups.mult_ac(2) Rings.ring_distribs(2)) have "\r. r + (c + - c) = r" by simp then have "c + x * (b + x * a) = 0" using f3 f2 f1 by (metis Groups.add_ac(3) Groups.mult_ac(1) Groups.mult_ac(2)) } hence "{x. a*x^2 + b*x + c = 0} \ {x. poly [:c, b, a:] x = 0}" by auto then have " \ finite {x. poly [:c, b, a:] x = 0}" using nfin2 using finite_subset by blast then have "[:c, b, a:] = 0" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show ?thesis by auto qed lemma have_inbetween_point_leq: fixes r::"real" assumes "(\((d::real), (e::real), (f::real))\set leq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof - have "(\(d, e, f)\set leq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>r. (\(d, e, f)\set leq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using leq_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set leq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set leq. \x\{r<..y}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set leq. d*q^2 + e*q + f \ 0)" using y_prop by auto then show ?thesis by auto qed lemma have_inbetween_point_neq: fixes r::"real" assumes "(\((d::real), (e::real), (f::real))\set neq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof - have "(\(d, e, f)\set neq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>r. (\(d, e, f)\set neq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using neq_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set neq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set neq. \x\{r<..y}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set neq. d*q^2 + e*q + f \ 0)" using y_prop by auto then show ?thesis by auto qed subsection "Setting up and Helper Lemmas" subsubsection "The les\\_qe lemma" lemma les_qe_forward : shows "(((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))) \ ((\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)))" proof - assume big_asm: "(((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))))" then have big_or: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) " by auto have h1_helper: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - show "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof (induct les) case Nil then show ?case by auto next case (Cons q les) have ind: " \a\set (q # les). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c < 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c < 0)" by auto have "\a\set les. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using ind by simp then have " \y. \xa\set les. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set les. case a of (a, ba, c) \ a * x^2 + ba * x + c < 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #les). case a of (a, ba, c) \ a * x^2 + ba * x + c < 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed then have h1: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" by (smt (z3) infzeros less_eq_real_def not_numeral_le_zero) (* apply (auto) by (metis (lifting) infzeros zero_neq_numeral) *) have h2: " (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "(\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" then obtain a' b' c' where abc_prop: "(a', b', c') \set les \ a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then show "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto qed have h3: " (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ ((\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)))" proof - assume asm: "\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set les \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then show "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto qed have h4: "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) " then obtain a' b' c' where abc_prop: "(a', b', c')\set les \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto then show ?thesis using asm by auto qed show ?thesis using big_or h1 h2 h3 h4 by blast qed (*sample points, some starter proofs below in comments *) lemma les_qe_backward : shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" proof - assume havex: "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then obtain x where x_prop: "\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0" by auto have h: "(\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))) \ False" proof - assume big: "(\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))" have notneginf: "\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" using big by auto have notlinroot: "\ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have notquadroot1: " \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have notquadroot2:" \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have nok: "\ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - have "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ False" proof - assume "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" then obtain k a b c where k_prop: "(a, b, c) \ set les \ a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto have azer: "a = 0 \ False" proof - assume azer: "a = 0" then have "b = 0 \ c = 0" using k_prop by auto then have bnonz: "b\ 0" using azer x_prop k_prop by auto then have "k = -c/b" using k_prop azer - by (metis (no_types, hide_lams) add.commute add.left_neutral add_uminus_conv_diff diff_le_0_iff_le divide_non_zero less_eq_real_def mult_zero_left neg_less_iff_less order_less_irrefl real_add_less_0_iff) + by (smt (verit, best) mult_eq_0_iff nonzero_mult_div_cancel_left) then have " (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using k_prop azer bnonz by auto then show "False" using notlinroot by auto qed have anonz: "a \ 0 \ False" proof - assume anonz: "a \ 0 " let ?r1 = "(- b - sqrt (b^2 - 4 * a * c)) / (2 * a)" let ?r2 = "(- b + sqrt (b^2 - 4 * a * c)) / (2 * a)" have discr: "4 * a * c \ b^2" using anonz k_prop discriminant_negative[of a b c] unfolding discrim_def by fastforce then have "k = ?r1 \ k = ?r2" using k_prop discriminant_nonneg[of a b c] unfolding discrim_def using anonz by auto then have "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using discr anonz notquadroot1 notquadroot2 k_prop by auto then show "False" using notquadroot1 notquadroot2 by auto qed show "False" using azer anonz by auto qed then show ?thesis by auto qed have finset: "finite (set les)" by blast have h1: "(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0) \ False" using x_prop by fastforce then have h2: "\(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0) \ False" proof - assume nozer: "\(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0)" then have same_set: "root_set (set les) = set (sorted_root_list_set (set les))" using root_set_finite finset set_sorted_list_of_set by (simp add: nozer root_set_finite sorted_root_list_set_def) have xnotin: "x \ root_set (set les)" unfolding root_set_def using x_prop by auto let ?srl = "sorted_root_list_set (set les)" have notinlist: "\ List.member ?srl x" using xnotin same_set by (simp add: in_set_member) then have notmem: "\n < (length ?srl). x \ nth_default 0 ?srl n" using nth_mem same_set xnotin nth_default_def by metis show ?thesis proof (induct ?srl) case Nil then have "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" proof clarsimp fix a b c assume noroots: "[] = sorted_root_list_set (set les)" assume inset: "(a, b, c) \ set les" have "{} = root_set (set les)" using noroots same_set by auto then have nozero: "\(\x. a*x^2 + b*x + c = 0)" using inset unfolding root_set_def by auto have "\y2 + b * y + c < 0" proof clarsimp fix y assume "y < x" then have "sign_num (a*x^2 + b*x + c) = sign_num (a*y^2 + b*y + c)" using nozero by (metis changes_sign_var) then show "a * y\<^sup>2 + b * y + c < 0" unfolding sign_num_def using x_prop inset by (smt split_conv) qed then show "\x. \y2 + b * y + c < 0" by auto qed then show ?case using notneginf by auto next case (Cons w xa) (* Need to argue that x isn't greater than the largest element of ?srl *) (* that if srl has length \ 2, x isn't in between any of the roots of ?srl*) (* and that x isn't less than the lowest root in ?srl *) then have lengthsrl: "length ?srl > 0" by auto have neginf: "x < nth_default 0 ?srl 0 \ False" proof - assume xlt: "x < nth_default 0 ?srl 0" have all: "(\(a, b, c)\set les. \y2 + b * y + c < 0)" proof clarsimp fix a b c y assume inset: "(a, b, c) \ set les" assume "y < x" have xl: "a*x^2 + b*x + c < 0" using x_prop inset by auto have "\(\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0)" proof - have "(\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0) \ False" proof - assume "\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0" then obtain q where q_prop: "q < nth_default 0 ?srl 0 \a*q^2 + b*q + c = 0" by auto then have " q \ root_set (set les)" unfolding root_set_def using inset by auto then have "List.member ?srl q" using same_set by (simp add: in_set_member) then have "q \ nth_default 0 ?srl 0" using sorted_sorted_list_of_set[where A = "root_set (set les)"] unfolding sorted_root_list_set_def by (metis \q \ root_set (set les)\ in_set_conv_nth le_less_linear lengthsrl not_less0 nth_default_nth same_set sorted_nth_mono sorted_root_list_set_def) then show "False" using q_prop by auto qed then show ?thesis by auto qed then have "\(\q. q < x \ a*q^2 + b*q + c = 0)" using xlt by auto then show " a * y\<^sup>2 + b * y + c < 0" using xl changes_sign_var[where a = "a", where b = "b", where c = "c", where x = "y", where y = "x"] unfolding sign_num_def using \y < x\ less_eq_real_def zero_neq_numeral by fastforce qed have "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" proof clarsimp fix a b c assume "(a, b, c)\set les" then show "\x. \y2 + b * y + c < 0" using all by blast qed then show "False" using notneginf by auto qed have "x > nth_default 0 ?srl (length ?srl - 1) \ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - assume xgt: "x > nth_default 0 ?srl (length ?srl - 1)" let ?lg = "nth_default 0 ?srl (length ?srl - 1)" have "List.member ?srl ?lg" by (metis diff_less in_set_member lengthsrl nth_default_def nth_mem zero_less_one) then have "?lg \ root_set (set les) " using same_set in_set_member[of ?lg ?srl] by auto then have exabc: "\(a, b, c)\set les. a*?lg^2 + b*?lg + c = 0" unfolding root_set_def by auto have "(\(d, e, f)\set les. \q\{?lg<..x}. d * q^2 + e * q + f < 0)" proof clarsimp fix d e f q assume inset: "(d, e, f) \ set les" assume qgt: "(nth_default 0) (sorted_root_list_set (set les)) (length (sorted_root_list_set (set les)) - Suc 0) < q" assume qlt: "q \ x" have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?lg)" proof - have "(\r. d * r^2 + e * r + f = 0 \ r > ?lg) \ False " proof - assume "\r. d * r^2 + e * r + f = 0 \ r > ?lg" then obtain r where r_prop: "d*r^2 + e*r + f = 0 \ r > ?lg" by auto then have "r \ root_set (set les)" using inset unfolding root_set_def by auto then have "List.member ?srl r" using same_set in_set_member by (simp add: in_set_member) then have " r \ ?lg" using sorted_sorted_list_of_set nth_default_def by (metis One_nat_def Suc_pred \r \ root_set (set les)\ in_set_conv_nth lengthsrl lessI less_Suc_eq_le same_set sorted_nth_mono sorted_root_list_set_def) then show "False" using r_prop by auto qed then show ?thesis by auto qed then have xltz_helper: "\(\r. r \ q \ d * r^2 + e * r + f = 0)" using qgt by auto then have xltz: "d*x^2 + e*x + f < 0" using inset x_prop by auto show "d * q\<^sup>2 + e * q + f < 0" using qlt qgt nor changes_sign_var[of d _ e f _] xltz xltz_helper unfolding sign_num_def apply (auto) by smt qed then have " (\(d, e, f)\set les. \y'>?lg. \x\{?lg<..y'}. d * x\<^sup>2 + e * x + f < 0)" using xgt by auto then have "(\(a, b, c)\set les. a*?lg^2 + b*?lg + c = 0 \ (\(d, e, f)\set les. \y'>?lg. \x\{?lg<..y'}. d * x\<^sup>2 + e * x + f < 0))" using exabc by auto then show "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" by auto qed then have posinf: "x > nth_default 0 ?srl (length ?srl - 1) \ False" using nok by auto have "(\n. (n+1) < (length ?srl) \ x > (nth_default 0 ?srl) n \ x < (nth_default 0 ?srl (n + 1))) \ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - assume "\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)" then obtain n where n_prop: "(n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)" by auto let ?elt = "nth_default 0 ?srl n" let ?elt2 = "nth_default 0 ?srl (n + 1)" have "List.member ?srl ?elt" using n_prop nth_default_def by (metis add_lessD1 in_set_member nth_mem) then have "?elt \ root_set (set les) " using same_set in_set_member[of ?elt ?srl] by auto then have exabc: "\(a, b, c)\set les. a*?elt^2 + b*?elt + c = 0" unfolding root_set_def by auto then obtain a b c where "(a, b, c)\set les \ a*?elt^2 + b*?elt + c = 0" by auto have xltel2: "x < ?elt2" using n_prop by auto have xgtel: "x > ?elt " using n_prop by auto have "(\(d, e, f)\set les. \q\{?elt<..x}. d * q^2 + e * q + f < 0)" proof clarsimp fix d e f q assume inset: "(d, e, f) \ set les" assume qgt: "nth_default 0 (sorted_root_list_set (set les)) n < q" assume qlt: "q \ x" have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?elt \r < ?elt2)" proof - have "(\r. d * r^2 + e * r + f = 0 \ r > ?elt \ r < ?elt2) \ False " proof - assume "\r. d * r^2 + e * r + f = 0 \ r > ?elt \ r < ?elt2" then obtain r where r_prop: "d*r^2 + e*r + f = 0 \ r > ?elt \ r < ?elt2" by auto then have "r \ root_set (set les)" using inset unfolding root_set_def by auto then have "List.member ?srl r" using same_set in_set_member by (simp add: in_set_member) then have "\i < (length ?srl). r = nth_default 0 ?srl i" by (metis \r \ root_set (set les)\ in_set_conv_nth same_set nth_default_def) then obtain i where i_prop: "i < (length ?srl) \ r = nth_default 0 ?srl i" by auto have "r > ?elt" using r_prop by auto then have igt: " i > n" using i_prop sorted_sorted_list_of_set by (smt add_lessD1 leI n_prop nth_default_def sorted_nth_mono sorted_root_list_set_def) have "r < ?elt2" using r_prop by auto then have ilt: " i < n + 1" using i_prop sorted_sorted_list_of_set by (smt leI n_prop nth_default_def sorted_nth_mono sorted_root_list_set_def) then show "False" using igt ilt by auto qed then show ?thesis by auto qed then have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?elt \r \ x)" using xltel2 xgtel by auto then have xltz: "d*x^2 + e*x + f < 0" using inset x_prop by auto show "d * q\<^sup>2 + e * q + f < 0" using qlt qgt nor changes_sign_var[of d _ e f _] xltz unfolding sign_num_def by smt qed then have " (\(d, e, f)\set les. \y'>?elt. \x\{?elt<..y'}. d * x\<^sup>2 + e * x + f < 0)" using xgtel xltel2 by auto then have "(\(a, b, c)\set les. a*?elt^2 + b*?elt + c = 0 \ (\(d, e, f)\set les. \y'>?elt. \x\{?elt<..y'}. d * x\<^sup>2 + e * x + f < 0))" using exabc by auto then show "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" by auto qed then have inbetw: "(\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)) \ False" using nok by auto have lenzer: "length xa = 0 \ False" proof - assume "length xa = 0" have xis: "x > w \ x < w" using notmem Cons.hyps by (smt list.set_intros(1) same_set xnotin) have xgt: "x > w \ False" proof - assume xgt: "x > w" show "False" using posinf Cons.hyps by (metis One_nat_def Suc_eq_plus1 \length xa = 0\ cancel_comm_monoid_add_class.diff_cancel list.size(4) nth_default_Cons_0 xgt) qed have xlt: "x < w \ False" proof - assume xlt: "x < w" show "False" using neginf Cons.hyps by (metis nth_default_Cons_0 xlt) qed show "False" using xis xgt xlt by auto qed have lengt: "length xa > 0 \ False" proof - assume "length xa > 0" have "x \ nth_default 0 ?srl 0" using neginf by fastforce then have xgtf: "x > nth_default 0 ?srl 0" using notmem using Cons.hyps(2) by fastforce have "x \ nth_default 0 ?srl (length ?srl - 1)" using posinf by fastforce then have "(\n. (n+1) < (length ?srl) \ x \ nth_default 0 ?srl n \ x \ nth_default 0 ?srl (n + 1))" using lengthsrl xgtf notmem sorted_list_prop[where l = ?srl, where x = "x"] by (metis add_lessD1 diff_less nth_default_nth sorted_root_list_set_def sorted_sorted_list_of_set zero_less_one) then obtain n where n_prop: "(n+1) < (length ?srl) \ x \ nth_default 0 ?srl n \ x \ nth_default 0 ?srl (n + 1)" by auto then have "x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n+1)" using notmem by (metis Suc_eq_plus1 Suc_lessD less_eq_real_def) then have "(\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1))" using n_prop by blast then show "False" using inbetw by auto qed then show ?case using lenzer lengt by auto qed qed show "False" using h1 h2 by auto qed then have equiv_false: "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))) \ False" by linarith have "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))) \ False" proof - assume "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" then have "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))" by auto then show ?thesis using equiv_false by auto qed then show ?thesis by blast qed lemma les_qe : shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) = ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" proof - have first: "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" using les_qe_backward by auto have second: "((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) " using les_qe_forward by auto have "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" using first second by meson then show ?thesis by blast qed subsubsection "equiv\\_lemma" lemma equiv_lemma: assumes big_asm: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" shows "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - let ?t = " ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" have h1: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ ?t" by auto have h2: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ ?t" by auto have h3: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ?t" by auto show ?thesis using big_asm h1 h2 h3 by presburger qed subsubsection "The eq\\_qe lemma" lemma eq_qe_forwards: shows "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - let ?big_or = "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" assume asm: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) " then obtain x where x_prop: "(\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" by auto have "\ (\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ False" proof - assume big_conj: "\ (\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" have not_lin: "\(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0))" using big_conj by auto have not_quad1: "\(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)))" using big_conj by auto have not_quad2: "\(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" using big_conj by auto have not_zer: "\((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using big_conj by auto then have not_zer1: "\(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" by auto have "(\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" using asm by auto then have "\(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0)" using not_zer1 by auto then have "\ (d, e, f)\set eq. d \ 0 \ e \ 0 \ f \ 0 " by auto then obtain d e f where def_prop: "(d, e, f) \ set eq \ (d \ 0 \ e \ 0 \ f \ 0)" by auto then have eval_at_x: "d*x^2 + e*x + f = 0" using x_prop by auto have dnonz: "d \ 0 \ False" proof - assume dneq: "d \ 0" then have discr: "-(e^2) + 4 *d *f \ 0" using discriminant_negative[of d e f x] eval_at_x unfolding discrim_def by linarith let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 *d *f)) / (2 * d)" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 *d *f)) / (2 * d)" have xis: "x = ?r1 \ x = ?r2" using dneq discr discriminant_nonneg[of d e f x] eval_at_x unfolding discrim_def by auto have xr1: "x = ?r1 \ False" using not_quad2 x_prop discr def_prop dneq by auto have xr2: "x = ?r2 \ False" using not_quad1 x_prop discr def_prop dneq by auto show "False" using xr1 xr2 xis by auto qed then have dz: "d = 0" by auto have enonz: "e \ 0 \ False" proof - assume enonz: "e\ 0" then have "x = -f/e" using dz eval_at_x by (metis add.commute minus_add_cancel mult.commute mult_zero_class.mult_zero_left nonzero_eq_divide_eq) then show "False" using not_lin x_prop enonz dz def_prop by auto qed then have ez: "e = 0" by auto have fnonz: "f \ 0 \ False" using ez dz eval_at_x by auto show "False" using def_prop dnonz enonz fnonz by auto qed then have h: "\(?big_or) \ False" by auto then show ?thesis using equiv_lemma by presburger qed lemma eq_qe_backwards: "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. ((\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))) " proof - assume "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then have bigor: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" by auto have h1: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(x::real). (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set eq \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)" by auto let ?x = "(-c' /b')::real" have "(\(d, e, f)\set eq. d * ?x\<^sup>2 + e * ?x + f = 0) \ (\(d, e, f)\set les. d * ?x^2 + e * ?x + f < 0)" using abc_prop by auto then show ?thesis using abc_prop by blast qed have h2: " (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)))" then obtain a' b' c' where abc_prop: "(a', b', c')\set eq \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" by auto let ?x = "((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')::real)" have anonz: "a' \ 0" using abc_prop by auto then have "\(q::real). q = ?x" by auto then obtain q where q_prop: "q = ?x" by auto have "(\(d, e, f)\set eq. d * (?x)\<^sup>2 + e * (?x) + f = 0) \ (\(d, e, f)\set les. d * (?x)\<^sup>2 + e * (?x) + f < 0)" using abc_prop by auto then have "(\(d, e, f)\set eq. d * q\<^sup>2 + e * q + f = 0) \ (\(d, e, f)\set les. d * q\<^sup>2 + e * q + f < 0)" using q_prop by auto then show ?thesis by auto qed have h3: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" then obtain a' b' c' where abc_prop: "a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (a', b', c')\set eq \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)" by auto let ?x = "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have anonz: "a' \ 0" using abc_prop by auto then have "\(q::real). q = ?x" by auto then obtain q where q_prop: "q = ?x" by auto have "(\(d, e, f)\set eq. d * (?x)\<^sup>2 + e * (?x) + f = 0) \ (\(d, e, f)\set les. d * (?x)\<^sup>2 + e * (?x) + f < 0)" using abc_prop by auto then have "(\(d, e, f)\set eq. d * q\<^sup>2 + e * q + f = 0) \ (\(d, e, f)\set les. d * q\<^sup>2 + e * q + f < 0)" using q_prop by auto then show ?thesis by auto qed have h4: "((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then have allzer: "(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0)" by auto have "(\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" using asm by auto then obtain x where x_prop: " \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0" by auto then have "\(d, e, f)\set eq. d*x^2 + e*x + f = 0" using allzer by auto then show ?thesis using x_prop by auto qed show ?thesis using bigor h1 h2 h3 h4 by blast qed lemma eq_qe : "(\x. ((\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))) = ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - have h1: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using eq_qe_forwards by auto have h2: "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using eq_qe_backwards by auto have h3: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using h1 h2 by smt then show ?thesis by (auto) qed subsubsection "The qe\\_forwards lemma" lemma qe_forwards_helper_gen: fixes r:: "real" assumes f8: "\(\((a'::real), (b'::real), (c'::real))\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have h1: "(\(d, e, f)\set a. d * r^2 + e * r + f = 0)" using alleqset by blast have c_prop: "(\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h2: "(\(d, e, f)\set c. d *r^2 + e * r + f \ 0)" proof - have c1: "\ (d, e, f) \ set c. d * (r)\<^sup>2 + e * (r) + f > 0 \ False" proof - assume "\ (d, e, f) \ set c. d * (r)\<^sup>2 + e * (r) + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set c \ d * (r)\<^sup>2 + e * r + f > 0" by auto have "\y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0" using def_prop c_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\x\{r<..y'}. d*x^2 + e*x + f > 0" using def_prop continuity_lem_gt0_expanded[of "r" y' d e f] using y_prop by linarith then show "False" using y_prop by auto qed then show ?thesis by fastforce qed have b_prop: "(\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" using abc_prop by auto have h3: "(\(d, e, f)\set b. d * r\<^sup>2 + e * r + f < 0)" proof - have c1: "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f > 0 \ False" proof - assume "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set b \ d * r\<^sup>2 + e * r + f > 0" by auto then have "\y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0" using b_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > r \ k < y' \ d * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then obtain k where k_prop: "k > r \ k < y' \ d * k^2 + e * k + f < 0" by auto then have "\(\x>r. x < y' \ d * x\<^sup>2 + e * x + f = 0)" using y_prop by force then show "False" using k_prop def_prop y_prop poly_IVT_neg[of "r" "k" "[:f, e, d:]"] poly_IVT_pos[of "-c'/b'" "k" "[:f, e, d:]"] by (smt quadratic_poly_eval) qed have c2: "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f = 0 \ False" proof - assume "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set b \ d' * r\<^sup>2 + e * r + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = r)" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * r + f = 0" using def_prop by auto then show "-f/e = r" using asm by (metis (no_types) add.commute diff_0 divide_minus_left minus_add_cancel nonzero_mult_div_cancel_left uminus_add_conv_diff) qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set b \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f5 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>r. \x\{r<..y'}. d' * x\<^sup>2 + e * x + f < 0" using b_prop def_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d' * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > r \ k < y' \ d' * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def using def_prop by fastforce then have two_cases: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ r = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f7 by auto qed have cf2: "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f6 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed show ?thesis using c1 c2 by fastforce qed have d_prop: "(\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h4: "(\(d, e, f)\set d. d * r\<^sup>2 + e * r + f \ 0)" proof - have "(\(d, e, f)\set d. d * r\<^sup>2 + e * r + f = 0) \ False" proof - assume "\ (d, e, f) \ set d. d * r\<^sup>2 + e * r + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set d \ d' * r\<^sup>2 + e * r + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = r)" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * r + f = 0" using def_prop by auto then show "-f/e = r" using asm by (metis (no_types) add.commute diff_0 divide_minus_left minus_add_cancel nonzero_mult_div_cancel_left uminus_add_conv_diff) qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set d \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f10 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'> r. \x\{ r<..y'}. d' * x\<^sup>2 + e * x + f \ 0" using d_prop def_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d' * x\<^sup>2 + e * x + f \ 0)" by auto then have "\k. k > r \ k < y' \ d' * k^2 + e * k + f \ 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ r = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f12 by auto qed have cf2: "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f11 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed then show ?thesis by auto qed have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r\<^sup>2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r\<^sup>2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r\<^sup>2 + e * r + f \ 0))" using h1 h2 h3 h4 abc_prop by auto then show "False" using f8 by auto qed then show ?thesis by auto qed lemma qe_forwards_helper_lin: assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f8: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then have bnonz: "b'\0" by auto have h1: "(\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0)" using bnonz alleqset by blast have c_prop: "(\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h2: "(\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" proof - have c1: "\ (d, e, f) \ set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0 \ False" proof - assume "\ (d, e, f) \ set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set c \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0" using def_prop c_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\x\{(-c'/b')<..y'}. d*x^2 + e*x + f > 0" using def_prop continuity_lem_gt0_expanded[of "(-c'/b')" y' d e f] using y_prop by linarith then show "False" using y_prop by auto qed then show ?thesis by fastforce qed have b_prop: "(\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0)" using abc_prop by auto have h3: "(\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)" proof - have c1: "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0 \ False" proof - assume "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set b \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" by auto then have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0" using b_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then obtain k where k_prop: "k > -c'/b' \ k < y' \ d * k^2 + e * k + f < 0" by auto then have "\(\x>(-c'/b'). x < y' \ d * x\<^sup>2 + e * x + f = 0)" using y_prop by force then show "False" using k_prop def_prop y_prop poly_IVT_neg[of "-c'/b'" "k" "[:f, e, d:]"] poly_IVT_pos[of "-c'/b'" "k" "[:f, e, d:]"] by (smt quadratic_poly_eval) qed have c2: "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0 \ False" proof - assume "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set b \ d' * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = -c'/b')" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * (- c' / b') + f = 0" using def_prop by auto then show "-f/e = -c'/b'" using asm by auto qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set b \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f5 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f < 0" using b_prop def_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d' * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ (- c' / b') = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f7 by auto qed have cf2: "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f6 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed show ?thesis using c1 c2 by fastforce qed have d_prop: "(\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h4: "(\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" proof - have "(\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ False" (* begin *) proof - assume "\ (d, e, f) \ set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set d \ d' * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = -c'/b')" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * (- c' / b') + f = 0" using def_prop by auto then show "-f/e = -c'/b'" using asm by auto qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set d \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f10 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f \ 0" using d_prop def_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f \ 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d' * k^2 + e * k + f \ 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ (- c' / b') = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f12 by auto qed have cf2: "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f11 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed then show ?thesis by auto qed have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using h1 h2 h3 h4 bnonz abc_prop by auto then show "False" using f8 by auto qed then show ?thesis by auto qed lemma qe_forwards_helper: assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f1: "\((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f8: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" assumes f13: "\(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" assumes f9: "\(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - have nor: "\r. \(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" proof clarsimp fix r t u v assume inset: "(t, u, v) \ set c" assume eo: "t = 0 \ u \ 0 " assume zero_eq: "t*r^2 + u*r + v = 0" assume ah: "\x\set a. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f = 0" assume bh: "\x\set b. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f < 0" assume ch: "\x\set c. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f \ 0" assume dh: "\x\set d. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f \ 0" have two_cases: "t \ 0 \ (t = 0 \ u \ 0)" using eo by auto have c1: "t \ 0 \ False" proof - assume tnonz: "t \ 0" then have discr_prop: "- u\<^sup>2 + 4 * t * v \ 0 " using discriminant_negative[of t u v] zero_eq unfolding discrim_def by force then have ris: "r = ((-u + - 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ r = ((-u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) " using tnonz discriminant_nonneg[of t u v] zero_eq unfolding discrim_def by auto let ?r1 = "((-u + - 1 * sqrt (u^2 - 4 * t * v)) / (2 * t))" let ?r2 = "((-u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t))" have ris1: "r = ?r1 \ False" proof - assume "r = ?r1" then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using inset ah bh ch dh discr_prop tnonz by auto then show ?thesis using f9 by auto qed have ris2: "r = ?r2 \ False" proof - assume "r = ?r2" then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" using inset ah bh ch dh discr_prop tnonz by auto then show ?thesis using f13 by auto qed show "False" using ris ris1 ris2 by auto qed have c2: "(t = 0 \ u \ 0) \ False" proof - assume asm: "t = 0 \ u \ 0" then have "r = -v/u" using zero_eq add.right_neutral nonzero_mult_div_cancel_left by (metis add.commute divide_divide_eq_right divide_eq_0_iff neg_eq_iff_add_eq_0) then have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using asm inset ah bh ch dh by auto then show "False" using f8 by auto qed then show "False" using two_cases c1 c2 by auto qed have keyh: "\r. \(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - fix r have h8: "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" using nor by auto show "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using qe_forwards_helper_gen[of c r a b d] alleqset f5 f6 f7 h8 f10 f11 f12 by auto qed have f8a: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using qe_forwards_helper_lin[of a b c d] alleqset f5 f6 f7 f8 f10 f11 f12 by blast have f13a: "\ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" proof - have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ False" proof - assume "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?r = "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have somek: "\k. k = ?r" by auto then obtain k where k_prop: "k = ?r" by auto have "(a'\ 0 \ b'\ 0) \ (a'*?r^2 + b'*?r + c' = 0)" using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def apply (auto) by (metis (mono_tags, lifting) times_divide_eq_right) then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*?r^2 + b'*?r + c' = 0) \ (\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop by auto then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using k_prop by auto then have "\k. (\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then show "False" using keyh by auto qed then show ?thesis by auto qed have f9a: "\ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto let ?r = "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have somek: "\k. k = ?r" by auto then obtain k where k_prop: "k = ?r" by auto have "(a'\ 0 \ b'\ 0) \ (a'*?r^2 + b'*?r + c' = 0)" using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def apply (auto) by (metis (mono_tags, lifting) times_divide_eq_right) then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*?r^2 + b'*?r + c' = 0) \ (\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop by auto then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using k_prop by auto then have "\k. (\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then show "False" using keyh by auto qed then show ?thesis by auto qed (* We need to show that the point is in one of these ranges *) have "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) \ False" proof - assume "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" then obtain x where x_prop: "(\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" by auto (* Need this sorted_nonzero_root_list_set in case some of the tuples from set c are (0, 0, 0) *) let ?srl = "sorted_nonzero_root_list_set (((set b) \ (set c))\ (set d))" have alleqsetvar: "\(t, u, v) \ set a. t = 0 \ u = 0 \ v = 0" proof clarsimp fix t u v assume "(t, u, v) \ set a" then have "\x. t*x^2 + u*x + v = 0" using alleqset by auto then have "\x\{0<..1}. t * x\<^sup>2 + u * x + v = 0" by auto then show "t = 0 \ u = 0 \ v = 0" using continuity_lem_eq0[of 0 1 t u v] by auto qed (* Should violate f1 *) have lenzero: "length ?srl = 0 \ False" proof - assume lenzero: "length ?srl = 0" have ina: "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have inb: "(\(a, b, c)\set b. \y. a * y\<^sup>2 + b * y + c < 0)" proof clarsimp fix t u v y assume insetb: "(t, u, v) \ set b" then have "t * x\<^sup>2 + u * x + v < 0" using x_prop by auto then have tuv_prop: "t \ 0 \ u \ 0 \ v \ 0" by auto then have tuzer: "(t = 0 \ u = 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" by simp then have tunonz: "(t \ 0 \ u \ 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" proof - assume tuv_asm: "t \ 0 \ u \ 0" have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetb tuv_asm tuv_prop by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1" using insetb unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1" using samesn by auto then show "t * y\<^sup>2 + u * y + v < 0" unfolding sign_num_def by smt qed have inc: "(\(a, b, c)\set c. \y. a * y\<^sup>2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetc: "(t, u, v) \ set c" then have "t * x\<^sup>2 + u * x + v \ 0" using x_prop by auto then have tuzer: "t = 0 \ u = 0 \ t*y^2 + u*y + v \ 0 " proof - assume tandu: "t = 0 \ u = 0" then have "v \ 0" using insetc x_prop by auto then show "t*y^2 + u*y + v \ 0" using tandu by auto qed have tunonz: "t \ 0 \ u \ 0 \ t*y^2 + u*y + v \ 0" proof - assume tuv_asm: "t \ 0 \ u \ 0" have insetcvar: "t*x^2 + u*x + v < 0" proof - have "t*x^2 + u*x + v = 0 \ False" proof - assume zer: "t*x^2 + u*x + v = 0" then have xin: "x \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetc tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "x \ set ?srl" using xin unfolding nonzero_root_set_def by auto then have "List.member ?srl x" using in_set_member[of x ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis using \t * x\<^sup>2 + u * x + v \ 0\ by fastforce qed then have tunonz: "\(\q. t * q\<^sup>2 + u * q + v = 0)" proof - have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetc tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1" using insetcvar unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1" using samesn by auto then show "t * y\<^sup>2 + u * y + v \ 0" unfolding sign_num_def by smt qed then show "t * y\<^sup>2 + u * y + v \ 0" using tuzer tunonz by blast qed have ind: "(\(a, b, c)\set d. \y. a * y\<^sup>2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetd: "(t, u, v) \ set d" assume falseasm: "t * y\<^sup>2 + u * y + v = 0" then have snz: "sign_num (t*y^2 + u*y + v) = 0" unfolding sign_num_def by auto have "t * x\<^sup>2 + u * x + v \ 0" using insetd x_prop by auto then have tuv_prop: "t \ 0 \ u \ 0 \ v \ 0" by auto then have tuzer: "(t = 0 \ u = 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" by simp then have tunonz: "(t \ 0 \ u \ 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" proof - assume tuv_asm: "t \ 0 \ u \ 0" have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetd tuv_asm tuv_prop by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 1 " using insetd unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1 \ sign_num (t*y^2 + u*y + v) = 1" using samesn by auto then show "False" using snz by auto qed (* Show all the polynomials never change sign *) have "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \y. a * y\<^sup>2 + b * y + c < 0) \ (\(a, b, c)\set c. \y. a * y\<^sup>2 + b * y + c \ 0) \ (\(a, b, c)\set d. \y. a * y\<^sup>2 + b * y + c \ 0))" using ina inb inc ind by auto then show "False" using f1 by auto qed have cases_mem: "(List.member ?srl x) \ False" proof - assume "(List.member ?srl x)" then have "x \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using set_sorted_list_of_set nonzero_root_set_finite in_set_member by (metis List.finite_set finite_Un nonzero_root_set_def sorted_nonzero_root_list_set_def) then have "\ (a, b, c) \ (((set b) \ (set c))\ (set d)) . (a \ 0 \ b \ 0) \ a*x^2 + b*x + c = 0" by blast then obtain t u v where def_prop: "(t, u, v) \ (((set b) \ (set c))\ (set d)) \ (t \ 0 \ u \ 0) \ t*x^2 + u*x + v = 0" by auto have notinb: "(t, u, v) \ (set b)" proof - have "(t, u, v) \ (set b ) \ False" proof - assume "(t, u, v) \ (set b)" then have "t*x^2 + u*x + v < 0" using x_prop by blast then show "False" using def_prop by simp qed then show ?thesis by auto qed have notind: "(t, u, v) \ (set d)" proof - have "(t, u, v) \ (set d) \ False" proof - assume "(t, u, v) \ (set d)" then have "t*x^2 + u*x + v \ 0" using x_prop by blast then show "False" using def_prop by simp qed then show ?thesis by auto qed then have inset: "(t, u, v) \ (set c)" using def_prop notinb notind by blast have case1: "t \ 0 \ False" proof - assume tnonz: "t \ 0" then have r1or2:"x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ x = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) " using def_prop discriminant_negative[of t u v] discriminant_nonneg[of t u v] apply (auto) using notinb apply (force) apply (simp add: discrim_def discriminant_iff) using notind by force have discrh: "-1*u^2 + 4 * t * v \ 0" using tnonz discriminant_negative[of t u v] unfolding discrim_def using def_prop by force have r1: "x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ False" proof - assume xis: "x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)" have " t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x\<^sup>2 + e * x + f \ 0)" using tnonz alleqset discrh x_prop by auto then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using xis inset by auto then show "False" using f9 by auto qed have r2: "x = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ False" proof - assume xis: "x = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)" have " t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x\<^sup>2 + e * x + f \ 0)" using tnonz alleqset discrh x_prop by auto then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using xis inset by auto then show "False" using f13 by auto qed then show "False" using r1or2 r1 r2 by auto qed have case2: "(t = 0 \ u \ 0) \ False" proof - assume asm: "t = 0 \ u \ 0" then have xis: "x = - v / u" using def_prop notinb add.commute diff_0 divide_non_zero minus_add_cancel uminus_add_conv_diff by (metis mult_zero_left) have "((t = 0 \ u \ 0) \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x^2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x^2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x^2 + e * x + f \ 0))" using asm x_prop alleqset by auto then have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using xis inset by auto then show "False" using f8 by auto qed show "False" using def_prop case1 case2 by auto qed have lengt0: "length ?srl \ 1 \ False" proof- assume asm: "length ?srl \ 1" (* should violate f1 *) have cases_lt: "x < ?srl ! 0 \ False" proof - assume xlt: "x < ?srl ! 0" have samesign: "\ (a, b, c) \ (set b \ set c \ set d). (\y < x. sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ylt: "y < x" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have "\(\q. q < ?srl ! 0 \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qlt: "q < sorted_nonzero_root_list_set (set b \ set c \ set d) ! 0" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have lm: "List.member ?srl q" using in_set_member[of q ?srl] by auto then have " List.member (sorted_list_of_set (nonzero_root_set (set b \ set c \ set d))) q \ q < sorted_list_of_set (nonzero_root_set (set b \ set c \ set d)) ! 0 \ (\x xs. (x \ set xs) = (\i (\x xs. (x \ set xs) = List.member xs x) \ (\y x. \ y \ x \ x < y) \ (\xs. sorted xs = (\i j. i \ j \ j < length xs \ xs ! i \ xs ! j)) \ (\p. sorted_nonzero_root_list_set p \ sorted_list_of_set (nonzero_root_set p)) \ False" proof - assume a1: "List.member (sorted_list_of_set (nonzero_root_set (set b \ set c \ set d))) q" assume a2: "q < sorted_list_of_set (nonzero_root_set (set b \ set c \ set d)) ! 0" have f3: "List.member (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q" using a1 by (metis nonzero_root_set_def) have f4: "q < sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! 0" using a2 by (metis nonzero_root_set_def) have f5: "q \ set (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)})" using f3 by (meson in_set_member) have "\rs r. \n. ((r::real) \ set rs \ n < length rs) \ (r \ set rs \ rs ! n = r)" by (metis in_set_conv_nth) then obtain nn :: "real list \ real \ nat" where f6: "\r rs. (r \ set rs \ nn rs r < length rs) \ (r \ set rs \ rs ! nn rs r = r)" by moura then have "sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q = q" using f5 by blast then have "\n. \ sorted (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) \ \ n \ nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q \ \ nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q < length (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) \ sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! n \ q" using not_less not_less0 sorted_iff_nth_mono by (metis (no_types, lifting)) then show ?thesis using f6 f5 f4 by (meson le0 not_less sorted_sorted_list_of_set) qed then show "False" using lm qlt in_set_conv_nth in_set_member not_le_imp_less not_less0 sorted_iff_nth_mono sorted_nonzero_root_list_set_def sorted_sorted_list_of_set by auto qed then have "\(\q. q \ x \ t * q\<^sup>2 + u * q + v = 0)" using xlt by auto then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using ylt changes_sign_var[of t y u v x] by blast qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed have bseth: "(\(a, b, c)\set b. \y2 + b * y + c < 0)" proof clarsimp fix t u v y assume insetb: "(t, u, v) \ set b" assume yltx: "y < x" have "(t, u, v) \ (set b \ set c \ set d)" using insetb by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetb yltx by blast have "sign_num (t*x^2 + u*x + v) = -1" using x_prop insetb unfolding sign_num_def by auto then show "t * y\<^sup>2 + u * y + v < 0" using samesn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have bset: " (\(a, b, c)\set b. \x. \y2 + b * y + c < 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set b" then have " \y2 + u * y + v < 0 " using bseth by auto then show "\x. \y2 + u * y + v < 0" by auto qed have cseth: "(\(a, b, c)\set c. \y2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetc: "(t, u, v) \ set c" assume yltx: "y < x" have "(t, u, v) \ (set b \ set c \ set d)" using insetc by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetc yltx by blast have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 0" using x_prop insetc unfolding sign_num_def by auto then show "t * y\<^sup>2 + u * y + v \ 0" using samesn unfolding sign_num_def using zero_neq_one by fastforce qed have cset: " (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set c" then have " \y2 + u * y + v \ 0 " using cseth by auto then show "\x. \y2 + u * y + v \0" by auto qed have dseth: "(\(a, b, c)\set d. \y2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetd: "(t, u, v) \ set d" assume yltx: "y < x" assume contrad: "t * y\<^sup>2 + u * y + v = 0" have "(t, u, v) \ (set b \ set c \ set d)" using insetd by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetd yltx by blast have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 1" using x_prop insetd unfolding sign_num_def by auto then have "t * y\<^sup>2 + u * y + v \ 0" using samesn unfolding sign_num_def by auto then show "False" using contrad by auto qed have dset: " (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set d" then have " \y2 + u * y + v \ 0 " using dseth by auto then show "\x. \y2 + u * y + v \ 0" by auto qed have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto then have "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" using bset cset dset by auto then show "False" using f1 by auto qed (* should violate one of the infinitesmials *) have cases_gt: " x > ?srl ! (length ?srl - 1) \ False" proof - assume xgt: "x > ?srl ! (length ?srl - 1)" let ?bgrt = "?srl ! (length ?srl - 1)" have samesign: "\ (a, b, c) \ (set b \ set c \ set d). (\y > ?bgrt. sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ygt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < y" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have "\(\q. q > ?srl ! (length ?srl - 1) \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qgt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < q" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using qgt in_set_conv_nth in_set_member not_le_imp_less not_less0 sorted_iff_nth_mono sorted_nonzero_root_list_set_def sorted_sorted_list_of_set by (smt (z3) Suc_diff_Suc Suc_n_not_le_n \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ in_set_conv_nth length_0_conv length_greater_0_conv length_sorted_list_of_set lenzero less_Suc_eq_le minus_nat.diff_0 not_le sorted_nth_mono sorted_sorted_list_of_set) qed then have nor: "\(\q. q > ?bgrt \ t * q\<^sup>2 + u * q + v = 0)" using xgt by auto have c1: " x > y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t y u v x] xgt ygt by fastforce then have c2: " y > x \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t x u v y] xgt ygt by force then have c3: " x = y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto then show "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using c1 c2 c3 by linarith qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have " ?bgrt \ set ?srl" using set_sorted_list_of_set nonzero_root_set_finite in_set_member using asm by auto then have "?bgrt \ nonzero_root_set (set b \ set c \ set d )" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set nonzero_root_set_finite by auto then have "\t u v. (t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" unfolding nonzero_root_set_def by auto then obtain t u v where tuvprop1: "(t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" by auto then have tuvprop: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) " by auto have tnonz: "t\ 0 \ (-1*u^2 + 4 * t * v \ 0 \ (?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) \ ?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)))" proof - assume "t\ 0" have "-1*u^2 + 4 * t * v \ 0 " using tuvprop1 discriminant_negative[of t u v] unfolding discrim_def using \t \ 0\ by force then show ?thesis using tuvprop discriminant_nonneg[of t u v] unfolding discrim_def using \t \ 0\ by auto qed have unonz: "(t = 0 \ u \ 0) \ ?bgrt = - v / u" proof - assume "(t = 0 \ u \ 0)" then have "u*?bgrt + v = 0" using tuvprop1 by simp then show "?bgrt = - v / u" by (simp add: \t = 0 \ u \ 0\ eq_minus_divide_eq mult.commute) qed have allpropb: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set b" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign apply (auto) by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 < 0" using xsn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have allpropbvar: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set b" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using allpropb by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using less_add_one by (metis One_nat_def) qed have allpropc: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set c" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 0" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign One_nat_def proof - have "case (t1, u1, v1) of (r, ra, rb) \ \raa>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1). sign_num (r * raa\<^sup>2 + ra * raa + rb) = sign_num (r * x\<^sup>2 + ra * x + rb)" by (smt (z3) Un_iff ins samesign) then show ?thesis by (simp add: x1gt) qed then show "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by (metis equal_neg_zero less_numeral_extra(3) linorder_not_less zero_neq_one) qed have allpropcvar: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set c" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropc by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using less_add_one One_nat_def - by (metis (no_types, hide_lams)) + by metis qed have allpropd: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set d" assume contrad:"t1 * x1\<^sup>2 + u1 * x1 + v1 = 0" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign apply (auto) by blast then have "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by auto then show "False" using contrad by auto qed have allpropdvar: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set d" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropd by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" - using less_add_one - by (metis (no_types, hide_lams) One_nat_def) + using less_add_one by force qed have "\x. (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0)" using alleqsetvar by auto then have ast: "(\(d, e, f)\set a. \x\{?bgrt<..(?bgrt + 1)}. d * x\<^sup>2 + e * x + f = 0)" by auto have allpropavar: "(\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set a" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 = 0 " using ast by auto then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 = 0" using less_add_one One_nat_def by metis qed have quadsetb: "((t, u, v) \ set b \ t\ 0) \ False" proof - assume asm: "(t, u, v) \ set b \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f6 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f7 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetb: "((t, u, v) \ set b \ (t = 0 \ u \ 0)) \ False" proof - assume asm: "(t, u, v) \ set b \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set b \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f5 by auto qed have insetb: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetb linsetb by auto have quadsetc: "(t, u, v) \ set c \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set c \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f13a bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f9a bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetc: "(t, u, v) \ set c \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set c \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set c \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f8a by auto qed have insetc: "((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetc linsetc by auto have quadsetd: "(t, u, v) \ set d \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set d \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f11 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f12 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetd: "(t, u, v) \ set d \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set d \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set d \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f10 by auto qed have insetd: "((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetd linsetd by auto then show "False" using insetb insetc insetd tuvprop by auto qed have len1: "length ?srl = 1 \ False" proof - assume len1: "length ?srl = 1" have cases: "(List.member ?srl x) \ x < ?srl ! 0 \ x > ?srl ! 0" using in_set_member lenzero nth_mem by fastforce then show "False" using len1 cases_mem cases_lt cases_gt by auto qed have lengtone: "length ?srl > 1 \ False" proof - assume lengt1: "length ?srl > 1" have cases: "(List.member ?srl x) \ x < ?srl ! 0 \ x > ?srl ! (length ?srl -1) \ (\k \ (length ?srl - 2). (?srl ! k < x \ x x > ?srl ! (length ?srl -1) \ (x \ ?srl ! 0 \ x \ ?srl ! (length ?srl -1))" by auto have ifo: "(x \ ?srl ! 0 \ x \ ?srl ! (length ?srl -1)) \ ((List.member ?srl x) \ (\k \ (length ?srl - 2). ?srl ! k < x \ x ?srl ! 0 \ x \ ?srl ! (length ?srl -1)" then have "\(List.member ?srl x) \ (\k \ (length ?srl - 2). ?srl ! k < x \ x (List.member ?srl x)" have "\(\k \ (length ?srl - 2). ?srl ! k < x \ x False" proof clarsimp assume "\k. sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x \ k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2 \ \ x < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" then have allk: "(\k \ length ?srl - 2. ?srl ! k < x \ \ x < ?srl ! Suc k)" by auto have basec: "x \ ?srl ! 0" using xinbtw by auto have "\k \ length ?srl - 2. ?srl ! k < x" proof clarsimp fix k assume klteq: "k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2" show "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x" using nonmem klteq basec proof (induct k) case 0 then show ?case using in_set_member lenzero nth_mem by fastforce next case (Suc k) then show ?case by (smt Suc_leD Suc_le_lessD \\k. sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x \ k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2 \ \ x < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k\ diff_less in_set_member length_0_conv length_greater_0_conv lenzero less_trans_Suc nth_mem pos2) qed qed then have "x \ ?srl ! (length ?srl -1)" using allk by (metis One_nat_def Suc_diff_Suc lengt1 less_eq_real_def less_or_eq_imp_le one_add_one plus_1_eq_Suc xinbtw) then have "x > ?srl ! (length ?srl - 1)" using nonmem by (metis One_nat_def Suc_le_D asm diff_Suc_Suc diff_zero in_set_member lessI less_eq_real_def nth_mem) then show "False" using xinbtw by auto qed then show "(\k \ (length ?srl - 2). ?srl ! k < x \ x (\k \ (length ?srl - 2). ?srl ! k < x \ x k \ (length ?srl - 2). ?srl ! k < x \ x False" proof - assume "(\k \ (length ?srl - 2). ?srl ! k < x \ x (length ?srl - 2) \ ?srl ! k < x \ x (a, b, c) \ (set b \ set c \ set d). (\y. (?srl ! k < y \ y sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ygt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < y" assume ylt: "y < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have nor: "\(\q. q > ?srl ! k \ q < ?srl ! (k + 1) \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qlt: "q < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" assume qgt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < q" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then have "\n < length ?srl. q = ?srl ! n" by (metis \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ in_set_conv_nth) then obtain n where nprop: "n < length ?srl \ q = ?srl ! n" by auto then have ngtk: "n > k" proof - have sortedh: "sorted ?srl" by (simp add: sorted_nonzero_root_list_set_def) then have nlteq: "n \ k \ ?srl ! n \ ?srl ! k" using nprop k_prop sorted_iff_nth_mono using sorted_nth_mono - by (metis (no_types, hide_lams) Suc_1 \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ diff_Suc_less length_pos_if_in_set sup.absorb_iff2 sup.strict_boundedE) + by (metis Suc_1 \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ diff_less length_pos_if_in_set sup.absorb_iff2 sup.strict_boundedE zero_less_Suc) have "?srl ! n > ?srl ! k" using nprop qgt by auto then show ?thesis using nlteq by linarith qed then have nltkp1: "n < k+1" proof - have sortedh: "sorted ?srl" by (simp add: sorted_nonzero_root_list_set_def) then have ngteq: "k+1 \ n \ ?srl ! (k+1) \ ?srl ! n" using nprop k_prop sorted_iff_nth_mono by auto have "?srl ! n < ?srl ! (k + 1)" using nprop qlt by auto then show ?thesis using ngteq by linarith qed then show "False" using ngtk nltkp1 by auto qed have c1: " x > y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t y u v x] k_prop ygt ylt by fastforce then have c2: " y > x \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t x u v y] k_prop ygt ylt by force then have c3: " x = y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto then show "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using c1 c2 c3 by linarith qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed let ?bgrt = "?srl ! k" have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have " ?bgrt \ set ?srl" using set_sorted_list_of_set nonzero_root_set_finite in_set_member k_prop asm by (smt diff_Suc_less le_eq_less_or_eq less_le_trans nth_mem one_add_one plus_1_eq_Suc zero_less_one) then have "?bgrt \ nonzero_root_set (set b \ set c \ set d )" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set nonzero_root_set_finite by auto then have "\t u v. (t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" unfolding nonzero_root_set_def by auto then obtain t u v where tuvprop1: "(t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" by auto then have tuvprop: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) " by auto have tnonz: "t\ 0 \ (-1*u^2 + 4 * t * v \ 0 \ (?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) \ ?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)))" proof - assume "t\ 0" have "-1*u^2 + 4 * t * v \ 0 " using tuvprop1 discriminant_negative[of t u v] unfolding discrim_def using \t \ 0\ by force then show ?thesis using tuvprop discriminant_nonneg[of t u v] unfolding discrim_def using \t \ 0\ by auto qed have unonz: "(t = 0 \ u \ 0) \ ?bgrt = - v / u" proof - assume "(t = 0 \ u \ 0)" then have "u*?bgrt + v = 0" using tuvprop1 by simp then show "?bgrt = - v / u" by (simp add: \t = 0 \ u \ 0\ eq_minus_divide_eq mult.commute) qed have "\y'. y' > x \ y' < ?srl ! (k+1)" using k_prop dense by blast then obtain y1 where y1_prop: "y1 > x \ y1 < ?srl ! (k+1)" by auto then have y1inbtw: "y1 > ?srl ! k \ y1 < ?srl ! (k+1)" using k_prop by auto have allpropb: "(\(d, e, f)\set b. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set b" assume x1gt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 < 0" using xsn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have allpropbvar: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set b" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using allpropb by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using y1inbtw by blast qed have allpropc: "(\(d, e, f)\set c. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set c" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 0" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def - by (metis (no_types, hide_lams) equal_neg_zero less_eq_real_def linorder_not_less zero_neq_one) + by (smt (verit, del_insts)) qed have allpropcvar: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set c" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropc by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using y1inbtw by blast qed have allpropd: "(\(d, e, f)\set d. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set d" assume contrad:"t1 * x1\<^sup>2 + u1 * x1 + v1 = 0" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then have "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by auto then show "False" using contrad by auto qed have allpropdvar: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set d" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropd by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using y1inbtw by blast qed have "\x. (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0)" using alleqsetvar by auto then have ast: "(\(d, e, f)\set a. \x\{?bgrt<..(?bgrt + 1)}. d * x\<^sup>2 + e * x + f = 0)" by auto have allpropavar: "(\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set a" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 = 0 " using ast by auto then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 = 0" using less_add_one by blast qed have quadsetb: "((t, u, v) \ set b \ t\ 0) \ False" proof - assume asm: "(t, u, v) \ set b \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f6 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f7 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetb: "((t, u, v) \ set b \ (t = 0 \ u \ 0)) \ False" proof - assume asm: "(t, u, v) \ set b \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set b \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f5 by auto qed have insetb: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetb linsetb by auto have quadsetc: "(t, u, v) \ set c \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set c \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f13a bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f9a bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetc: "(t, u, v) \ set c \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set c \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set c \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f8a by auto qed have insetc: "((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetc linsetc by auto have quadsetd: "(t, u, v) \ set d \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set d \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f11 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f12 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetd: "(t, u, v) \ set d \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set d \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set d \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f10 by auto qed have insetd: "((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetd linsetd by auto then show "False" using insetb insetc insetd tuvprop by auto qed show "False" using cases cases_btw cases_mem cases_lt cases_gt by auto qed show "False" using asm len1 lengtone by linarith qed show "False" using lenzero lengt0 by linarith qed then show ?thesis by blast qed lemma qe_forwards: assumes "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" shows "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" (* using eq_qe_1 les_qe_1 *) proof - let ?e2 = "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" let ?f10orf11orf12 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f8orf9 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f5orf6orf7 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f2orf3orf4 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?e1 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" let ?f1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" let ?f2 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f3 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f4 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) " let ?f5 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f6 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f7 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f8 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f13 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f9 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f10 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f11 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f12 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have h1a: "(?f1 \ ?f2orf3orf4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" by auto have h2: "(?f2 \ ?f3 \ ?f4) \ ?f2orf3orf4" by auto then have h1b: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" using h1a by auto have h3: "(?f5 \ ?f6 \ ?f7) \ ?f5orf6orf7" by auto then have h1c: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" using h1b by smt have h4: "(?f8 \ ?f9 \ ?f13) \ ?f8orf9" by auto then have h1d: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f9 \ ?f13 \ ?f10orf11orf12) \ ?e2" using h1c by smt have h5: "(?f10 \ ?f11 \ ?f12) \ ?f10orf11orf12" by auto then have bigor: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) \ ?e2 " using h1d by smt then have bigor_var: "\?e2 \ \(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) " using contrapos_nn by smt have not_eq: "\(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) =(\?f1 \ \?f2 \ \?f3 \ \?f4 \ \?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12) " by linarith obtain x where x_prop: "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using assms by auto have "(\?f1 \ \?f2 \ \?f3 \ \?f4 \ \?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12) \ False" proof - assume big_not: " \ ((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)) \ \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)) \ \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)) \ \ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have c1: "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ False" proof - assume "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c') \ set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0" by auto then have "a'*x^2 + b'*x + c' = 0" using x_prop by auto then have xis: "x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a') \ x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a') " using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def by auto then have "((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ ((\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using x_prop by auto then have "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using abc_prop xis by auto then show "False" using big_not by auto qed have c2: "(\ (d, e, f) \ set a. d = 0 \ e \ 0) \ False" proof - assume "(\ (d, e, f) \ set a. d = 0 \ e \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c') \ set a \ a' = 0 \ b' \ 0" by auto then have "a'*x^2 + b'*x + c' = 0" using x_prop by auto then have "b'*x + c' = 0" using abc_prop by auto then have xis: "x = - c' / b'" using abc_prop by (smt divide_non_zero) then have "(\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" using x_prop by auto then have "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using abc_prop xis by auto then show "False" using big_not by auto qed have c3: "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0) \ False" proof - assume "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" then have equalset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" using case_prodE by auto have "\?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12" using big_not by auto then have "\(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" using equalset big_not qe_forwards_helper[of a b c d] by auto then show "False" using x_prop by auto qed have eo: "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof - have "(\ (d, e, f) \ set a. (d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0))" proof clarsimp fix d e f assume in_set: " (d, e, f) \ set a" assume dnonz: "d \ 0" have "d*x^2 + e*x + f = 0" using in_set x_prop by auto then show " 4 * d * f \ e\<^sup>2" using dnonz discriminant_negative[of d e f] unfolding discrim_def by fastforce qed then have discrim_prop: "\(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ \(\ (d, e, f) \ set a. d \ 0)" by auto have "\(\ (d, e, f) \ set a. d \ 0) \ \(\ (d, e, f) \ set a. d = 0 \ e \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof - assume ne: "\(\ (d, e, f) \ set a. d \ 0) \ \(\ (d, e, f) \ set a. d = 0 \ e \ 0)" show "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof clarsimp fix d e f assume in_set: "(d, e, f) \set a" then have xzer: "d*x^2 + e*x + f = 0" using x_prop by auto have dzer: "d = 0" using ne in_set by auto have ezer: "e = 0" using ne in_set by auto show "d = 0 \ e = 0 \ f = 0" using xzer dzer ezer by auto qed qed then show ?thesis using discrim_prop by auto qed show "False" using c1 c2 c3 eo by auto qed then have " \?e2 \ False" using bigor_var not_eq by presburger (* Takes a second *) then have " \?e2 \ False" using impI[of "\?e2" "False"] by blast then show ?thesis by auto qed subsubsection "Some Cases and Misc" lemma quadratic_linear : assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "b * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + c = 0" assumes "\x\set eq. case x of (d, e, f) \ d * ((sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a))\<^sup>2 + e * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + f = 0" assumes "(aaa, aaaa, baa) \ set eq" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa = 0" proof- have h: "-(c/b) = (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set eq. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f = 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear1: assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "(b::real) * (sqrt ((aa::real)\<^sup>2 - 4 * (a::real) * (ba::real)) - (aa::real)) / (2 * a) + (c::real) = 0" assumes " (\x\set (les::(real*real*real)list). case x of (d, e, f) \ d * ((sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a))\<^sup>2 + e * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + f < 0)" assumes "(aaa, aaaa, baa) \ set les" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa < 0" proof- have h: "-(c/b) = (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set les. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f < 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear2 : assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "b * (- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + c = 0" assumes "\x\set eq. case x of (d, e, f) \ d * ((- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a))\<^sup>2 + e * (- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + f = 0" assumes "(aaa, aaaa, baa) \ set eq" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa = 0" proof- have h: "-((c::real)/(b::real)) = (- (aa::real) -sqrt (aa\<^sup>2 - 4 * (a::real) * (ba::real))) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set eq. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f = 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear3: assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "(b::real) * (- (aa::real)- sqrt ((aa::real)\<^sup>2 - 4 * (a::real) * (ba::real)) ) / (2 * a) + (c::real) = 0" assumes "(\x\set (les::(real*real*real)list). case x of (d, e, f) \ d * ((- aa - sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a))\<^sup>2 + e * (- aa - sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + f < 0)" assumes "(aaa, aaaa, baa) \ set les" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa < 0" proof- have h: "-((c::real)/(b::real)) = (- (aa::real) -sqrt (aa\<^sup>2 - 4 * (a::real) * (ba::real))) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set les. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f < 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma h1b_helper_les: "(\((a::real), (b::real), (c::real))\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - show "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof (induct les) case Nil then show ?case by auto next case (Cons q les) have ind: " \a\set (q # les). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c < 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c < 0)" by auto have "\a\set les. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using ind by simp then have " \y. \xa\set les. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set les. case a of (a, ba, c) \ a * x^2 + ba * x + c < 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #les). case a of (a, ba, c) \ a * x^2 + ba * x + c < 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma h1b_helper_leq: "(\((a::real), (b::real), (c::real))\set leq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof - show "(\(a, b, c)\set leq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof (induct leq) case Nil then show ?case by auto next case (Cons q leq) have ind: " \a\set (q # leq). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c \ 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c \ 0)" by auto have "\a\set leq. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using ind by simp then have " \y. \xa\set leq. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set leq. case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #leq). case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma h1b_helper_neq: "(\((a::real), (b::real), (c::real))\set neq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof - show "(\(a, b, c)\set neq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof (induct neq) case Nil then show ?case by auto next case (Cons q neq) have ind: " \a\set (q # neq). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c \ 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c \ 0)" by auto have "\a\set neq. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using ind by simp then have " \y. \xa\set neq. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set neq. case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #neq). case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma min_lem: fixes r::"real" assumes a1: "(\y'>r. (\((d::real), (e::real), (f::real))\set b. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" assumes a2: "(\y'>r. (\((d::real), (e::real), (f::real))\set c. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes a3: "(\y'>r. (\((d::real), (e::real), (f::real))\set d. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - obtain y1 where y1_prop: "y1 > r \ (\(d, e, f)\set b. \x\{r<..y1}. d * x\<^sup>2 + e * x + f < 0)" using a1 by auto obtain y2 where y2_prop: "y2 > r \ (\(d, e, f)\set c. \x\{r<..y2}. d * x\<^sup>2 + e * x + f \ 0)" using a2 by auto obtain y3 where y3_prop: "y3 > r \ (\(d, e, f)\set d. \x\{r<..y3}. d * x\<^sup>2 + e * x + f \ 0)" using a3 by auto let ?y = "(min (min y1 y2) y3)" have "?y > r" using y1_prop y2_prop y3_prop by auto then have "\x. x > r \ x < ?y" using dense[of r ?y] by auto then obtain x where x_prop: "x > r \ x < ?y" by auto have bp: "(\(a, b, c)\set b. a *x\<^sup>2 + b * x + c < 0)" using x_prop y1_prop by auto have cp: "(\(a, b, c)\set c. a * x^2 + b * x + c \ 0)" using x_prop y2_prop by auto have dp: "(\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using x_prop y3_prop by auto then have "(\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using bp cp dp by auto then show ?thesis by auto qed lemma qe_infinitesimals_helper: fixes k::"real" assumes asm: "(\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - have "\(d, e, f)\set a. d = 0 \ e = 0 \ f = 0" proof clarsimp fix d e f assume "(d, e, f) \ set a" then have "\y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0" using asm by auto then obtain y' where y_prop: "y'>k \ (\x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0)" by auto then show "d = 0 \ e = 0 \ f = 0" using continuity_lem_eq0[of "k" "y'" d e f] by auto qed then have eqprop: "\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) " by auto have lesprop: "(\y'>k. (\(d, e, f)\set b. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" using les_qe_inf_helper[of b "k"] asm by blast have leqprop: "(\y'>k. (\(d, e, f)\set c. \x\{(k)<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using leq_qe_inf_helper[of c "k"] asm by blast have neqprop: "(\y'>(k). (\(d, e, f)\set d. \x\{(k)<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using neq_qe_inf_helper[of d "k"] asm by blast then have "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) " using lesprop leqprop neqprop min_lem[of "k" b c d] by auto then show ?thesis using eqprop by auto qed subsubsection "The qe\\_backwards lemma" lemma qe_backwards: assumes "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" shows " (\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - let ?e2 = "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" let ?f10orf11orf12 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f8orf9 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f5orf6orf7 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f2orf3orf4 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?e1 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" let ?f1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" let ?f2 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f3 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f4 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) " let ?f5 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f6 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f7 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f8 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f13 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f9 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f10 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f11 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f12 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have h1a: "?e2 \ (?f1 \ ?f2orf3orf4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12)" by auto have h2: "?f2orf3orf4 \ (?f2 \ ?f3 \ ?f4)" by auto then have h1b: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) " using h1a by auto have h3: "?f5orf6orf7 \ (?f5 \ ?f6 \ ?f7)" by auto then have h1c: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8orf9 \ ?f10orf11orf12) " using h1b by smt have h4: "?f8orf9 \ (?f8 \ ?f9 \ ?f13)" by auto then have h1d: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f9 \ ?f13 \ ?f10orf11orf12) " using h1c by smt have h5: "?f10orf11orf12 \ (?f10 \ ?f11 \ ?f12)" by auto then have bigor: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) " using h1d by smt have "?f1 \ ?e1" proof - assume asm: "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)" then have eqprop: "\x. \(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0" by auto have "\y. \x(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0" using asm h1b_helper_les by auto then obtain y1 where y1_prop: "\x(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0" by auto have "\y. \x(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0" using asm h1b_helper_leq by auto then obtain y2 where y2_prop: "\x(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0" by auto have "\y. \x(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0" using asm h1b_helper_neq by auto then obtain y3 where y3_prop: "\x(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0" by auto let ?y = "(min (min y1 y2) y3) - 1" have y_prop: "?y < y1 \ ?y < y2 \ ?y < y3" by auto have ap: "(\(a, b, c)\set a. a * ?y\<^sup>2 + b * ?y + c = 0)" using eqprop by auto have bp: "(\(a, b, c)\set b. a * ?y\<^sup>2 + b * ?y + c < 0)" using y_prop y1_prop by auto have cp: "(\(a, b, c)\set c. a * ?y\<^sup>2 + b * ?y + c \ 0)" using y_prop y2_prop by auto have dp: "(\(a, b, c)\set d. a * ?y\<^sup>2 + b * ?y + c \ 0)" using y_prop y3_prop by auto then have "(\(a, b, c)\set a. a * ?y\<^sup>2 + b * ?y + c = 0) \ (\(a, b, c)\set b. a * ?y\<^sup>2 + b * ?y + c < 0) \ (\(a, b, c)\set c. a * ?y\<^sup>2 + b * ?y + c \ 0) \ (\(a, b, c)\set d. a * ?y\<^sup>2 + b * ?y + c \ 0)" using ap bp cp dp by auto then show ?thesis by auto qed then have h1: "?f1 \ ?e1" by auto have "?f2 \ ?e1" proof - assume " \(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" by auto then have "\(x::real). x = -c'/b'" by auto then obtain x where x_prop: "x = - c' / b'" by auto then have "(\xa\set a. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c = 0) \ (\xa\set b. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c < 0) \ (\xa\set c. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c \ 0) \ (\xa\set d. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h2: "?f2 \ ?e1" by auto have "?f3 \ ?e1" proof - assume "\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h3: "?f3 \ ?e1" by auto have "?f4 \ ?e1" proof - assume " \(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have "?f4 \ ?e1" by auto have "?f5 \ ?e1" proof - assume asm: "\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "- c' / b'" b c d] by auto qed then have h5: "?f5 \ ?e1" by auto have "?f6 \ ?e1" proof - assume "\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h6: "?f6 \ ?e1" by auto have "?f7 \ ?e1" proof - assume "\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h7: "?f7 \ ?e1" by auto have "?f8 \ ?e1" proof - assume "\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" by auto then have "\(x::real). x = (- c' / b')" by auto then obtain x where x_prop: " x = - c' / b'" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h8: "?f8 \ ?e1" by auto have "?f9 \ ?e1" proof - assume "\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h9: "?f9 \ ?e1" by auto have "?f10 \ ?e1" proof - assume asm: "\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "- c' / b'" b c d] by auto qed then have h10: "?f10 \ ?e1" by auto have "?f11 \ ?e1" proof - assume "\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h11: "?f11 \ ?e1" by auto have "?f12 \ ?e1" proof - assume "\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h12: "?f12 \ ?e1" by auto have "?f13 \ ?e1" proof - assume " \(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h13: "?f13 \ ?e1" by auto show ?thesis using bigor h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 h11 h12 h13 using assms by (smt \\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ \x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)\) (* by force *) qed subsection "General QE lemmas" lemma qe: "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) = ((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" proof - let ?e1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" let ?e2 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" have h1: "?e1 \ ?e2" using qe_backwards by auto have h2: "?e2 \ ?e1" using qe_forwards by auto have "(?e2 \ ?e1) \ (?e1 \ ?e2) " using h1 h2 by force then have "?e2 \ ?e1" using iff_conv_conj_imp[of ?e1 ?e2] by presburger then show ?thesis by auto qed fun eq_fun :: "real \ real \ real \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ bool" where "eq_fun a' b' c' eq les leq neq = ((a' = 0 \ b' \ 0) \ (\a\set eq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\a\set les. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\a\set eq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\a\set les. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set eq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\a\set les. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" fun les_fun :: "real \ real \ real \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ bool" where "les_fun a' b' c' eq les leq neq = ((a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set eq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" lemma general_qe' : (* Direct substitution F(x) *) assumes "F = (\x. (\(a,b,c)\set eq . a*x\<^sup>2+b*x+c=0)\ (\(a,b,c)\set les. a*x\<^sup>2+b*x+c<0)\ (\(a,b,c)\set leq. a*x\<^sup>2+b*x+c\0)\ (\(a,b,c)\set neq. a*x\<^sup>2+b*x+c\0))" (* Substitution of r+\ into F *) assumes "F\ = (\r. (\(a,b,c)\set eq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c=0) \ (\(a,b,c)\set les. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c<0) \ (\(a,b,c)\set leq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c\0) \ (\(a,b,c)\set neq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c\0) )" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = ( (\(a,b,c)\set eq. \x. \y2+b*y+c=0) \ (\(a,b,c)\set les. \x. \y2+b*y+c<0) \ (\(a,b,c)\set leq. \x. \y2+b*y+c\0) \ (\(a,b,c)\set neq. \x. \y2+b*y+c\0) )" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b}::real set else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* all the root of each atom *) assumes "A = \(roots ` (set eq))" assumes "B = \(roots ` (set les))" assumes "C = \(roots ` (set leq))" assumes "D = \(roots ` (set neq))" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f\(\r\A. F r)\(\r\B. F\ r)\(\r\C. F r)\(\r\D. F\ r))" proof- { fix X have "(\(a, b, c)\set X. eq_fun a b c eq les leq neq) = (\x\F ` \(roots ` (set X)). x)" proof(induction X) case Nil then show ?case by auto next case (Cons p X) have h1: "(\x\F ` \ (roots ` set (p # X)). x) = ((\x\F ` roots p. x) \ (\x\F ` \ (roots ` set X). x))" by auto have h2 :"(case p of (a,b,c) \ eq_fun a b c eq les leq neq) = (\x\F ` roots p. x)" apply(cases p) unfolding assms apply simp by linarith show ?case unfolding h1 Cons[symmetric] using h2 by auto qed } then have eq : "\X. (\(a, b, c)\set X. eq_fun a b c eq les leq neq) = (\x\F ` \ (roots ` set X). x)" by auto { fix X have "(\(a, b, c)\set X. les_fun a b c eq les leq neq) = (\x\F\ ` \(roots ` (set X)). x)" proof(induction X) case Nil then show ?case by auto next case (Cons p X) have h1: "(\x\F\ ` \ (roots ` set (p # X)). x) = ((\x\F\ ` roots p. x) \ (\x\F\ ` \ (roots ` set X). x))" by auto have h2 :"(case p of (a,b,c) \ les_fun a b c eq les leq neq) = (\x\F\ ` roots p. x)" apply(cases p) unfolding assms apply simp by linarith show ?case unfolding h1 Cons[symmetric] using h2 by auto qed } then have les : "\X. (\(a, b, c)\set X. les_fun a b c eq les leq neq) = (\x\F\ ` \ (roots ` set X). x)" by auto have inf : "(\(a, b, c)\set eq. a = 0 \ b = 0 \ c = 0) = (\x\set eq. case x of (a, b, c) \ \x. \y2 + b * y + c = 0)" proof(induction eq) case Nil then show ?case by auto next case (Cons p eq) then show ?case proof(cases p) case (fields a b c) show ?thesis unfolding fields using infzeros[of _ a b c] Cons by auto qed qed show ?thesis using qe[of "eq" "les" "leq" "neq"] using eq[of eq] eq[of leq] les[of les] les[of neq] unfolding inf assms by auto qed lemma general_qe'' : (* Direct substitution F(x) *) assumes "S = {(=), (<), (\), (\)}" assumes "finite (X (=))" assumes "finite (X (<))" assumes "finite (X (\))" assumes "finite (X (\))" assumes "F = (\x. \rel\S. \(a,b,c)\(X rel). rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of r+\ into F *) assumes "F\ = (\r. \rel\S. \(a,b,c)\(X rel). \y>r.\x\{r<..y}. rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = (\rel\S. \(a,b,c)\(X rel). \x. \y2+b*y+c) 0)" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b}::real set else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* all the root of each atom *) assumes "A = \(roots ` ((X (=))))" assumes "B = \(roots ` ((X (<))))" assumes "C = \(roots ` ((X (\))))" assumes "D = \(roots ` ((X (\))))" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f\(\r\A. F r)\(\r\B. F\ r)\(\r\C. F r)\(\r\D. F\ r))" proof- define less where "less = (\(a::real,b::real,c::real).\(a',b',c'). a (a=a'\ (b(b=b'\cx.\y. x=y \ less x y)" have linorder: "class.linorder leq less" unfolding class.linorder_def class.order_def class.preorder_def class.order_axioms_def class.linorder_axioms_def less_def leq_def by auto show ?thesis using assms(6-8) unfolding assms(1) apply simp using general_qe'[OF _ _ _ assms(9), of F "List.linorder.sorted_list_of_set leq (X (=))" "List.linorder.sorted_list_of_set leq (X (<))" "List.linorder.sorted_list_of_set leq (X (\))" "List.linorder.sorted_list_of_set leq (X (\))" F\ F\<^sub>i\<^sub>n\<^sub>f A B C D] unfolding List.linorder.set_sorted_list_of_set[OF linorder assms(2)] List.linorder.set_sorted_list_of_set[OF linorder assms(3)] List.linorder.set_sorted_list_of_set[OF linorder assms(4)] List.linorder.set_sorted_list_of_set[OF linorder assms(5)] using assms(10-13)by auto qed theorem general_qe : (* finite sets of atoms involving = < \ and \*) assumes "R = {(=), (<), (\), (\)}" assumes "\rel\R. finite (Atoms rel)" (* Direct substitution F(x) *) assumes "F = (\x. \rel\R. \(a,b,c)\(Atoms rel). rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of r+\ into F *) assumes "F\ = (\r. \rel\R. \(a,b,c)\(Atoms rel). \y>r.\x\{r<..y}. rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = (\rel\R. \(a,b,c)\(Atoms rel). \x. \y2+b*y+c) 0)" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b} else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f \ (\r\\(roots ` (Atoms (=) \ Atoms (\))). F r) \ (\r\\(roots ` (Atoms (<) \ Atoms (\))). F\ r))" using general_qe''[OF assms(1) _ _ _ _ assms(3-6) refl refl refl refl] using assms(2) unfolding assms(1) by auto end \ No newline at end of file diff --git a/thys/Virtual_Substitution/QuadraticCase.thy b/thys/Virtual_Substitution/QuadraticCase.thy --- a/thys/Virtual_Substitution/QuadraticCase.thy +++ b/thys/Virtual_Substitution/QuadraticCase.thy @@ -1,969 +1,969 @@ subsection "Quadratic Case" theory QuadraticCase imports VSAlgos begin (*-------------------------------------------------------------------------------------------------------------*) lemma quad_part_1_eq : assumes lLength : "length L > var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Eq p) (list_update L var ((A+B*C)/D)) = aEval (Eq(quadratic_part_1 var a b d (Eq p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1 : "\i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have h2 : "((\i = 0..i = 0..i = 0..i = 0..i = 0..i = 0..i = 0.. var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Less p) (list_update L var ((A+B*C)/D)) = aEval (Less(quadratic_part_1 var a b d (Less p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1a : "((\i = 0..i = 0..i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have "((\i = 0..i = 0..i = 0..i = 0..i = 0.. var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Leq p) (list_update L var ((A+B*C)/D)) = aEval (Leq(quadratic_part_1 var a b d (Leq p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1a : "((\i = 0..i = 0..i = 0..i = 0..i = 0.. 0) =((\i = 0.. 0)" using h1a h1b nonzero by auto have h4a : "\i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have "((\i = 0.. 0)= ((\i = 0.. 0)" using h1c f_def by auto also have "...=((\i = 0.. 0)" by auto also have "...=((\i = 0.. 0)" by (metis (no_types, lifting) power_divide sum.cong) also have "...=((\i = 0..0)" using h4a by auto also have "... = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) p\0)" using sum_over_degree_insertion hdeg lLength by auto finally show ?thesis by(simp add: hdeg lLength insertion_add insertion_mult ha hb hd insertion_sum insertion_pow insertion_var) qed (*------------------------------------------------------------------------------------------------*) lemma quad_part_1_neq : assumes lLength : "length L > var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Neq p) (list_update L var ((A+B*C)/D)) = aEval (Neq(quadratic_part_1 var a b d (Neq p))) (list_update L var C)" proof - have "aEval (Eq(quadratic_part_1 var a b d (Eq p))) (list_update L var C) = aEval (Eq p) (list_update L var ((A+B*C)/D))" using quad_part_1_eq assms by blast then show ?thesis by auto qed (*------------------------------------------------------------------------------------------------*) lemma sqrt_case : assumes detGreater0 : "SQ \ 0" shows "((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" proof - have h1 : "i mod 2 = 0 \ (odd i \ (i mod 2 = 1))" by auto have h2 : "i mod 2 = 0 \ ((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" using detGreater0 apply auto by (simp add: real_sqrt_power_even) have h3 : "(odd i \ (i mod 2 = 1)) \ ((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" using detGreater0 apply auto by (smt One_nat_def add_Suc_right mult.commute nat_arith.rule0 odd_two_times_div_two_succ power.simps(2) power_mult real_sqrt_pow2) show ?thesis using h1 h2 h3 by linarith qed lemma sum_over_sqrt : assumes detGreater0 : "SQ \ 0" shows "(\i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Eq p) (list_update L var (sqrt SQ)) = aEval (Eq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Less p) (list_update L var (sqrt SQ)) = aEval (Less(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Neq p) (list_update L var (sqrt SQ)) = aEval (Neq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Leq p) (list_update L var (sqrt SQ)) = aEval (Leq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0..vars(sq::real mpoly)" shows "MPoly_Type.degree (quadratic_part_2 var sq p) var \ 1" proof - define deg where "deg = MPoly_Type.degree (p::real mpoly) var" define f where "f i = isolate_variable_sparse p var i * sq ^ (i div 2) * Const (real (i mod 2)) * Var var" for i define g where "g i = isolate_variable_sparse p var i * sq ^ (i div 2) * Const (1 - real (i mod 2))" for i have h1a : "\i. MPoly_Type.degree (isolate_variable_sparse p var i) var = 0" by (simp add: varNotIn_degree not_in_isovarspar) have h1b : "\i. MPoly_Type.degree (sq ^ (i div 2)) var = 0" by (simp add: sqfree varNotIn_degree not_in_pow) have h1c : "\i. MPoly_Type.degree (Const (real (i mod 2))) var = 0" using degree_const by blast have h1d : "MPoly_Type.degree (Var var :: real mpoly) var = 1" using degree_one by auto have h1 : "\i 1" using f_def degree_mult h1a h1b h1c h1d by (smt ExecutiblePolyProps.degree_one add.right_neutral mult.commute mult_eq_0_iff nat_le_linear not_one_le_zero) have h2a : "\i. MPoly_Type.degree (Const (1 - real (i mod 2))) var = 0" using degree_const by blast have h2 : "\ii 1" using h1 h2 by (simp add: degree_add_leq) show ?thesis using atLeastLessThanSuc_atLeastAtMost degree_sum f_def g_def h3 deg_def by auto qed (*------------------------------------------------------------------------------------------------*) lemma quad_equality_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*C)))) (list_update L var (sqrt Cv))" proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2a : "(Av + Bv * sqrt Cv = 0) = (Av = - Bv * sqrt Cv)" by auto have h2b : "(Av = - Bv * sqrt Cv) \ (Av^2 = (- Bv * sqrt Cv)^2)" by simp have h2c : "(Av^2 = (- Bv * sqrt Cv)^2) = (Av^2 = Bv^2 * (sqrt Cv)^2)" by (simp add: power_mult_distrib) have h2d : "(Av^2 = Bv^2 * (sqrt Cv)^2) = (Av^2 = Bv^2 * Cv)" by (simp add: detGreat0) have h2 : "(Av + Bv * sqrt Cv = 0) \ (Av^2 = Bv^2 * Cv)" using h2a h2b h2c h2d by blast have h3a : "(Av*Bv > 0) \ (Av + Bv * sqrt Cv \ 0)" by (smt detGreat0 mult_nonneg_nonneg real_sqrt_ge_zero zero_less_mult_iff) have h3 : "(Av + Bv * sqrt Cv = 0) \ (Av*Bv\ 0)" using h3a by linarith have h4 : "(Av * Bv \ 0 \ Av\<^sup>2 = Bv\<^sup>2 * Cv) \ (Av + Bv * sqrt Cv = 0)" apply(cases "Av>0") apply (metis detGreat0 h2a h2c h2d mult_minus_left not_le power2_eq_iff real_sqrt_lt_0_iff zero_less_mult_iff) by (smt h2a real_sqrt_abs real_sqrt_mult zero_less_mult_iff) show ?thesis apply(simp add: hA hB h1 insertion_add insertion_mult insertion_var lLength) using h2 h3 h4 by blast qed lemma quadratic_sub_eq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Eq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Eq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Eq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" using freeC quad_part_2_deg p2_def by (meson le_neq_implies_less less_one) have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have "aEval (Eq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Eq p1) (list_update L var (sqrt Cv))" using p1_def quad_part_1_eq nonzero ha hb hd lLength by blast also have h2 : "... = aEval (Eq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_eq lLength detGreater0 hc by metis also have "... = aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))" using quad_equality_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Eq p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(1) by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_less_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Less (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * C - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * C)))))) (list_update L var (sqrt Cv)) " proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2 : "\x. insertion (nth_default 0 (list_update L var x)) ((B^2)*C-A^2) = (Bv^2)*Cv-Av^2" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h3 : "Av=0 \ Bv=0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by simp have h4 : "Av<0 \ Bv\0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by (metis add.right_neutral add_mono_thms_linordered_field(5) detGreat0 less_eq_real_def mult_less_0_iff mult_zero_class.mult_zero_left mult_zero_class.mult_zero_right real_sqrt_eq_zero_cancel_iff real_sqrt_gt_0_iff) have h5a : "Av\0 \ Bv\0 \ (Av < -Bv * sqrt Cv) \ (Av\<^sup>2 < Bv\<^sup>2 * Cv)" proof - assume a1: "0 \ Av" assume a2: "Av < - Bv * sqrt Cv" assume "Bv \ 0" then have "Av < sqrt (Cv * (Bv * Bv))" using a2 by (simp add: mult.commute real_sqrt_mult) then show ?thesis using a1 by (metis (no_types) mult.commute power2_eq_square real_sqrt_less_iff real_sqrt_mult real_sqrt_pow2_iff) qed have h5b : "Av\0 \ Bv\0 \ (Av\<^sup>2 < Bv\<^sup>2 * Cv) \ (Av < -Bv * sqrt Cv)" using real_less_rsqrt real_sqrt_mult by fastforce have h5 : "Av\0 \ Bv\0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" using h5a h5b by linarith have h6 : "Av\0 \ Bv>0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by (smt detGreat0 mult_nonneg_nonneg real_sqrt_ge_zero) have h7a : "Av<0 \ Bv>0 \ (Av < -Bv * sqrt Cv) \ (Bv\<^sup>2 * Cv < Av\<^sup>2)" by (smt mult_minus_left real_sqrt_abs real_sqrt_le_mono real_sqrt_mult) have h7b : "Av<0 \ Bv>0 \ (Bv\<^sup>2 * Cv < Av\<^sup>2) \ (Av < -Bv * sqrt Cv)" by (metis abs_of_nonneg abs_real_def add.commute less_eq_real_def mult.assoc mult_minus_left power2_eq_square real_add_less_0_iff real_sqrt_less_iff real_sqrt_mult real_sqrt_mult_self) have h7 : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" using h7a h7b by linarith show ?thesis apply(simp add: hA hB h1 h2 insertion_add insertion_mult insertion_var lLength) using h3 h4 h5 h6 h7 by smt qed lemma quadratic_sub_less : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Less p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Less p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Less p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def by blast then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0(list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0(list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have h1 : "aEval (Less p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Less (quadratic_part_1 var a b d (Less p))) (list_update L var (sqrt Cv))" using quad_part_1_less assms by blast also have "... = aEval (Less p1) (list_update L var (sqrt Cv))" using p1_def by auto also have "... = aEval (Less (quadratic_part_2 var c p1)) (list_update L var (sqrt Cv))" using quad_part_2_less assms by blast also have "... = aEval (Less p2) (list_update L var (sqrt Cv))" using p2_def by auto also have "... = aEval (Less (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c)))))) (list_update L var (sqrt Cv))" using quadratic_sub_less_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Less p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(2) by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_leq_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Leq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*C-A^2))))(And (Atom(Leq B)) (Atom(Leq (A^2-B^2*C))))) (list_update L var (sqrt Cv))" proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2 : "\x. insertion (nth_default 0 (list_update L var x)) ((B^2)*C-A^2) = (Bv^2)*Cv-Av^2" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h3 : "Av=0 \ Bv=0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by simp have h4 : "Av<0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 real_sqrt_ge_zero zero_less_mult_iff) have h5 : "Av=0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 real_sqrt_ge_zero zero_less_mult_iff) have h6 : "Av\0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 mult_nonneg_nonneg mult_pos_pos real_sqrt_gt_0_iff real_sqrt_zero zero_le_power2 zero_less_mult_pos zero_less_power2) have h7a : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) \ Bv\<^sup>2 * Cv \ Av\<^sup>2" by (smt real_sqrt_abs real_sqrt_less_mono real_sqrt_mult) have h7b : "Av<0 \ Bv>0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ (Av + Bv * sqrt Cv \ 0) " by (smt real_sqrt_abs real_sqrt_less_mono real_sqrt_mult) have h7 : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" using h7a h7b by linarith have h8c : "(-Bv * sqrt Cv)^2 = Bv\<^sup>2 * Cv" by (simp add: detGreat0 power_mult_distrib) have h8a : "Av>0 \ Bv\0 \ (Av \ -Bv * sqrt Cv) \ Av\<^sup>2 \ Bv\<^sup>2 * Cv" using detGreat0 h8c power_both_sides by smt have h8b : "Av>0 \ Bv\0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv \ (Av + Bv * sqrt Cv \ 0) " using detGreat0 h8c power_both_sides by (smt mult_minus_left real_sqrt_ge_zero zero_less_mult_iff) have h8 : "Av>0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" using h8a h8b by linarith show ?thesis apply(simp add: hA hB h1 h2 insertion_add insertion_mult insertion_var lLength) using h3 h4 h5 h6 h7 h8 by smt qed lemma quadratic_sub_leq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Leq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Leq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Leq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def lLength by metis then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have "aEval (Leq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Leq p1) (list_update L var (sqrt Cv))" using quad_part_1_leq nonzero ha hb hd p1_def lLength by metis also have "... = aEval (Leq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_leq hc detGreater0 lLength by metis also have "... = aEval (Leq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have h4 : "... = eval (Or (And (Atom(Leq(A))) (Atom (Leq(B^2*c-A^2)))) (And (Atom(Leq B)) (Atom(Leq (A^2-B^2*c))))) (list_update L var (sqrt Cv))" using quadratic_sub_leq_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Leq p)) (list_update L var (sqrt Cv))" using p1_def quadratic_sub.simps(3) p2_def A_def B_def by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_neq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Neq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Neq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Neq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def lLength by metis then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have h4c : "aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))" using quad_equality_helper hc detGreater0 h4a h4b lLength by blast have h4d : "aEval (Neq (A + B * Var var)) (list_update L var (sqrt Cv)) = (\ (eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))))" using aEval.simps(1) aEval.simps(4) h4c by blast have h4e : "(\ (eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv)))) = eval (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c)))) (list_update L var (sqrt Cv))" by (metis aNeg.simps(2) aNeg.simps(3) aNeg_aEval eval.simps(1) eval.simps(4) eval.simps(5) mult_minus_left) have "aEval (Neq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Neq p1) (list_update L var (sqrt Cv))" using quad_part_1_neq nonzero ha hb hd p1_def lLength by blast also have "... = aEval (Neq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_neq hc detGreater0 lLength by metis also have "... = aEval (Neq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c)))) (list_update L var (sqrt Cv))" using h4c h4d h4e by auto also have "... = eval (quadratic_sub var a b c d (Neq p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(4) quadratic_part_1.simps(1) quadratic_part_1.simps(4) by (metis (no_types, lifting)) finally show ?thesis by blast qed (*-----------------------------------------------------------------------------------------------*) theorem free_in_quad : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn var (quadratic_sub var a b c d A)" proof(cases A) case (Less p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Less p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Less p)) = freeIn var (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c))))))" using p2_def A_def B_def p1_def quadratic_sub.simps(2) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var ((Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c)))))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Less by blast next case (Eq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Eq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Eq p)) = freeIn var (And (Atom(Leq (A*B))) (Atom (Eq (A\<^sup>2 - B\<^sup>2 * c))))" using p2_def A_def B_def p1_def quadratic_sub.simps(1) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (And (Atom(Leq (A*B))) (Atom (Eq (A\<^sup>2 - B\<^sup>2 * c))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Eq by blast next case (Leq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Leq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Leq p)) = freeIn var (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*c-A^2))))(And(Atom(Leq B))(Atom(Leq (A^2-B^2*c)))))" using p2_def A_def B_def p1_def quadratic_sub.simps(3) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*c-A^2))))(And(Atom(Leq B))(Atom(Leq (A^2-B^2*c)))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Leq by blast next case (Neq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Neq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Neq p)) = freeIn var (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c))))" using p2_def A_def B_def p1_def quadratic_sub.simps(4) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Neq by blast qed theorem quadratic_sub : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval A (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d A) (list_update L var (sqrt Cv))" proof(cases A) case (Less x1) then show ?thesis using quadratic_sub_less assms by blast next case (Eq x2) then show ?thesis using quadratic_sub_eq assms by blast next case (Leq x3) then show ?thesis using quadratic_sub_leq assms by blast next case (Neq x4) then show ?thesis using quadratic_sub_neq assms by blast qed lemma free_in_quad_fm_helper : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn (var+z) (quadratic_sub_fm_helper var a b c d F z)" proof(induction F arbitrary: z) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom x) then show ?case using free_in_quad[OF not_in_lift[OF assms(1)] not_in_lift[OF assms(2)] not_in_lift[OF assms(3)] not_in_lift[OF assms(4)], of z] by auto next case (And F1 F2) then show ?case by auto next case (Or F1 F2) then show ?case by auto next case (Neg F) then show ?case by auto next case (ExQ F) show ?case using ExQ[of "z+1"] by simp next case (AllQ F) show ?case using AllQ[of "z+1"] by simp next case (ExN x1 F) then show ?case by (metis (no_types, lifting) add.assoc freeIn.simps(13) liftmap.simps(10) quadratic_sub_fm_helper.simps) next case (AllN x1 F) then show ?case by (metis (no_types, lifting) freeIn.simps(12) group_cancel.add1 liftmap.simps(9) quadratic_sub_fm_helper.simps) qed theorem free_in_quad_fm : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn var (quadratic_sub_fm var a b c d A)" using free_in_quad_fm_helper[OF assms, of 0] by auto lemma quadratic_sub_fm_helper : assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes lLength : "length L > var+z" assumes ha : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (d::real mpoly) = (Dv :: real)" shows "eval F (list_update L (var+z) ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub_fm_helper var a b c d F z) (list_update L (var+z) (sqrt Cv))" using assms proof(induction F arbitrary: z L) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom x) define L1 where "L1 = drop z L" define L2 where "L2 = take z L" have L_def : "L = L2 @ L1" using L1_def L2_def by auto have lengthl2 : "length L2 = z" using L2_def using Atom.prems(4) by auto have "eval (Atom(Eq (a-Const Av))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (a- Const Av)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (a - Const Av)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(5) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z a) = Av" using lift_minus by blast then have a1 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z a) = Av" - unfolding L_def + unfolding L_def by (metis (no_types, lifting) Atom.prems(5) L1_def add.right_neutral add_diff_cancel_right' append_eq_append_conv append_eq_append_conv2 length_append lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) have "eval (Atom(Eq (b-Const Bv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (b- Const Bv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (b - Const Bv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(6) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z b) = Bv" using lift_minus by blast then have a2 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z b) = Bv" unfolding L_def using Atom(6) L1_def by (metis L_def add_diff_cancel_right' append.simps(1) lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) have "eval (Atom(Eq (c-Const Cv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (c- Const Cv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (c - Const Cv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(7) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = Cv" using lift_minus by blast then have a3 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z c) = Cv" unfolding L_def proof - obtain nn :: "(nat \ real) \ (nat \ real) \ real mpoly \ nat" where "\x0 x1 x2. (\v3. v3 \ vars x2 \ x1 v3 \ x0 v3) = (nn x0 x1 x2 \ vars x2 \ x1 (nn x0 x1 x2) \ x0 (nn x0 x1 x2))" by moura then have f1: "\m f fa. nn fa f m \ vars m \ f (nn fa f m) \ fa (nn fa f m) \ insertion f m = insertion fa m" by (meson insertion_irrelevant_vars) obtain rr :: real where "(\v0. insertion (nth_default 0 ((L2 @ L1)[var + z := v0])) (liftPoly 0 z c) \ Cv) = (insertion (nth_default 0 ((L2 @ L1)[var + z := rr])) (liftPoly 0 z c) \ Cv)" by blast moreover { assume "var + z \ nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)" moreover { assume "(nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))) \ ((L2 @ L1) ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = (L2 @ L1)[var + z := rr] ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" then have "nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) \ (L2 @ L1)[var + z := rr] ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) \ (L2 @ L1) ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)" by linarith then have "nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ vars (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" by (metis (no_types) append_Nil2 length_list_update nth_default_append) } ultimately have "nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ vars (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" by force } ultimately show "\r. insertion (nth_default 0 ((L2 @ L1)[var + z := r])) (liftPoly 0 z c) = Cv" using f1 by (metis (full_types) Atom.prems(3) \insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = Cv\ not_in_lift) qed have "eval (Atom(Eq (d-Const Dv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (d- Const Dv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (d - Const Dv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(8) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z d) = Dv" using lift_minus by blast then have a4 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z d) = Dv" unfolding L_def by (metis Atom(8) L1_def L_def add_diff_cancel_right' append.simps(1) lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) then show ?case apply(simp) using quadratic_sub[OF Atom(4) Atom(1) Atom(2) not_in_lift[OF Atom(3)], of "(liftPoly 0 z a)" Av "(liftPoly 0 z b)" Bv "(liftPoly 0 z d)" x , OF a1 a2 a3 a4] . next case (And F1 F2) then show ?case by auto next case (Or F1 F2) then show ?case by auto next case (Neg F) then show ?case by auto next case (ExQ F) have lengthG : "var + (z + 1) < length (x#L)" for x using ExQ(5) by auto have forall : "\x. insertion (nth_default 0 ((drop z L)[var := x])) a = Av \ \x. insertion (nth_default 0 ((drop (z + 1) (x1 # L))[var := x])) a = Av" for x1 a Av by auto have l : "x # L[var + z := v] = ((x#L)[var+(z+1):=v])" for x v by auto have "eval (ExQ F) (L[var + z := (Av + Bv * sqrt Cv) / Dv]) = (\x. eval F (x # L[var + z := (Av + Bv * sqrt Cv) / Dv]))" by(simp) also have "... = (\x. eval (liftmap (\x. quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d)) F (z + 1)) (x # L[var + z := sqrt Cv]))" apply(rule ex_cong1) unfolding l using ExQ(1)[OF ExQ(2) ExQ(3) ExQ(4) lengthG forall[OF ExQ(6)] forall[OF ExQ(7)] forall[OF ExQ(8)] forall[OF ExQ(9)]] unfolding quadratic_sub_fm_helper.simps liftmap.simps by simp also have "... = eval (quadratic_sub_fm_helper var a b c d (ExQ F) z) (L[var + z := sqrt Cv])" unfolding quadratic_sub_fm_helper.simps liftmap.simps eval.simps by auto finally show ?case by simp next case (AllQ F) have lengthG : "var + (z + 1) < length (x#L)" for x using AllQ(5) by auto have forall : "\x. insertion (nth_default 0 ((drop z L)[var := x])) a = Av \ \x. insertion (nth_default 0 ((drop (z + 1) (x1 # L))[var := x])) a = Av" for x1 a Av by auto have l : "x # L[var + z := v] = ((x#L)[var+(z+1):=v])" for x v by auto have "eval (AllQ F) (L[var + z := (Av + Bv * sqrt Cv) / Dv]) = (\x. eval F (x # L[var + z := (Av + Bv * sqrt Cv) / Dv]))" by(simp) also have "... = (\x. eval (liftmap (\x. quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d)) F (z + 1)) (x # L[var + z := sqrt Cv]))" apply(rule all_cong1) unfolding l using AllQ(1)[OF AllQ(2) AllQ(3) AllQ(4) lengthG forall[OF AllQ(6)] forall[OF AllQ(7)] forall[OF AllQ(8)] forall[OF AllQ(9)]] unfolding quadratic_sub_fm_helper.simps liftmap.simps by simp also have "... = eval (quadratic_sub_fm_helper var a b c d (AllQ F) z) (L[var + z := sqrt Cv])" unfolding quadratic_sub_fm_helper.simps liftmap.simps eval.simps by auto finally show ?case by simp next case (ExN x1 F) have list : "\l. length l=x1 \ ((drop (z + x1) l @ drop (z + x1 - length l) L)) = ((drop z L))" by auto have map : "\ z L. eval (liftmap (\x A. (quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d) A)) F (z + x1)) L = eval (liftmap (\x A. (quadratic_sub (var + x1 + x) (liftPoly 0 (x+x1) a) (liftPoly 0 (x+x1) b) (liftPoly 0 (x+x1) c) (liftPoly 0 (x+x1) d) A)) F z) L" apply(induction F) apply(simp_all add:add.commute add.left_commute) apply force apply force by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1))+ show ?case apply simp apply(rule ex_cong1) subgoal for l using map[of z] list[of l] ExN(1)[OF ExN(2-4), of "z+x1" "l@L"] ExN(5-9) list_update_append apply auto by (simp add: list_update_append) + done next case (AllN x1 F) have list : "\l. length l=x1 \ ((drop (z + x1) l @ drop (z + x1 - length l) L)) = ((drop z L))" by auto have map : "\ z L. eval (liftmap (\x A. (quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d) A)) F (z + x1)) L = eval (liftmap (\x A. (quadratic_sub (var + x1 + x) (liftPoly 0 (x+x1) a) (liftPoly 0 (x+x1) b) (liftPoly 0 (x+x1) c) (liftPoly 0 (x+x1) d) A)) F z) L" apply(induction F) apply(simp_all add:add.commute add.left_commute) apply force apply force by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1))+ show ?case apply simp apply(rule all_cong1) subgoal for l using map[of z] list[of l] AllN(1)[OF AllN(2-4), of "z+x1" "l@L"] AllN(5-9) apply auto by (simp add: list_update_append) + done qed theorem quadratic_sub_fm : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "eval F (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub_fm var a b c d F) (list_update L var (sqrt Cv))" unfolding quadratic_sub_fm.simps using quadratic_sub_fm_helper[OF assms(2) assms(3) assms(4), of 0 L a Av b Bv d F] assms(1) assms(5) assms(6) assms(7) assms(8) by (simp add: lLength) end \ No newline at end of file