(* transcendence-20241221.ml D. J. Bernstein This file is hereby placed into the public domain. SPDX-License-Identifier: LicenseRef-PD-hp OR CC0-1.0 OR 0BSD OR MIT-0 OR MIT ===== To re-verify under Debian (about an hour depending on CPU): sudo apt install opam xdot -y time opam init github git+https://github.com/ocaml/opam-repository.git --disable-sandboxing -a time opam switch create 5.2.0 time opam install 'hol_light=3.0.0' -y eval `opam env` time LINE_EDITOR=env hol.sh \ < transcendence-20241221.ml \ > transcendence-20241221.out & ===== Highlighted theorems: e_is_irrational (1744 Euler): ~(cexp (Cx (&1)) IN QinC) e_is_transcendental (1873 Hermite): ~algebraic_number (cexp (Cx (&1))) pi_is_transcendental (1882 Lindemann): ~algebraic_number (Cx pi) transcendental_if_exp_nonzero_algebraic (1882 Lindemann): algebraic_number a /\ algebraic_number (cexp a) ==> a = Cx (&0) zero_sum_algebraic_exp_algebraic (1882 Lindemann): forall S B. FINITE S /\ S SUBSET algebraic_number /\ (forall s. s IN S ==> algebraic_number (B s)) /\ ring_sum complex_ring S (\s. B s * cexp s) = Cx (&0) ==> (forall s. s IN S ==> B s = Cx (&0)) Traditional Greek presentation of the last theorem: if sum_α β_α exp α = 0, where each α is algebraic and each β_α is algebraic, then each β_α = 0. This easily implies the other highlighted theorems. Some of the proofs below act as sanity checks on the definitions (e.g., algebraic_number_QinC; algebraic_number_ii). Another sanity check would be to prove that, for reals, the definition of algebraic numbers here is equivalent to the definition in 100/liouville.ml, which internally uses real polys defined via Library/poly.html with polys constrained to have integer coefficients. ===== Historical notes: Various sources attribute the last theorem to 1885 Weierstrass as the "Lindemann--Weierstrass theorem", or sometimes the "Hermite--Lindemann--Weierstrass theorem". My reading of the Lindemann and Weierstrass papers is that Lindemann had this theorem and Weierstrass was, with Lindemann's agreement, just giving a more detailed exposition of the result, not assuming so much reader familiarity with 1873 Hermite. Baker's book refers to the second result as "Lindemann's theorem", saying that Lindemann sketched the proof and that Weierstrass gave a rigorous proof. I don't know what's supposedly missing from the sketch. Some sources instead claim that this theorem is a simplification due to Baker of the original "Lindemann--Weierstrass theorem", namely the statement that if α_1,... are algebraic numbers linearly independent over \Q then exp(α_1),... are algebraically independent over \Q. No evident signs of anyone asking the obvious questions: "Um, are we sure that people didn't find this theorem first, with the algebraic-independence version as a corollary later? Maybe we should check what Lindemann actually wrote?" ===== Proof notes: The high-level structure of the proofs here follows Lindemann and most of the literature in symmetrizing β to reduce to the case that β_α is rational, and symmetrizing α to reduce further to the core case that β_α is rational and invariant under conjugation of α. The first major proof stage is the core proof, producing transcendence_all_0_QinC here: if sum_α β_α exp α = 0, where each α is algebraic, each β_α is rational, and β_α is invariant under conjugation of α, then each β_α = 0. The core proof here follows 1990 Beukers--Bezivin--Robba. Most of the literature instead follows 1893 Hilbert. The core proof here uses polynomial reversals. Laurent series would avoid this but would need more infrastructure. For eventually optimizing an explicit form of the transcendence proof (proving a lower bound on weighted sum of exp of algebraic numbers), one wants an explicit lower bound on factorials, such as factorial_lower_bound here (which says (n/e)^n <= n!). More work (as in, e.g., 100/stirling.ml) produces only marginally better bounds. For the current target of just saying sum of exp of algebraic is nonzero, it suffices to say that factorials eventually beat any exponential. This doesn't need the explicit bound, but isn't much simpler, so the proof here uses the explicit bound. The proofs of transcendence of e and pi here are immediately after the first stage. This is common in the literature: it is a warmup for parts of the second stage, and it gives proofs of those facts that are shorter overall than deriving the facts after the second stage (or the third). The pi proof here uses explicit resolvent identities, proven via expformal and Newton's identities, whereas all of the literature that I've found instead uses the fundamental theorem of symmetric functions. (Some sources credit Newton's identities to Girard. I don't see the rationale for this credit. 1629 Girard had _formulas_ for first few coeffs in terms of power sums. 1707 Newton had a bilinear identity relating coeffs to power sums, which can be used as a _recurrence_ for all coeffs given power sums. One way to prove Girard's formulas is to iterate Newton's identities; this doesn't mean that Girard published Newton's identities.) The second major proof stage uses α symmetrization to obtain zero_sum_QinC_exp_algebraic here: if sum_α β_α exp α = 0, where each α is algebraic and each β_α is rational, then each β_α = 0. Consider the product prod_q (sum_α β_α exp(q(α))) where q runs through all permutations of a finite Galois-stable set containing all α. This product is 0: consider the identity permutation. Expanding the product produces a sum that fits the core case. As in the literature, the proof here uses the fundamental theorem of symmetric functions (which is also proven here). Applying the core theorem then shows that sum_D ((prod β_α^(D_α)) #{f: sum_q q(f(q)) = z}) = 0 for each z, where f runs through functions with distribution D. There is then a denouement concluding that β_α = 0 for each α. The typical denouement in the literature identifies the maximum z with nonzero prod β_α^(D_α), where maximum refers to lex order on \C. A different denouement appears in 1952 Steinberg--Redheffer "Analytic proof of the Lindemann theorem". The denouement here follows 1952 Steinberg--Redheffer but works with expformal to skip a detour through convergence. The point is that, working backwards, one has prod_q (sum_α β_α E(q(α))) = 0 and thus sum_α β_α E(q(α)) = 0 for some q, for any function E mapping addition to multiplication in some field. The 1952 proof considers E(α) = z^α for various complex z; the proof here takes E = expformal. The third major proof stage symmetrizes β, producing zero_sum_algebraic_exp_algebraic here: if sum_α β_α exp α = 0, where each α is algebraic and each β_α is algebraic, then each β_α = 0. For β symmetrization, analogously to α symmetrization, the proof here follows the literature in using the fundamental theorem of symmetric functions, but again uses a simplified version of the 1952 denouement. As in, e.g., 2017 Bernard, the β symmetrization here works with one polynomial having all β as roots, rather than separately symmetrizing each β. The shortest-for-papers symmetrization approach (used in, e.g., 1990 Beukers--Bezivin--Robba) picks a finite-degree Galois extension of \Q containing all α and all β and then applies all automorphisms of the extension to symmetrize all α and β simultaneously, but this relies on more background. The proof here of transcendental_if_exp_nonzero_algebraic is immediately after the second stage. This is common in the literature. Waiting until after the third stage would replace the minpoly over \Q in this proof with a linear poly over \Qbar, giving a slightly shorter last step of the proof but again longer proofs overall. Computationally, the β symmetrization, if optimized, would give back the same minpoly over \Q. ===== Prerequisite notes: The big import of Multivariate/cauchy.ml is for the fundamental theorem of algebra (FTA). This also brings in various complex-number basics, the complex exponential function (cexp), and a bound on the tail of the exp series (TAYLOR_CEXP). Most proofs of Lindemann's theorem have bigger analytic parts, for example to put bounds on integrals. The proof here focuses more on identities of formal power series. *) (* ===== imports *) needs "Library/products.ml";; needs "Library/ringtheory.ml";; needs "Multivariate/cauchy.ml";; prioritize_real();; (* to ensure portability *) (* ===== tactics *) let rw = REWRITE_TAC;; let once_rw = ONCE_REWRITE_TAC;; let qed = ASM_MESON_TAC;; let simp = ASM_SIMP_TAC;; let simporqed = (simp[] THEN qed[]) ORELSE (simp[]);; let pass = ALL_TAC;; let case = ASM_CASES_TAC;; let conjunction = CONJ_TAC;; let splitiff = EQ_TAC;; let know = ASSUME;; let sufficesby_bare x = MATCH_MP_TAC x;; let sufficesby x = sufficesby_bare(REWRITE_RULE[GSYM IMP_CONJ] x);; let intro = REPEAT STRIP_TAC;; let intro_genonly = REPEAT GEN_TAC;; let intro_gendisch = REPEAT GEN_TAC THEN REPEAT DISCH_TAC;; let subgoal t = SUBGOAL_THEN t ASSUME_TAC;; let witness = EXISTS_TAC;; let proven_if t why = case t THENL [qed why; pass];; let labelhavetac L t tac = SUBGOAL_THEN t MP_TAC THENL [tac; pass] THEN DISCH_THEN(fun th -> LABEL_TAC L th);; let havetac = labelhavetac "";; let labelhave L t why = labelhavetac L t(qed why);; let have t why = havetac t(qed why);; let have_rw t why = have t why THEN once_rw[know t];; let labelspecialize L v th = LABEL_TAC L (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL v th)));; let specialize = labelspecialize "";; let specialize_assuming v th = ASSUME_TAC(REWRITE_RULE [IMP_CONJ] (ISPECL v th));; let specialize_assuming_nosplit v th = ASSUME_TAC(UNDISCH_ALL(ISPECL v th));; let specialize_forward v th = ASSUME_TAC(UNDISCH_ALL (fst (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL v th))))));; let specialize_reverse v th = ASSUME_TAC(UNDISCH_ALL (snd (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL v th))))));; let recall = specialize[];; let complex_field_fact t = recall(COMPLEX_FIELD t);; let real_field_fact t = recall(REAL_FIELD t);; let real_linear_fact t = recall(REAL_ARITH t);; let int_linear_fact t = recall(INT_ARITH t);; let num_linear_fact t = recall(ARITH_RULE t);; let set_fact_using t why = havetac t(SET_TAC why);; let set_fact t = recall(SET_RULE t);; let set_fact_assuming t = specialize_assuming[](SET_RULE t);; let def n d = X_CHOOSE_TAC n(MESON [] (mk_exists (n, mk_eq (n, d))));; let removelabeled L = REMOVE_THEN L (fun th -> ALL_TAC);; let choose n p why = labelhave "choosetmp" (mk_exists (n, p)) why THEN X_CHOOSE_TAC n(UNDISCH (TAUT (mk_imp (mk_exists (n, p), mk_exists (n, p))))) THEN removelabeled "choosetmp";; let choose2 n1 n2 p why = choose n1 (mk_exists (n2, p)) why THEN choose n2 p [] ;; (* ===== sets *) let subset_y = prove(` !S U:X->bool. S SUBSET U <=> (!y. y IN S ==> y IN U) `, qed[SUBSET] );; let extension_z = prove(` !S:X->bool T. S = T <=> (!z. z IN S <=> z IN T) `, qed[EXTENSION] );; let insert_delete_nonmember = prove(` !(x:X) S. ~(x IN S) ==> (x INSERT S) DELETE x = S `, SET_TAC[] );; let surjective_finite = prove(` !f p:T->bool. (!t:T. ?s:S. f(s) = t) ==> FINITE {s | p(f(s))} ==> FINITE {t | p(t)} `, intro THEN specialize[`f:S->T`;`{s:S | p(f(s):T)}`]FINITE_IMAGE_EXPAND THEN subgoal `{t:T | p t} SUBSET {t | ?s:S. s IN {x | p(f(x))} /\ t = f s}` THENL [ rw[SUBSET;IN_ELIM_THM] THEN qed[] ; qed[FINITE_SUBSET] ] );; let is_inters = prove(` !(x:X->bool) u. x IN u /\ (!s. s IN u ==> x SUBSET s) ==> x = INTERS u `, SET_TAC[] );; let card_empty = prove(` CARD ({}:X->bool) = 0 `, qed[CARD_CLAUSES] );; (* just to avoid SUC from CARD_CLAUSES *) let card_insert = prove(` !x:X S. FINITE S ==> CARD(x INSERT S) = (if x IN S then CARD S else CARD S + 1) `, simp[CARD_CLAUSES] THEN ARITH_TAC );; let subset_full_card = prove(` !S:X->bool U. FINITE S ==> ( (U SUBSET S /\ CARD U = CARD S) <=> U = S ) `, intro THEN qed[SUBSET_REFL;CARD_SUBSET_EQ] );; let finite_subsets_card = prove(` !S:X->bool n. FINITE S ==> FINITE {U | U SUBSET S /\ CARD U = n} `, intro THEN set_fact `{U:X->bool | U SUBSET S /\ CARD U = n} SUBSET {U:X->bool | U SUBSET S}` THEN qed[FINITE_POWERSET;FINITE_SUBSET] );; let subsets_full_card = prove(` !S:X->bool. FINITE S ==> {U | U SUBSET S /\ CARD U = CARD S} = {S} `, rw[EXTENSION;IN_ELIM_THM] THEN rw[IN_SING] THEN qed[subset_full_card] );; let subsets_card_0 = prove(` !S:X->bool. FINITE S ==> {U | U SUBSET S /\ CARD U = 0} = {{}} `, rw[EXTENSION;IN_ELIM_THM] THEN rw[IN_SING] THEN intro THEN qed[CARD_EQ_0;FINITE_SUBSET;EMPTY_SUBSET;card_empty] );; let subsets_full_card_empty = prove(` {U:X->bool | U SUBSET {} /\ CARD U = 0} = {{}} `, have `FINITE({}:X->bool)` [FINITE_EMPTY] THEN specialize[`{}:X->bool`]subsets_card_0 THEN qed[] );; let subsets_card_toobig = prove(` !S:X->bool n. FINITE S ==> ~(n <= CARD S) ==> {U | U SUBSET S /\ CARD U = n} = {} `, rw[EXTENSION;IN_ELIM_THM;EMPTY] THEN intro THEN qed[CARD_SUBSET] );; let powerset_insert_disjoint = prove(` !S:X->bool t. ~(t IN S) ==> DISJOINT {A | A SUBSET S} (IMAGE (\A. t INSERT A) {A | A SUBSET S}) `, intro THEN rw[DISJOINT] THEN once_rw[EXTENSION] THEN rw[INTER;IMAGE;IN_ELIM_THM;EMPTY] THEN intro THEN qed[IN_INSERT;SUBSET] );; let image_card_powerset = prove(` !S:X->bool. FINITE S ==> IMAGE CARD {A | A SUBSET S} = (0..CARD S) `, intro THEN rw[EXTENSION;IN_ELIM_THM;IN_IMAGE;IN_NUMSEG_0] THEN intro THEN splitiff THENL [ qed[CARD_SUBSET] ; intro THEN have `{} SUBSET S:X->bool` [EMPTY_SUBSET] THEN have `FINITE({}:X->bool)` [FINITE_EMPTY] THEN have `CARD({}:X->bool) <= x` [card_empty;ARITH_RULE `0 <= n`] THEN have `FINITE S ==> x <= CARD(S:X->bool)` [] THEN specialize[ `x:num`; `{}:X->bool`; `S:X->bool` ]CHOOSE_SUBSET_BETWEEN THEN choose `t:X->bool` `{}:X->bool SUBSET t /\ t SUBSET S /\ t HAS_SIZE x` [] THEN witness `t:X->bool` THEN qed[HAS_SIZE] ] );; let le_lt_numseg = prove(` !a b. {i:num | a <= i /\ i < b} = if a < b then (a..b-1) else {} `, rw[numseg;EXTENSION;EMPTY;IN_ELIM_THM;COND_RAND] THEN ARITH_TAC );; let finite_le_lt = prove(` !a b. FINITE {i:num | a <= i /\ i < b} `, rw[le_lt_numseg;COND_RAND] THEN rw[FINITE_EMPTY;FINITE_NUMSEG] THEN qed[] );; (* warning: this is normal notation only if a<=b *) let card_le_lt = prove(` !a b. CARD {i:num | a <= i /\ i < b} = b-a `, rw[le_lt_numseg;COND_RAND] THEN rw[CARD_CLAUSES;CARD_NUMSEG] THEN rw[GSYM COND_RAND] THEN ARITH_TAC );; let insert_empty = prove(` !x:X. x INSERT {} = {x} `, SET_TAC[] );; let max_finite = prove(` !S:num->bool. FINITE S ==> ( S = {} \/ (?m. m IN S /\ (!n. m < n ==> ~(n IN S))) ) `, sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ qed[] ; DISJ2_TAC THEN witness `x:num` THEN simp[insert_empty;IN_SING] THEN ARITH_TAC ; DISJ2_TAC THEN rw[IN_INSERT] THEN case `x <= m:num` THENL [ witness `m:num` THEN intro THENL [ qed[] ; ASM_ARITH_TAC ; qed[] ] ; witness `x:num` THEN intro THENL [ qed[] ; ASM_ARITH_TAC ; num_linear_fact `~(x <= m:num) ==> x < n ==> m < n` THEN qed[] ] ] ] );; let image_surj = prove(` !f:X->Y A B. SURJ f A B ==> IMAGE f A = B `, rw[SURJ;IMAGE;EXTENSION;IN_ELIM_THM] THEN qed[] );; let elim_image_subset_u = prove(` !A B f:X->Y. IMAGE f A SUBSET B <=> (!u. u IN A ==> f u IN B) `, SET_TAC[] );; let elim_image_subset_v = prove(` !A B f:X->Y. IMAGE f A SUBSET B <=> (!v. v IN A ==> f v IN B) `, SET_TAC[] );; let select_foreach = prove(` !S P. (!s:X. s IN S ==> ?t:Y. P s t) ==> (!s:X. s IN S ==> P s ((\s. @t. P s t) s)) `, qed[] );; (* XXX: there must be a better way to do renaming *) let o_def_s = prove(` !f:Y->Z g:X->Y. f o g = (\s. f (g s)) `, qed[o_DEF] );; let in_image_cd = prove(` !d s f:X->Y. d IN IMAGE f s <=> (?c. d = f c /\ c IN s) `, qed[IN_IMAGE] );; let in_image_vw = prove(` !w s f:X->Y. w IN IMAGE f s <=> (?v. w = f v /\ v IN s) `, qed[IN_IMAGE] );; (* ===== pairs *) let lambda_pair_ab = prove(` !f:A#B->C. (\ab. f ab) = (\(a,b). f (a,b)) `, qed[LAMBDA_PAIR_THM] );; let injective_pair_rewrite = prove(` !f:A#B->L. (!(a,b) (c,d). f (a,b) = f (c,d) ==> (a,b) = (c,d)) ==> !x y. f x = f y ==> x = y `, rw[FORALL_PAIRED_THM] THEN qed[PAIR_SURJECTIVE] );; let image_pair = prove(` !P f:A->B->C. IMAGE (\(a,b). f a b) {a,b | P a b} = {f a b | P a b} `, rw[EXTENSION;IN_ELIM_THM;IN_IMAGE] THEN intro THEN splitiff THENL [ STRIP_TAC THEN witness `a:A` THEN witness `b:B` THEN have `(x:C) = (\((a:A),(b:B)). f a b) (a,b)` [] THEN simp[] ; STRIP_TAC THEN witness `a:A,b:B` THEN conjunction THENL [ simp[] ; witness `a:A` THEN witness `b:B` THEN simp[] ] ] );; (* ===== functions *) let functions = new_definition ` functions (A:X->bool) (B:Y->bool) = {f | IMAGE f A SUBSET B /\ (!z. ~(z IN A) ==> f z = ARB)} `;; let in_functions = prove(` !A:X->bool B:Y->bool f. f IN functions A B <=> IMAGE f A SUBSET B /\ (!z. ~(z IN A) ==> f z = ARB) `, SET_TAC[functions] );; let fun_eq_thm_e = prove(` !f g:X->Y. f = g <=> (!e. f e = g e) `, qed[FUN_EQ_THM] );; let fun_eq_thm_v = prove(` !f g:X->Y. f = g <=> (!v. f v = g v) `, qed[FUN_EQ_THM] );; let image_functions_subset = prove(` !A:X->bool B:Y->bool f. f IN functions A B ==> IMAGE f A SUBSET B `, SET_TAC[functions] );; (* XXX: factor most image_functions_subset applications via this *) let functions_to = prove(` !A:X->bool B:Y->bool f x. f IN functions A B ==> x IN A ==> f x IN B `, SET_TAC[image_functions_subset] );; let finite_functions = prove(` !A:X->bool B:Y->bool. FINITE A ==> FINITE B ==> FINITE (functions A B) `, rw[functions] THEN intro THEN subgoal `{f | IMAGE f A SUBSET B /\ (!x:X. ~(x IN A) ==> f x = ARB:Y)} = {f | IMAGE f A SUBSET B /\ {x | ~(f x = ARB)} SUBSET A}` THENL [ SET_TAC[] ; pass ] THEN simp[] THEN specialize[ `A:X->bool`; `B:Y->bool`; `\a:X. ARB:Y` ]FINITE_RESTRICTED_FUNSPACE THEN qed[] );; let functions_empty = prove(` !B:Y->bool. functions {} B = {\x:X. ARB} `, rw[functions] THEN rw[EXTENSION;IN_SING;IN_ELIM_THM;IMAGE_CLAUSES;EMPTY_SUBSET;SUBSET;EMPTY] THEN qed[FUN_EQ_THM] );; let functions_insert = prove(` !A:X->bool B:Y->bool i:X. ~(i IN A) ==> functions (i INSERT A) B = IMAGE (\((b:Y),(f:X->Y)) (a:X). if a = i then b else f a) (B CROSS (functions A B)) `, rw[functions;EXTENSION;IN_ELIM_THM;in_image_vw] THEN intro THEN splitiff THENL [ intro THEN witness `(x(i:X):Y),(\a:X. if a = i then ARB:Y else x a)` THEN rw[BETA_THM;IN_CROSS] THEN intro THENL [ qed[FUN_EQ_THM] ; ASM SET_TAC[] ; rw[IN_ELIM_THM] THEN ASM SET_TAC[] ] ; rw[IMAGE_CLAUSES;INSERT_SUBSET;EXISTS_PAIR_THM;IN_CROSS;IN_ELIM_THM] THEN ASM SET_TAC[] ] );; let functions_insert_injective = prove(` !A:X->bool B:Y->bool i:X v w. ~(i IN A) ==> v IN B CROSS functions A B ==> w IN B CROSS functions A B ==> (\((b:Y),(f:X->Y)) (a:X). if a = i then b else f a) v = (\((b:Y),(f:X->Y)) (a:X). if a = i then b else f a) w ==> v = w `, rw[FORALL_PAIR_THM;IN_CROSS;PAIR_EQ;FUN_EQ_THM] THEN rw[functions;IN_ELIM_THM] THEN qed[] );; (* ===== permutations *) let perm = new_definition ` perm (S:X->bool) (f:X->X) <=> f permutes S `;; let perm_in_permutes = prove(` !S:X->bool f. perm S f <=> f IN {p | p permutes S} `, rw[IN_ELIM_THM;perm] );; let perm_set_permutes = prove(` !S:X->bool. perm S = {p | p permutes S} `, rw[EXTENSION;IN;perm_in_permutes] );; let card_perm = prove(` !S:X->bool. FINITE S ==> CARD(perm S) = FACT(CARD S) `, rw[perm_set_permutes;CARD_PERMUTATIONS] );; let finite_perm = prove(` !S:X->bool. FINITE S ==> FINITE(perm S) `, rw[perm_set_permutes;FINITE_PERMUTATIONS] );; let permutes_o_inverse_refl_o = prove(` !A:X->bool f:X->X g:Y->X. f permutes A ==> f o inverse f o g = g `, intro THEN have `f:X->X o inverse f = I` [PERMUTES_INVERSES_o] THEN rw[o_ASSOC] THEN qed[I_O_ID] );; let inverse_permutes_o_refl_o = prove(` !A:X->bool f:X->X g:Y->X. f permutes A ==> inverse f o f o g = g `, intro THEN have `inverse f o f:X->X = I` [PERMUTES_INVERSES_o] THEN rw[o_ASSOC] THEN qed[I_O_ID] );; let image_permutes_o_perm = prove(` !A:X->bool f:X->X. f permutes A ==> IMAGE (\i. f o i) (perm A) = perm A `, rw[EXTENSION;in_image_vw;IN_ELIM_THM;perm_set_permutes] THEN intro THEN splitiff THENL [ qed[PERMUTES_COMPOSE] ; intro THEN witness `inverse f:X->X o x:X->X` THEN qed[permutes_o_inverse_refl_o;PERMUTES_INVERSE;PERMUTES_COMPOSE] ] );; let image_permutes_o_functions_perm = prove(` !A:X->bool B:Y->bool f:X->X. f permutes A ==> IMAGE (\g i. g(f o i)) (functions (perm A) B) = functions (perm A) B `, rw[EXTENSION;functions;in_image_vw;IN_ELIM_THM] THEN rw[elim_image_subset_u] THEN intro THEN have `inverse f:X->X permutes A` [PERMUTES_INVERSE] THEN have `!a:X->X. inverse f o f:X->X o a = a` [inverse_permutes_o_refl_o] THEN have `!a:X->X. f o inverse f:X->X o a = a` [permutes_o_inverse_refl_o] THEN splitiff THENL [ intro THENL [ have `u:X->X permutes A` [IN;perm] THEN have `f:X->X o u:X->X permutes A` [PERMUTES_COMPOSE] THEN qed[IN;perm] ; case `f:X->X o z:X->X permutes A` THENL [ have `z:X->X permutes A` [PERMUTES_COMPOSE] THEN qed[IN;perm] ; pass ] THEN qed[IN;perm] ] ; intro THEN def `y:(X->X)->Y` `\i:X->X. x (inverse f:X->X o i):Y` THEN witness `y:(X->X)->Y` THEN intro THENL [ simp[fun_eq_thm_e] ; qed[IN;perm;PERMUTES_COMPOSE] ; case `inverse f:X->X o z:X->X permutes A` THENL [ have `z:X->X permutes A` [PERMUTES_COMPOSE] THEN qed[IN;perm] ; pass ] THEN qed[IN;perm] ] ] );; let injective_permutes_arbo_functions = prove(` !A:X->bool B:Y->bool p:Y->Y f g. p permutes B ==> f IN functions A B ==> g IN functions A B ==> (\a. if a IN A then p (f a) else ARB) = (\a. if a IN A then p (g a) else ARB) ==> f = g `, rw[in_functions;elim_image_subset_u] THEN rw[fun_eq_thm_e] THEN qed[permutes] );; let image_permutes_arbo_functions = prove(` !A:X->bool B:Y->bool p:Y->Y. p permutes B ==> IMAGE (\ab a. if a IN A then p(ab a) else ARB) (functions A B) = functions A B `, rw[EXTENSION;IN_IMAGE;in_functions;elim_image_subset_u] THEN rw[fun_eq_thm_e] THEN intro THEN have `inverse (p:Y->Y) permutes B` [PERMUTES_INVERSE] THEN splitiff THENL [ intro THENL [ simp[] THEN qed[PERMUTES_IN_IMAGE] ; qed[permutes] ] ; intro THEN witness `\a:X. if a IN A then (inverse p)(x a:Y) else ARB:Y` THEN rw[BETA_THM] THEN intro THENL [ simp[] THEN qed[PERMUTES_INVERSES;permutes] ; qed[PERMUTES_IN_IMAGE] ; qed[] ] ] );; let injective_permutes_arbo_functions_functions = prove(` !A:X->bool B:Y->bool C:Z->bool p:Y->Y f g. p permutes B ==> f IN functions (functions A B) C ==> g IN functions (functions A B) C ==> (\ab. if ab IN functions A B then f (\a. if a IN A then p (ab a) else ARB) else ARB) = (\ab. if ab IN functions A B then g (\a. if a IN A then p (ab a) else ARB) else ARB) ==> f = g `, rw[fun_eq_thm_e] THEN intro THEN case `e:X->Y IN functions A B` THENL [ have `inverse (p:Y->Y) permutes B` [PERMUTES_INVERSE] THEN subgoal `(\a:X. if a IN A then (inverse p)(e a:Y):Y else ARB) IN functions A B` THENL [ rw[functions;IN_ELIM_THM;SUBSET;in_image_vw] THEN qed[functions_to;PERMUTES_IN_IMAGE] ; pass ] THEN subgoal `(\a:X. if a IN A then p(if a IN A then (inverse p)(e a:Y):Y else ARB) else ARB) = e` THENL [ rw[FUN_EQ_THM] THEN qed[in_functions;PERMUTES_INVERSES] ; pass ] THEN specialize[ `\a:X. if a IN A then (inverse p)(e a:Y):Y else ARB` ](know `!e:X->Y. (if e IN functions A B then f (\a. if a IN A then p (e a) else ARB:Y) else ARB) = (if e IN functions A B then g (\a. if a IN A then p (e a:Y) else ARB) else ARB):Z`) THEN qed[] ; pass ] THEN qed[in_functions] );; let image_permutes_arbo_functions_functions = prove(` !A:X->bool B:Y->bool C:Z->bool p:Y->Y. p permutes B ==> IMAGE (\bc ab. if ab IN functions A B then bc(\a. if a IN A then p(ab(a)) else ARB) else ARB) (functions (functions A B) C) = functions (functions A B) C `, rw[EXTENSION;functions;in_image_vw;IN_ELIM_THM] THEN once_rw[elim_image_subset_u] THEN once_rw[elim_image_subset_v] THEN rw[IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THENL [ simp[] THEN sufficesby(know `!u. (!v:X. v IN A ==> u v:Y IN B) /\ (!z. ~(z IN A) ==> u z = ARB) ==> v u:Z IN C`) THEN simp[BETA_THM;o_DEF] THEN qed[PERMUTES_IN_IMAGE] ; simp[] ] ; intro THENL [ witness `\ab:X->Y. if ab IN functions A B then x(\a. if a IN A then (inverse p:Y->Y)(ab(a)) else ARB) else ARB:Z` THEN rw[functions;elim_image_subset_v;IN_ELIM_THM] THEN intro THENL [ rw[fun_eq_thm_e] THEN intro THEN case `(!u:X. u IN A ==> e u:Y IN B) /\ (!z. ~(z IN A) ==> e z = ARB)` THENL [ case `!v:X. v IN A ==> p (e v:Y):Y IN B` THENL [ subgoal `e:X->Y = (\a. if a IN A then inverse p ((p:Y->Y) (e a)) else ARB)` THENL [ rw[FUN_EQ_THM] THEN qed[PERMUTES_INVERSES] ; pass ] THEN simp[] THEN qed[] ; pass ] THEN simp[] THEN sufficesby(know `!z. ~((!u:X. u IN A ==> z u:Y IN B) /\ (!a. ~(a IN A) ==> z a = ARB)) ==> x z = ARB:Z`) THEN qed[permutes] ; pass ] THEN simp[] ; simp[] THEN sufficesby(know `!u. (!v:X. v IN A ==> u v:Y IN B) /\ (!z. ~(z IN A) ==> u z = ARB) ==> x u:Z IN C`) THEN rw[BETA_THM] THEN simp[] THEN qed[PERMUTES_INVERSE;PERMUTES_IN_IMAGE] ; qed[] ] ] ] );; (* ===== naturals *) let min_le = prove(` !a b. MIN a b <= a /\ MIN a b <= b `, rw[MIN] THEN ARITH_TAC );; let le_min = prove(` !a b c. a <= MIN b c <=> a <= b /\ a <= c `, rw[MIN] THEN ARITH_TAC );; let max_le = prove(` !a b c. MAX a b <= c <=> a <= c /\ b <= c `, rw[MAX] THEN ARITH_TAC );; let fact_1 = prove(` FACT 1 = 1 `, rw[ARITH_RULE `1 = SUC 0`;FACT] THEN ARITH_TAC );; let binom_stair_sum = prove(` !e n. binom(n+e,e) = if e = 0 then 1 else nsum (0..n) (\a. binom(a+e-1,e-1)) `, GEN_TAC THEN INDUCT_TAC THENL [ rw[ARITH_RULE `0+e=e`;BINOM_REFL] THEN rw[NSUM_SING_NUMSEG] THEN rw[ARITH_RULE `0+e-1 = e-1`;BINOM_REFL] THEN rw[COND_ID] ; pass ] THEN intro THEN rw[NSUM_CLAUSES_NUMSEG;ARITH_RULE `0 <= SUC n`] THEN proven_if `e = 0` [binom] THEN num_linear_fact `~(e = 0) ==> e = SUC(e-1)` THEN have `binom(SUC(n + e),SUC(e-1)) = binom(n+e,e) + binom(n+e,e-1)` [binom] THEN have `binom(SUC(n + e),e) = binom(n+e,e) + binom(n+e,e-1)` [] THEN have `binom(SUC n + e,e) = binom(n+e,e) + binom(n+e,e-1)` [ARITH_RULE `SUC n+e = SUC(n+e)`] THEN num_linear_fact `~(e = 0) ==> SUC n + e - 1 = n+e` THEN qed[] );; (* simpler special case of NSUM_REFLECT *) let nsum_reflect_0 = prove(` !f n. nsum(0..n) f = nsum(0..n) (\i. f(n-i)) `, intro THEN num_linear_fact `~(n < 0)` THEN specialize[`f:num->num`;`0`;`n:num`]NSUM_REFLECT THEN qed[ARITH_RULE `n-0 = n`] );; (* GEN `b` NSUM_DELTA modulo order of variables *) let nsum_delta = prove(` !s a:X b. nsum s (\x. if x = a then b else 0) = (if a IN s then b else 0) `, qed[NSUM_DELTA] );; let term_le_nsum = prove(` !f S t:X. FINITE S ==> t IN S ==> f t <= nsum S f `, qed[NSUM_DELETE;ARITH_RULE `a <= a + b:num`] );; let binom_reverse_stair_sum = prove(` !e n. binom(n+e,e) = if e = 0 then 1 else nsum (0..n) (\a. binom(n-a+e-1,e-1)) `, intro THEN specialize[`e:num`;`n:num`]binom_stair_sum THEN specialize[`\a. binom(a+e-1,e-1)`;`n:num`]nsum_reflect_0 THEN qed[] );; let fact_binom_lemma_37 = prove(` !n i. i < n ==> FACT(n-i) * binom(n,i) = (FACT(n-1-i) * binom(n-1,i)) + ((n-1) * FACT(n-1-i) * binom(n-1,i)) `, intro THEN subgoal `FACT(n-i) = (n-i)*FACT(n-1-i)` THENL [ num_linear_fact `i < n ==> n-i = SUC(n-1-i)` THEN qed[FACT] ; pass ] THEN simp[] THEN subgoal `(n - i) * binom (n,i) = binom (n - 1,i) + (n - 1) * binom (n - 1,i)` THENL [ case `i = 0` THENL [ simp[binom] THEN ASM_ARITH_TAC ; pass ] THEN specialize[`n-1`;`i:num`]BINOM_TOP_STEP THEN num_linear_fact `i < n ==> n-1+1 = n` THEN have `(n-i)*binom(n,i) = n*binom(n-1,i)` [] THEN num_linear_fact `i < n ==> n = 1 + (n-1)` THEN have `1 * binom(n-1,i) + (n-1) * binom(n-1,i) = n * binom(n-1,i)` [RIGHT_ADD_DISTRIB] THEN have `binom(n-1,i) + (n-1) * binom(n-1,i) = n * binom(n-1,i)` [MULT_CLAUSES] THEN qed[] ; pass ] THEN have `((n - i) * binom (n,i)) * FACT(n-1-i) = (binom (n - 1,i)) * FACT(n-1-i) + ((n - 1) * binom (n - 1,i)) * FACT(n-1-i)` [RIGHT_ADD_DISTRIB] THEN ASM_ARITH_TAC );; let properly_le = prove(` properly (<=) = (<):num->num->bool `, rw[properly;FUN_EQ_THM] THEN ARITH_TAC );; let o_permutes_subset = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> A o f SUBSET S `, rw[SUBSET;IN;o_THM;permutes] THEN qed[] );; let image_permutes_subset = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> IMAGE f A SUBSET S `, rw[SUBSET;IN_IMAGE;permutes] THEN qed[] );; let image_inverse_permutes = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> IMAGE (inverse f) A = A o f `, rw[EXTENSION;IN_IMAGE;SUBSET;IN;o_THM] THEN rw[permutes;inverse] THEN qed[] );; let image_o_permutes = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> IMAGE f (A o f) = A `, rw[EXTENSION;SUBSET;IN_ELIM_THM;IN;IMAGE;o_THM;permutes] THEN qed[] );; let o_image_permutes = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> (IMAGE f A) o f = A `, rw[EXTENSION;SUBSET;IMAGE;IN;IN_ELIM_THM;o_THM;permutes] THEN qed[] );; let card_image_permutes = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> FINITE A ==> CARD(IMAGE f A) = CARD A `, intro THEN sufficesby CARD_IMAGE_INJ THEN qed[permutes] );; let card_o_permutes = prove(` !A S f:X->X. f permutes S ==> A SUBSET S ==> FINITE A ==> CARD(A o f) = CARD A `, intro THEN have `A o (f:X->X) = IMAGE (inverse f) A` [image_inverse_permutes] THEN simp[] THEN sufficesby CARD_IMAGE_INJ THEN have `inverse (f:X->X) permutes S` [PERMUTES_INVERSE] THEN qed[permutes] );; let o_permutes_cancel = prove(` !S f:X->X g h:X->Y. f permutes S ==> g o f = h o f ==> g = h `, intro THEN subgoal `g:X->Y = (g o (f:X->X)) o inverse f` THENL [ rw[GSYM o_ASSOC] THEN qed[PERMUTES_INVERSES_o;I_O_ID] ; pass ] THEN subgoal `h:X->Y = (h o (f:X->X)) o inverse f` THENL [ rw[GSYM o_ASSOC] THEN qed[PERMUTES_INVERSES_o;I_O_ID] ; pass ] THEN qed[] );; (* ===== sets of naturals *) let range = new_definition ` range = (>):num->num->bool `;; let range_lt = prove(` !n i. i IN range n <=> i < n `, rw[range;IN;GT] );; let range_set_lt = prove(` !n. range n = {i | i < n} `, rw[EXTENSION;range_lt;IN_ELIM_THM] );; let range_0 = prove(` range 0 = {} `, rw[EXTENSION;range_lt;EMPTY;IN_ELIM_THM] THEN ARITH_TAC );; let range_add_1_delete_refl = prove(` !n. range(n+1) DELETE n = range n `, rw[DELETE;range_lt;EXTENSION;IN_ELIM_THM] THEN ARITH_TAC );; let finite_range = prove(` !n. FINITE(range n) `, rw[range_set_lt] THEN qed[FINITE_NUMSEG_LT] );; let card_range = prove(` !n. CARD(range n) = n `, rw[range_set_lt] THEN qed[CARD_NUMSEG_LT] );; let swap_permutes_range = prove(` !n i j. i < n ==> j < n ==> swap(i,j) permutes range n `, rw[swap;permutes;range_lt] THEN qed[] );; (* maybe use ITSET? *) let finite_ordering = prove(` !S:X->bool. FINITE S ==> ?f. IMAGE f (range(CARD S)) = S `, sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[CARD_CLAUSES;range_0;IMAGE_CLAUSES] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN witness `\i:num. if i = CARD(S:X->bool) then x else f i:X` THEN simp[EXTENSION;in_image_vw;IN_INSERT;CARD_CLAUSES] THEN rw[ARITH_RULE `SUC n = 1+n`] THEN intro THEN splitiff THENL [ intro THEN have `v < 1+CARD(S:X->bool)` [range_lt] THEN proven_if `v = CARD(S:X->bool)` [] THEN num_linear_fact `v < 1+CARD(S:X->bool) ==> ~(v = CARD S) ==> v < CARD S` THEN have `v IN range(CARD(S:X->bool))` [range_lt] THEN qed[IN_IMAGE] ; intro THENL [ witness `CARD(S:X->bool)` THEN rw[range_lt] THEN simp[ARITH_RULE `n < 1+n`] ; choose `v:num` `v IN range(CARD(S:X->bool)) /\ f v = x':X` [IN_IMAGE] THEN witness `v:num` THEN have `v < CARD(S:X->bool)` [range_lt] THEN num_linear_fact `v < CARD(S:X->bool) ==> ~(v = CARD S)` THEN num_linear_fact `v < CARD(S:X->bool) ==> v < 1 + CARD S` THEN qed[range_lt] ] ] ] );; (* XXX: relies on CARD S always being finite *) let injective_finite_ordering = prove(` !S:X->bool f. IMAGE f (range(CARD S)) = S ==> (!x y. x < CARD S ==> y < CARD S ==> f x = f y ==> x = y) `, rw[GSYM range_lt] THEN intro THEN have `FINITE(range(CARD(S:X->bool)))` [finite_range] THEN have `CARD(range(CARD(S:X->bool))) = CARD S` [card_range] THEN qed[CARD_IMAGE_EQ_INJ] );; (* XXX: relies on CARD S always being finite *) let bij_finite_ordering = prove(` !S:X->bool f. IMAGE f (range(CARD S)) = S ==> BIJ f (range(CARD S)) S `, rw[BIJ;INJ;SURJ] THEN REPEAT GEN_TAC THEN DISCH_TAC THEN have `!x y. x < CARD(S:X->bool) ==> y < CARD S ==> f x = f y:X ==> x = y` [injective_finite_ordering] THEN ASM SET_TAC[range_lt] );; let image_numseg_antidiagonal = prove(` !d:num. IMAGE (\a:num. a,d-a) (0..d) = {a,b | a + b = d} `, rw[GSYM NUMSEG_LE;IMAGE;EXTENSION;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN witness `x':num` THEN witness `d - x':num` THEN qed[SUB_ADD;ADD_SYM] ; intro THEN witness `a:num` THEN qed[LE_EXISTS;ADD_SUB2] ] );; let numseg_le_lt_reflect = prove(` !m n:num. m <= n ==> IMAGE (\a. n - a) {i | i < n - m} = {i | m < i /\ i <= n} `, rw[IMAGE;EXTENSION;IN_ELIM_THM] THEN intro THEN splitiff THENL [ ASM_ARITH_TAC ; intro THEN witness `n - x:num` THEN ASM_ARITH_TAC ] );; let numseg_le_reflect_0 = prove(` !n:num. IMAGE (\a. n - a) (0..n) = (0..n) `, rw[GSYM NUMSEG_LE;IMAGE;EXTENSION;IN_ELIM_THM] THEN intro THEN splitiff THENL [ qed[ARITH_RULE `n - a <= n:num`] ; intro THEN witness `n - x:num` THEN qed[ARITH_RULE `n - a <= n:num`;ARITH_RULE `x <= n ==> x = n - (n - x):num`] ] );; (* ===== reals *) let real_of_num_plus_minus_minus = prove(` !a b. &(a+b) - &a - &b = &0:real `, intro THEN num_linear_fact `a <= (a+b):num` THEN num_linear_fact `b <= (a+b)-a:num` THEN num_linear_fact `(a+b)-a-b = 0` THEN simp[REAL_OF_NUM_SUB] );; let fact_binom_lemma_37_real = prove(` !n i. i < n ==> &(FACT(n-i) * binom(n,i)) - &(FACT(n-1-i) * binom(n-1,i)) - &((n-1) * FACT(n-1-i) * binom(n-1,i)) = &0:real `, qed[real_of_num_plus_minus_minus;fact_binom_lemma_37] );; (* ===== integer sums *) let isum_integer_sum = prove(` !(f:X->int) S. FINITE S ==> isum S f = ring_sum integer_ring S f `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[ISUM_CLAUSES;RING_SUM_CLAUSES] THEN rw[INTEGER_RING] ; simp[ISUM_CLAUSES;RING_SUM_CLAUSES] THEN rw[INTEGER_RING;IN_UNIV] ] );; let sum_real_of_int = prove(` !(f:X->int) S. FINITE S ==> sum S (\s. real_of_int (f s)) = real_of_int (isum S f) `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[SUM_CLAUSES;ISUM_CLAUSES;REAL_OF_INT_CLAUSES] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[SUM_CLAUSES;ISUM_CLAUSES;REAL_OF_INT_CLAUSES] ] );; let int_of_num_sum = prove(` !(f:X->num) S. FINITE S ==> &(nsum S f):int = isum S (\x. &(f x)) `, intro THEN rw[int_eq] THEN simp[INT_OF_REAL_THM REAL_OF_NUM_SUM;sum_real_of_int] );; (* ===== rings *) let ring_div_refl = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_div r c c = if ring_unit r c then ring_1 r else ring_0 r `, rw[ring_div;ring_inv] THEN intro THEN case `ring_unit(r:R ring) c` THENL [ choose `d:R` `d IN ring_carrier(r:R ring) /\ ring_mul r c d = ring_1 r` [ring_unit] THEN qed[] ; qed[RING_MUL_RZERO] ] );; let ring_add_sub_cancel = prove(` !(r:R ring) a b. a IN ring_carrier r ==> b IN ring_carrier r ==> ring_add r a (ring_sub r b a) = b `, RING_TAC );; let ring_sum_image_injective = prove(` !r (f:K->L) (g:L->A) s. (!x y. f x = f y ==> x = y) ==> ring_sum r (IMAGE f s) g = ring_sum r s (g o f) `, qed[RING_SUM_IMAGE] );; let ring_sum_image_injective_pair = prove(` !r (f:A#B->L) (g:L->X) s. (!(a,b) (c,d). f (a,b) = f (c,d) ==> (a,b) = (c,d)) ==> ring_sum r (IMAGE f s) g = ring_sum r s (g o f) `, intro THEN have `!x y. (f:A#B->L) x = f y ==> x = y` [injective_pair_rewrite] THEN qed[ring_sum_image_injective] );; let ring_sum_delete2 = prove(` !(r:R ring) S (f:X->R) s. FINITE S ==> s IN S ==> f s IN ring_carrier r ==> ring_sum r S f = ring_add r (f s) (ring_sum r (S DELETE s) f) `, intro THEN simp[RING_SUM_DELETE] THEN have `ring_sum (r:R ring) S (f:X->R) IN ring_carrier r` [RING_SUM] THEN qed[ring_add_sub_cancel] );; let ring_sum_shift1 = prove(` !(r:R ring) f n. f 0 = ring_0 r ==> ring_sum r (0..n+1) f = ring_sum r (0..n) (\a. f (a+1)) `, rw[GSYM NUMSEG_LE] THEN intro THEN set_fact_using `0 IN {a | a <= n+1}` [ARITH_RULE `0 <= n+1`] THEN have `f 0 IN ring_carrier(r:R ring)` [RING_0] THEN specialize[`n+1`]FINITE_NUMSEG_LE THEN specialize[`r:R ring`;`{a | a <= n+1}`;`f:num->R`;`0`]ring_sum_delete2 THEN simp[RING_ADD_LZERO;RING_SUM] THEN rw[NUMSEG_LE] THEN rw[GSYM RING_SUM_OFFSET] THEN have `0 INSERT ((0+1)..(n+1)) = 0..(n+1)` [NUMSEG_LREC;ARITH_RULE `0 <= n+1`] THEN have `~(0 IN ((0+1)..(n+1)))` [IN_NUMSEG;ARITH_RULE `~(0+1 <= 0)`] THEN have `(0..(n+1)) DELETE 0 = ((0+1)..(n+1))` [insert_delete_nonmember] THEN simp[] );; let ring_sum_insert_top = prove(` !(r:R ring) f n. f (n+1) = ring_0 r ==> ring_sum r (0..n+1) f = ring_sum r (0..n) f `, rw[GSYM NUMSEG_LE] THEN intro THEN set_fact_using `n+1 IN {a | a <= n+1}` [ARITH_RULE `n+1 <= n+1`] THEN have `f (n+1) IN ring_carrier(r:R ring)` [RING_0] THEN specialize[`n+1`]FINITE_NUMSEG_LE THEN specialize[`r:R ring`;`{a | a <= n+1}`;`f:num->R`;`n+1`]ring_sum_delete2 THEN simp[RING_ADD_LZERO;RING_SUM] THEN have `(n+1) INSERT (0..((n+1)-1)) = 0..(n+1)` [NUMSEG_RREC;ARITH_RULE `0 <= n+1`] THEN have `(n+1)-1 = n` [ADD_SUB] THEN have `(n+1) INSERT (0..n) = 0..(n+1)` [NUMSEG_RREC;ARITH_RULE `0 <= n+1`] THEN have `~((n+1) IN (0..n))` [IN_NUMSEG;ARITH_RULE `~(n+1 <= n)`] THEN have `(0..(n+1)) DELETE (n+1) = 0..n` [insert_delete_nonmember] THEN rw[NUMSEG_LE] THEN simp[] );; let ring_mul_sum_mul_delete = prove(` !(r:R ring) S:X->bool f:X->R g:X->R x. ~(x IN S) ==> FINITE S ==> g x IN ring_carrier r ==> (!s. s IN S ==> g s IN ring_carrier r) ==> (!s. s IN S ==> f s IN ring_carrier r) ==> ring_mul r (g x) (ring_sum r S (\s. ring_mul r (f s) (ring_product r (S DELETE s) g))) = ring_sum r S (\s. ring_mul r (f s) (ring_product r ((x INSERT S) DELETE s) g)) `, intro THEN have `!s:X. s IN S ==> ring_product r (S DELETE s) g IN ring_carrier(r:R ring)` [RING_PRODUCT] THEN have `!s:X. s IN S ==> ring_mul r (f s) (ring_product r (S DELETE s) g) IN ring_carrier(r:R ring)` [RING_MUL] THEN have `!s:X. s IN S ==> ring_product r ((x INSERT S) DELETE s) g IN ring_carrier(r:R ring)` [RING_PRODUCT] THEN specialize[`r:R ring`;`\s:X. ring_mul(r:R ring) (f s) (ring_product r (S DELETE s) g)`;`g(x:X):R`;`S:X->bool`](GSYM RING_SUM_LMUL) THEN simp[] THEN sufficesby RING_SUM_EQ THEN intro THEN simp[] THEN set_fact `!s:X. s IN S ==> ~(x IN S) ==> (x INSERT S) DELETE s = x INSERT (S DELETE s)` THEN have `(x:X INSERT S) DELETE a = x INSERT (S DELETE a)` [] THEN have `FINITE(S DELETE a:X)` [FINITE_DELETE] THEN set_fact `!s:X. s IN S ==> ~(x IN S) ==> ~(x IN S DELETE s)` THEN have `~(x:X IN S DELETE a)` [] THEN simp[RING_PRODUCT_CLAUSES] THEN have `g(x:X) IN ring_carrier(r:R ring)` [] THEN have `f(a:X) IN ring_carrier(r:R ring)` [] THEN have `ring_product r (S DELETE a:X) g IN ring_carrier(r:R ring)` [RING_PRODUCT] THEN specialize[`r:R ring`;`f(a:X):R`;`g(x:X):R`;`ring_product(r:R ring) (S DELETE a:X) g`]RING_MUL_ASSOC THEN specialize[`r:R ring`;`g(x:X):R`;`f(a:X):R`;`ring_product(r:R ring) (S DELETE a:X) g`]RING_MUL_ASSOC THEN specialize[`r:R ring`;`f(a:X):R`;`g(x:X):R`]RING_MUL_SYM THEN qed[] );; (* simpler special case of RING_SUM_REFLECT *) let ring_sum_numseg_le_reflect = prove(` !(r:R ring) n f:num->R. ring_sum r (0..n) f = ring_sum r (0..n) (\b. f(n - b)) `, intro THEN subgoal `!x:num y:num. x IN (0..n) /\ y IN (0..n) /\ n-x = n-y ==> x = y` THENL [ rw[IN_ELIM_THM;GSYM NUMSEG_LE] THEN qed[IN_ELIM_THM;ARITH_RULE `x <= n /\ y <= n /\ n-x = n-y ==> x = y:num`] ; pass ] THEN specialize_assuming[`r:R ring`;`\a:num. n - a`;`f:num->R`;`0..n`]RING_SUM_IMAGE THEN have `ring_sum (r:R ring) (IMAGE (\a:num. n - a) (0..n)) f = ring_sum r (0..n) (f o (\a. n - a))` [] THEN have `IMAGE (\a:num. n - a) (0..n) = (0..n)` [numseg_le_reflect_0] THEN have `ring_sum (r:R ring) (0..n) f = ring_sum r (0..n) (f o (\a. n - a))` [] THEN simp[o_THM] THEN sufficesby RING_SUM_EQ THEN qed[o_THM] );; let ring_sum_numseg_le_offset = prove(` !(r:R ring) m n f:num->R. ring_sum r (0..m) f = ring_sum r (0..m+n) (\b. if n <= b then f(b - n) else ring_0 r) `, intro THEN specialize[`n:num`;`r:R ring`;`\i:num. f(i - n):R`;`0:num`;`m:num`]RING_SUM_OFFSET THEN have `(0+n..m+n) SUBSET (0..m+n)` [SUBSET_NUMSEG;ARITH_RULE `0 <= 0+n`;ARITH_RULE `m+n <= m+n:num`] THEN subgoal `!x. x IN 0..m+n ==> ~(x IN 0+n..m+n) ==> (if n <= x then f(x-n) else ring_0(r:R ring)) = ring_0 r` THENL [ rw[numseg;IN_ELIM_THM] THEN qed[ARITH_RULE `x:num <= m+n /\ ~(0+n <= x /\ x <= m+n) ==> ~(n <= x)`] ; pass ] THEN specialize[`r:R ring`;`\b:num. if n <= b then f(b-n) else ring_0(r:R ring)`;`0+n..m+n`;`0..m+n`]RING_SUM_SUPERSET THEN subgoal `!a. a IN 0+n..m+n ==> (if n <= a then f(a-n) else ring_0 r) = f(a-n):R` THENL [ rw[numseg;IN_ELIM_THM] THEN qed[ARITH_RULE `0+n = n`] ; pass ] THEN specialize[`r:R ring`;`\i:num. if n <= i then f(i-n) else ring_0(r:R ring)`;`\i:num. f(i-n):R`;`0+n..m+n`]RING_SUM_EQ THEN simp[] THEN sufficesby RING_SUM_EQ THEN qed[ARITH_RULE `(b+n)-n = b:num`] );; let ring_sum_numseg_le_expand = prove(` !(r:R ring) m n f:num->R. m <= n ==> ring_sum r (0..m) f = ring_sum r (0..n) (\a. if a <= m then f(a) else ring_0 r) `, intro THEN subgoal `ring_sum(r:R ring) (0..m) f = ring_sum r (0..m) (\a. if a <= m then f(a) else ring_0 r)` THENL [ sufficesby RING_SUM_EQ THEN rw[GSYM NUMSEG_LE;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[] THEN sufficesby(GSYM RING_SUM_SUPERSET) THEN intro THENL [ rw[SUBSET_NUMSEG] THEN qed[ARITH_RULE `0 <= 0`] ; have `~(x <= m:num)` [ NUMSEG_LE;IN_ELIM_THM ] THEN simp[] ] );; let num_in_subring = prove(` !(r:R ring) s. s subring_of r ==> !n. ring_of_num r n IN s `, rw[subring_of] THEN GEN_TAC THEN GEN_TAC THEN DISCH_TAC THEN INDUCT_TAC THEN qed[ring_of_num] );; let int_in_subring = prove(` !(r:R ring) s. s subring_of r ==> !i. ring_of_int r i IN s `, rw[FORALL_INT_CASES] THEN rw[RING_OF_INT_NEG;RING_OF_INT_OF_NUM] THEN qed[num_in_subring;subring_of] );; let ring_sum_subring_generated = prove(` !(r:R ring) S A (f:X->R). S subring_of r /\ (!x. x IN A ==> f x IN S) ==> ring_sum (subring_generated r S) A f = ring_sum r A f `, simp[] THEN intro THEN set_fact `(!x:X. x IN A ==> (f x):R IN S) ==> {x | x IN A /\ f x IN S} = A` THEN simp[RING_SUM_SUBRING_GENERATED] );; let ring_sum_subring_generated_v2 = prove(` !(r:R ring) S A (f:X->R). (!x. x IN A ==> f x IN ring_carrier(subring_generated r S)) ==> ring_sum (subring_generated r S) A f = ring_sum r A f `, intro THEN have `ring_carrier(subring_generated(r:R ring) S) subring_of r` [SUBRING_SUBRING_GENERATED] THEN have `subring_generated(r:R ring) (ring_carrier(subring_generated r S)) = subring_generated r S` [SUBRING_GENERATED_BY_SUBRING_GENERATED] THEN qed[ring_sum_subring_generated] );; let ring_sum_in_subring = prove(` !(r:R ring) G S (f:X->R). (!s. s IN S ==> f s IN ring_carrier(subring_generated r G)) ==> ring_sum r S f IN ring_carrier(subring_generated r G) `, qed[ring_sum_subring_generated_v2;RING_SUM] );; let ring_product_subring_generated = prove(` !(r:R ring) S A (f:X->R). S subring_of r /\ (!x. x IN A ==> f x IN S) ==> ring_product (subring_generated r S) A f = ring_product r A f `, simp[] THEN intro THEN set_fact `(!x:X. x IN A ==> (f x):R IN S) ==> {x | x IN A /\ f x IN S} = A` THEN simp[RING_PRODUCT_SUBRING_GENERATED] );; let ring_product_subring_generated_v2 = prove(` !(r:R ring) S A (f:X->R). (!x. x IN A ==> f x IN ring_carrier(subring_generated r S)) ==> ring_product (subring_generated r S) A f = ring_product r A f `, intro THEN have `ring_carrier(subring_generated(r:R ring) S) subring_of r` [SUBRING_SUBRING_GENERATED] THEN have `subring_generated(r:R ring) (ring_carrier(subring_generated r S)) = subring_generated r S` [SUBRING_GENERATED_BY_SUBRING_GENERATED] THEN qed[ring_product_subring_generated] );; let ring_product_in_subring = prove(` !(r:R ring) G S (f:X->R). (!s. s IN S ==> f s IN ring_carrier(subring_generated r G)) ==> ring_product r S f IN ring_carrier(subring_generated r G) `, qed[ring_product_subring_generated_v2;RING_PRODUCT] );; let ring_pow_in_subring = prove(` !(r:R ring) G f n. f IN ring_carrier(subring_generated r G) ==> ring_pow r f n IN ring_carrier(subring_generated r G) `, qed[RING_POW_SUBRING_GENERATED;RING_POW] );; let ring_sum_1 = prove(` !(r:R ring) S. FINITE S ==> ring_sum r S (\s:X. ring_1 r) = ring_of_num r (CARD S) `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[card_empty;RING_SUM_CLAUSES;RING_OF_NUM_0] ; simp[card_insert;RING_SUM_CLAUSES;RING_1;ARITH_RULE `CARD S + 1 = SUC(CARD(S:X->bool))`;ring_of_num] THEN qed[RING_ADD_SYM;RING_1;RING_OF_NUM] ] );; let ring_sum_num = prove(` !(r:R ring) (f:X->num) S. FINITE S ==> ring_sum r S (\s. ring_of_num r (f s)) = ring_of_num r (nsum S f) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;NSUM_CLAUSES;ring_of_num] ; simp[RING_SUM_CLAUSES;NSUM_CLAUSES] THEN rw[RING_OF_NUM;RING_OF_NUM_ADD] ] );; let ring_pow_sub1 = prove(` !(r:R ring) c d. c IN ring_carrier r ==> ~(d = 0) ==> ring_mul r c (ring_pow r c (d-1)) = ring_pow r c d `, intro THEN have `ring_pow(r:R ring) c (1 + (d-1)) = ring_mul r (ring_pow r c 1) (ring_pow r c (d-1))` [RING_POW_ADD] THEN have `ring_pow(r:R ring) c 1 = c` [RING_POW_1] THEN num_linear_fact `~(d = 0) ==> 1 + (d-1) = d:num` THEN qed[] );; let ring_pow_product = prove(` !(r:R ring) (p:X->R) n S. FINITE S ==> (!s. s IN S ==> p s IN ring_carrier r) ==> ring_pow r (ring_product r S p) n = ring_product r S (\s:X. ring_pow r (p s) n) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN rw[RING_POW_ONE] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES;RING_POW] THEN have `p(x:X) IN ring_carrier(r:R ring)` [] THEN simp[RING_MUL_POW;RING_PRODUCT] ] );; let ring_sum_numseg_0_diff = prove(` !(r:R ring) m n f. m <= n ==> ring_sum r {i | m < i /\ i <= n} f = ring_sub r (ring_sum r (0..n) f) (ring_sum r (0..m) f) `, intro THEN subgoal `{i | m < i /\ i <= n} = (0..n) DIFF (0..m)` THENL [ rw[EXTENSION;IN_DIFF;IN_ELIM_THM;IN_NUMSEG] THEN ARITH_TAC ; pass ] THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN have `(0..m) SUBSET (0..n)` [SUBSET_NUMSEG;ARITH_RULE `0 <= 0`] THEN qed[RING_SUM_DIFF] );; let ring_sum_numseg_0_diff_reflect = prove(` !(r:R ring) m n f. m <= n ==> ring_sum r {i | i < n-m} (\i. f(n-i)) = ring_sub r (ring_sum r (0..n) f) (ring_sum r (0..m) f) `, simp[GSYM ring_sum_numseg_0_diff] THEN intro THEN subgoal `!x y. x IN {i:num | i < n - m} ==> y IN {i | i < n - m} ==> n - x = n - y ==> x = y` THENL [ rw[IN_ELIM_THM] THEN ARITH_TAC ; pass ] THEN specialize[`r:R ring`;`\i:num. n-i`;`f:num->R`;`{i | i < n - m:num}`]RING_SUM_IMAGE THEN have `IMAGE (\i:num. n - i) {i | i < n - m} = {i | m < i /\ i <= n}` [numseg_le_lt_reflect] THEN have `ring_sum(r:R ring) {i:num | m < i /\ i <= n} f = ring_sum r {i | i < n - m} (f o (\i. n - i))` [] THEN have `ring_sum(r:R ring) {i:num | i < n - m} (f o (\i. n - i)) = ring_sum r {i | i < n - m} (\i. f (n - i))` [RING_SUM_EQ;o_THM] THEN qed[] );; let ring_pow_is_product = prove(` !(r:R ring) a n. a IN ring_carrier r ==> ring_pow r a n = ring_product r (1..n) (\i. a) `, GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[NUMSEG_CLAUSES;ARITH_RULE `~(1 = 0)`] THEN rw[ring_pow;RING_PRODUCT_CLAUSES] ; rw[NUMSEG_CLAUSES;ARITH_RULE `1 <= SUC n`] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `~(SUC n IN 1..n)` [IN_NUMSEG;ARITH_RULE `~(SUC n <= n)`] THEN simp[ring_pow;RING_PRODUCT_CLAUSES] ] );; let ring_sum_sub = prove(` !(r:R ring) f g S. FINITE S ==> (!s:X. s IN S ==> f s IN ring_carrier r) ==> (!s:X. s IN S ==> g s IN ring_carrier r) ==> ring_sum r S (\s. ring_sub r (f s) (g s)) = ring_sub r (ring_sum r S f) (ring_sum r S g) `, intro THEN rw[ring_sub] THEN specialize_assuming[`r:R ring`;`f:X->R`;`(\s:X. ring_neg(r:R ring) (g s))`;`S:X->bool`]RING_SUM_ADD THEN have `ring_sum(r:R ring) S (\s:X. ring_add r (f s) (ring_neg r (g s))) = ring_add r (ring_sum r S f) (ring_sum r S (\s. ring_neg r (g s)))` [RING_NEG] THEN simp[] THEN qed[RING_SUM_NEG] );; (* XXX: merge this into ring_ord_unique *) let unique_prime_valuation_lemma = prove(` !(r:R ring) p d e f g. integral_domain r ==> ring_prime r p ==> f IN ring_carrier r ==> g IN ring_carrier r ==> ~(ring_divides r p g) ==> ring_mul r (ring_pow r p d) f = ring_mul r (ring_pow r p e) g ==> d <= e `, intro THEN have `p IN ring_carrier(r:R ring)` [ring_prime] THEN have `~(p = ring_0(r:R ring))` [ring_prime] THEN have `~(ring_pow r p e = ring_0(r:R ring))` [INTEGRAL_DOMAIN_POW_EQ_0] THEN case `~(d <= e:num)` THENL [ num_linear_fact `~(d <= e) ==> e < d:num` THEN num_linear_fact `e < d:num ==> d = e + (d-e)` THEN have `ring_pow(r:R ring) p d = ring_mul r (ring_pow r p e) (ring_pow r p (d-e))` [RING_POW_ADD] THEN have `ring_mul(r:R ring) (ring_mul r (ring_pow r p e) (ring_pow r p (d-e))) f = ring_mul r (ring_pow r p e) g` [] THEN have `ring_mul(r:R ring) (ring_pow r p e) (ring_mul r (ring_pow r p (d-e)) f) = ring_mul r (ring_pow r p e) g` [RING_MUL_ASSOC;RING_POW] THEN have `(ring_mul(r:R ring) (ring_pow r p (d-e)) f) = g` [INTEGRAL_DOMAIN_MUL_LCANCEL;RING_POW;RING_MUL] THEN num_linear_fact `e < d:num ==> d-e = 1 + (d-e-1)` THEN have `ring_pow(r:R ring) p (d-e) = ring_mul r (ring_pow r p 1) (ring_pow r p (d-e-1))` [RING_POW_ADD] THEN have `ring_pow(r:R ring) p (d-e) = ring_mul r p (ring_pow r p (d-e-1))` [RING_POW_1] THEN have `(ring_mul(r:R ring) (ring_mul r p (ring_pow r p (d-e-1))) f) = g` [] THEN have `ring_mul(r:R ring) p (ring_mul r (ring_pow r p (d-e-1)) f) = g` [RING_MUL_ASSOC;RING_POW] THEN qed[ring_divides;RING_POW;RING_MUL] ; pass ] THEN qed[] );; (* XXX: merge this into ring_ord_unique *) let unique_prime_valuation = prove(` !(r:R ring) p d e f g. integral_domain r ==> ring_prime r p ==> f IN ring_carrier r ==> g IN ring_carrier r ==> ~(ring_divides r p f) ==> ~(ring_divides r p g) ==> ring_mul r (ring_pow r p d) f = ring_mul r (ring_pow r p e) g ==> (d = e /\ f = g) `, intro THEN have `d <= e:num` [unique_prime_valuation_lemma] THEN have `e <= d:num` [unique_prime_valuation_lemma] THEN num_linear_fact `d <= e:num ==> e <= d ==> d = e` THENL [ qed[] ; have `p IN ring_carrier(r:R ring)` [ring_prime] THEN have `~(p = ring_0(r:R ring))` [ring_prime] THEN have `~(ring_pow r p e = ring_0(r:R ring))` [INTEGRAL_DOMAIN_POW_EQ_0] THEN qed[INTEGRAL_DOMAIN_MUL_LCANCEL;RING_POW] ] );; let ring_sum_const = prove(` !(r:R ring) c S. FINITE S ==> c IN ring_carrier r ==> ring_sum r S (\s:X. c) = ring_mul r (ring_of_num r (CARD S)) c `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;CARD_CLAUSES;RING_OF_NUM_0] THEN qed[RING_MUL_LZERO] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[RING_SUM_CLAUSES;CARD_CLAUSES;ring_of_num] THEN RING_TAC ] );; let ring_product_const = prove(` !(r:R ring) c S. FINITE S ==> c IN ring_carrier r ==> ring_product r S (\s:X. c) = ring_pow r c (CARD S) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;CARD_CLAUSES;RING_POW_0] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES;CARD_CLAUSES;ring_pow] ] );; let ring_of_num_injective_lemma = prove(` !(r:R ring) m n. ring_char r = 0 ==> ring_of_num r m = ring_of_num r n ==> n <= m `, intro THEN case `m < n:num` THENL [ num_linear_fact `m < n:num ==> n = (n-m)+m` THEN have `ring_of_num(r:R ring) n = ring_add r (ring_of_num r (n-m)) (ring_of_num r m)` [RING_OF_NUM_ADD] THEN have `ring_of_num(r:R ring) m = ring_add r (ring_of_num r (n-m)) (ring_of_num r m)` [] THEN have `ring_of_num r (n-m) = ring_0(r:R ring)` [RING_ADD_EQ_RIGHT;RING_OF_NUM] THEN have `n-m = 0` [RING_CHAR_EQ_0] THEN qed[ARITH_RULE `m < n ==> ~(n-m = 0)`] ; ASM_ARITH_TAC ] );; (* compare RING_CHAR_EQ_0, RING_OF_NUM_EQ *) let ring_of_num_injective = prove(` !(r:R ring) m n. ring_char r = 0 ==> ring_of_num r m = ring_of_num r n ==> m = n `, intro THEN have `n <= m:num` [ring_of_num_injective_lemma] THEN have `m <= n:num` [ring_of_num_injective_lemma] THEN ASM_ARITH_TAC );; let ring_of_num_nonzero = prove(` !(r:R ring) n. ring_char r = 0 ==> ring_of_num r n = ring_0 r ==> n = 0 `, qed[ring_of_num_injective;RING_OF_NUM_0] );; let neg_ring_of_num_nonzero = prove(` !(r:R ring) n. ring_char r = 0 ==> ring_neg r (ring_of_num r n) = ring_0 r ==> n = 0 `, qed[ring_of_num_nonzero;RING_NEG_EQ_0;RING_OF_NUM] );; let ring_pow_neg_1_mul_refl = prove(` !(r:R ring) n. ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) n ) ( ring_pow r (ring_neg r (ring_1 r)) n ) = ring_1 r `, intro THEN havetac `ring_mul(r:R ring) (ring_neg r (ring_1 r)) (ring_neg r (ring_1 r)) = ring_1 r` RING_TAC THEN simp[GSYM RING_MUL_POW;RING_1;RING_NEG] THEN qed[RING_POW_ONE] );; let ring_pow_neg_1_plus1 = prove(` !(r:R ring) n. ring_pow r (ring_neg r (ring_1 r)) (n + 1) = ring_neg r (ring_pow r (ring_neg r (ring_1 r)) n) `, intro THEN simp[RING_POW_ADD;RING_1;RING_NEG] THEN RING_TAC );; let ring_sum_cross_mul = prove(` !(r:R ring) P Q (f:X->R) (g:Y->R). FINITE P ==> FINITE Q ==> (!x:X. x IN P ==> f x IN ring_carrier r) ==> (!y:Y. y IN Q ==> g y IN ring_carrier r) ==> ring_sum r (P CROSS Q) (\(x,y). ring_mul r (f x) (g y)) = ring_mul r (ring_sum r P f) (ring_sum r Q g) `, intro THEN simp[GSYM RING_SUM_RMUL;RING_SUM] THEN simp[GSYM RING_SUM_LMUL;RING_SUM] THEN simp[RING_SUM_SUM_PRODUCT;RING_SUM;RING_MUL] THEN rw[CROSS] );; let ring_product_1_plus_expand = prove(` !(r:R ring) c S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_product r S (\s:X. ring_add r (ring_1 r) (c s)) = ring_sum r {A | A SUBSET S} (\A. ring_product r A c) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POWERSET_CLAUSES;RING_SUM_SING;RING_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_ADD_RDISTRIB;RING_SUM;RING_MUL_LID;RING_PRODUCT_CLAUSES;RING_ADD;RING_1;POWERSET_CLAUSES] THEN have `DISJOINT {A | A SUBSET S} (IMAGE (\A. (x:X) INSERT A) {A | A SUBSET S})` [powerset_insert_disjoint] THEN have `FINITE {A:X->bool | A SUBSET S}` [FINITE_POWERSET] THEN have `FINITE(IMAGE (\A. x INSERT A) {A:X->bool | A SUBSET S})` [FINITE_IMAGE] THEN specialize[ `r:R ring`; `\A. ring_product r A (c:X->R)`; `{A:X->bool | A SUBSET S}`; `IMAGE (\A. x INSERT A) {A:X->bool | A SUBSET S}` ]RING_SUM_UNION THEN simp[] THEN subgoal `!a b. a IN {A | A SUBSET S} ==> b IN {A | A SUBSET S} ==> (x:X) INSERT a = x INSERT b ==> a = b` THENL [ rw[IN_ELIM_THM;SUBSET;INSERT;EXTENSION] THEN intro THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `\A. (x:X) INSERT A`; `\A. ring_product r A (c:X->R)`; `{A:X->bool | A SUBSET S}` ]RING_SUM_IMAGE THEN subgoal `ring_sum r {A | A SUBSET S} ((\A. ring_product r A c) o (\A. x INSERT A)) = ring_sum r {A | A SUBSET S} (\A. ring_mul r (c(x:X):R) (ring_product r A c))` THENL [ sufficesby RING_SUM_EQ THEN rw[IN_ELIM_THM;BETA_THM;o_THM] THEN intro THEN have `FINITE(a:X->bool)` [FINITE_SUBSET] THEN have `~((x:X) IN a)` [SUBSET] THEN simp[RING_PRODUCT_CLAUSES] ; pass ] THEN subgoal `ring_sum r {A | A SUBSET S} (\A. ring_mul r (c x) (ring_product r A (c:X->R))) = ring_mul r (c x) (ring_sum r {A | A SUBSET S} (\A. ring_product r A c))` THENL [ specialize_assuming[ `r:R ring`; `\A. ring_product r A (c:X->R)`; `(c:X->R) x`; `{A:X->bool | A SUBSET S}` ]RING_SUM_LMUL THEN qed[RING_PRODUCT] ; pass ] THEN qed[] ] );; let ring_product_sum_expand = prove(` !(r:R ring) f Q P. FINITE P ==> FINITE Q ==> (!p q. p IN P ==> q IN Q ==> f p q IN ring_carrier r) ==> ring_product r P (\p:X. ring_sum r Q (\q:Y. f p q)) = ring_sum r (functions P Q) (\g. ring_product r P (\p. f p (g p))) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;functions_empty] THEN simp[RING_SUM_SING;RING_1] ; set_fact `(x:X) IN x INSERT P` THEN set_fact `!p:X. p IN P ==> p IN x INSERT P` THEN simp[functions_insert] THEN simp[RING_PRODUCT_CLAUSES;RING_SUM] THEN subgoal `ring_sum(r:R ring) (IMAGE (\(b,f) a:X. if a = x then b else f a:Y) (Q CROSS functions P Q)) (\a. if ~(f x (a x) IN ring_carrier r) then ring_product r P (\p. f p (a p)) else ring_mul r (f x (a x)) (ring_product r P (\p. f p (a p)))) = ring_sum r (IMAGE (\(b,f) a. if a = x then b else f a) (Q CROSS functions P Q)) (\a. ring_mul r (f x (a x)) (ring_product r P (\p. f p (a p))))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;in_image_vw;EXISTS_PAIR_THM;IN_CROSS;functions;IN_ELIM_THM;SUBSET] THEN intro THEN have `a(x:X):Y IN Q` [] THEN qed[] ; pass ] THEN subgoal `ring_sum(r:R ring) (IMAGE (\(b,f) a:X. if a = x then b else f a:Y) (Q CROSS functions P Q)) (\a. ring_mul r (f x (a x)) (ring_product r P (\p. f p (a p)))) = ring_sum r (Q CROSS functions P Q) ((\a. ring_mul r (f x (a x)) (ring_product r P (\p. f p (a p)))) o (\(b,f) a. if a = x then b else f a))` THENL [ specialize_assuming[ `r:R ring`; `(\(b,f) a:X. if a = x then b else f a:Y)`; `\a:X->Y. ring_mul(r:R ring) (f x (a x)) (ring_product r P (\p. f p (a p)))`; `Q CROSS functions (P:X->bool) (Q:Y->bool)` ]RING_SUM_IMAGE THEN specialize[ `P:X->bool`; `Q:Y->bool`; `x:X` ]functions_insert_injective THEN qed[] ; pass ] THEN simp[o_def_s;lambda_pair_ab] THEN simp[GSYM ring_sum_cross_mul;RING_PRODUCT;finite_functions] THEN sufficesby RING_SUM_EQ THEN rw[FORALL_PAIR_THM;IN_CROSS] THEN intro THEN subgoal `ring_product(r:R ring) P (\p:X. f p (p2 p:Y)) = ring_product r P (\p. f p (if p = x then p1 else p2 p))` THENL [ sufficesby RING_PRODUCT_EQ THEN qed[] ; pass ] THEN qed[] ] );; let sub_in_subring = prove(` !(r:R ring) G a b. a IN ring_carrier r ==> b IN ring_carrier(subring_generated r G) ==> ( ring_sub r a b IN ring_carrier(subring_generated r G) <=> a IN ring_carrier(subring_generated r G) ) `, intro THEN splitiff THENL [ intro THEN have `b:R IN ring_carrier r` [RING_CARRIER_SUBRING_GENERATED_SUBSET;SUBSET] THEN specialize[](RING_RULE `a = ring_add(r:R ring) (ring_sub r a b) b`) THEN have `ring_add(subring_generated r G) (ring_sub r a b) b IN ring_carrier(subring_generated(r:R ring) G)` [RING_ADD] THEN qed[SUBRING_GENERATED] ; intro THEN qed[RING_SUB_SUBRING_GENERATED;RING_SUB] ] );; let ring_sum_eq_name_d = prove(` !(r:R ring) f g S:X->bool. (!d. d IN S ==> f d = g d) ==> ring_sum r S f = ring_sum r S g `, qed[RING_SUM_EQ] );; let ring_product_delete = prove(` !(r:R ring) S t f:X->R. FINITE S ==> t IN S ==> f t IN ring_carrier r ==> ring_product r S f = ring_mul r (f t) (ring_product r (S DELETE t) f) `, intro THEN set_fact `(t:X) IN S ==> t INSERT (S DELETE t) = S` THEN have `FINITE (S DELETE (t:X))` [FINITE_DELETE] THEN have `~(t IN (S DELETE (t:X)))` [IN_DELETE] THEN qed[RING_PRODUCT_CLAUSES] );; let ring_divides_pow_pow = prove(` !(r:R ring) a e f. e <= f ==> a IN ring_carrier r ==> ring_divides r (ring_pow r a e) (ring_pow r a f) `, intro THEN num_linear_fact `e <= f ==> f = e+(f-e):num` THEN have `ring_pow r a f:R = ring_mul r (ring_pow r a e) (ring_pow r a (f-e))` [RING_POW_ADD] THEN rw[ring_divides] THEN qed[RING_POW] );; let ring_product_collect = prove(` !(r:R ring) S f:X->R. FINITE S ==> (!s:X. s IN S ==> f s IN ring_carrier r) ==> ring_product r S f = ring_product r (IMAGE f S) (\y. ring_pow r y (CARD {s | s IN S /\ f s = y})) `, intro THEN specialize[ `r:R ring`; `f:X->R`; `f:X->R`; `S:X->bool` ]RING_PRODUCT_IMAGE_GEN THEN simp[] THEN sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `ring_product r {x | x IN S /\ f x = a} f = ring_product r {x | x IN S /\ f x = a} (\x:X. a:R)` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[IN_ELIM_THM] THEN qed[] ; pass ] THEN have `a:R IN ring_carrier r` [IN_IMAGE] THEN set_fact `{x:X | x IN S /\ f x = a:R} SUBSET S` THEN have `FINITE {x:X | x IN S /\ f x = a:R}` [FINITE_SUBSET] THEN simp[ring_product_const] );; let ring_coprime_1 = prove(` !(r:R ring) a. a IN ring_carrier r ==> (ring_coprime r (ring_1 r,a) /\ ring_coprime r (a,ring_1 r) ) `, rw[ring_coprime] THEN qed[RING_1;RING_UNIT_DIVIDES] );; let ring_coprime_product_waterfall = prove(` !(r:R ring) f:X->R a. (UFD r \/ integral_domain r /\ bezout_ring r) ==> a IN ring_carrier r ==> !S. FINITE S ==> (!s. s IN S ==> ring_coprime r (a,f s)) ==> ring_coprime r (a,ring_product r S f) `, intro_gendisch THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ simp[RING_PRODUCT_CLAUSES] THEN qed[ring_coprime_1] ; have `x:X IN x INSERT S` [IN_INSERT] THEN have `f(x:X):R IN ring_carrier r` [SUBSET;ring_coprime] THEN simp[RING_PRODUCT_CLAUSES] THEN have `!s:X. s IN S ==> s IN x INSERT S` [IN_INSERT] THEN have `ring_coprime r (a,ring_product r S (f:X->R))` [] THEN qed[RING_COPRIME_RMUL;RING_PRODUCT;ring_coprime] ] );; let ring_coprime_product = prove(` !(r:R ring) f:X->R a S. (UFD r \/ integral_domain r /\ bezout_ring r) ==> a IN ring_carrier r ==> FINITE S ==> (!s. s IN S ==> ring_coprime r (a,f s)) ==> ring_coprime r (a,ring_product r S f) `, simp[ring_coprime_product_waterfall] );; let ring_product_divides_if_coprime_waterfall = prove(` !(r:R ring) f:X->R a. (UFD r \/ integral_domain r /\ bezout_ring r) ==> a IN ring_carrier r ==> !S. FINITE S ==> (!s t. s IN S ==> t IN S ==> ~(s = t) ==> ring_coprime r (f s,f t)) ==> (!s. s IN S ==> ring_divides r (f s) a) ==> ring_divides r (ring_product r S f) a `, intro_gendisch THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ simp[RING_PRODUCT_CLAUSES] THEN qed[RING_DIVIDES_1] ; have `x:X IN x INSERT S` [IN_INSERT] THEN have `f(x:X):R IN ring_carrier r` [SUBSET;ring_divides] THEN simp[RING_PRODUCT_CLAUSES] THEN have `!s:X. s IN S ==> s IN x INSERT S` [IN_INSERT] THEN have `ring_divides r (ring_product r S (f:X->R)) a` [] THEN subgoal `ring_coprime r (f x,ring_product r S (f:X->R))` THENL [ specialize_assuming[ `r:R ring`; `f:X->R`; `f(x:X):R`; `S:X->bool` ]ring_coprime_product THEN qed[] ; pass ] THEN qed[RING_DIVIDES_MUL] ] );; (* compare product_coprime_primes_divides below *) let ring_product_divides_if_coprime = prove(` !(r:R ring) f:X->R a S. (UFD r \/ integral_domain r /\ bezout_ring r) ==> a IN ring_carrier r ==> FINITE S ==> (!s. s IN S ==> ring_divides r (f s) a) ==> (!s t. s IN S ==> t IN S ==> ~(s = t) ==> ring_coprime r (f s,f t)) ==> ring_divides r (ring_product r S f) a `, simp[ring_product_divides_if_coprime_waterfall] );; let ring_coprime_lpow = prove(` !(r:R ring) a b n. (UFD r \/ integral_domain r /\ bezout_ring r) ==> ring_coprime r (a,b) ==> ring_coprime r ((ring_pow r a n),b) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ qed[RING_POW_0;ring_coprime_1;ring_coprime] ; rw[ARITH_RULE `SUC n = n+1`] THEN intro THEN have `a:R IN ring_carrier r` [ring_coprime] THEN simp[RING_POW_ADD;RING_POW_1] THEN have `b:R IN ring_carrier r` [ring_coprime] THEN qed[RING_COPRIME_LMUL;RING_POW] ] );; let ring_coprime_rpow = prove(` !(r:R ring) a b n. (UFD r \/ integral_domain r /\ bezout_ring r) ==> ring_coprime r (a,b) ==> ring_coprime r (a,(ring_pow r b n)) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ qed[RING_POW_0;ring_coprime_1;ring_coprime] ; rw[ARITH_RULE `SUC n = n+1`] THEN intro THEN have `b:R IN ring_carrier r` [ring_coprime] THEN simp[RING_POW_ADD;RING_POW_1] THEN have `a:R IN ring_carrier r` [ring_coprime] THEN qed[RING_COPRIME_RMUL;RING_POW] ] );; let ring_coprime_lrpow = prove(` !(r:R ring) a b m n. (UFD r \/ integral_domain r /\ bezout_ring r) ==> ring_coprime r (a,b) ==> ring_coprime r ((ring_pow r a m),(ring_pow r b n)) `, qed[ring_coprime_lpow;ring_coprime_rpow] );; let ring_coprime_associates_prime = prove(` !(r:R ring) p q. integral_domain r ==> ring_prime r p ==> ring_prime r q ==> ( ring_coprime r (p,q) <=> ~ring_associates r p q ) `, intro THEN splitiff THENL [ intro THEN have `ring_associates r p (p:R)` [RING_ASSOCIATES_REFL;ring_prime] THEN have `ring_coprime r (p,p:R)` [RING_ASSOCIATES_COPRIME] THEN have `ring_unit r (p:R)` [RING_COPRIME_REFL] THEN qed[ring_prime] ; intro THEN proven_if `ring_coprime r (p,q:R)` [] THEN have `q:R IN ring_carrier r` [ring_prime] THEN have `ring_divides r (p:R) q` [INTEGRAL_DOMAIN_PRIME_DIVIDES_OR_COPRIME] THEN have `p:R IN ring_carrier r` [ring_prime] THEN have `ring_divides r (q:R) p` [INTEGRAL_DOMAIN_PRIME_DIVIDES_OR_COPRIME;RING_COPRIME_SYM] THEN qed[ring_associates] ] );; let ring_divides_associates_prime = prove(` !(r:R ring) p q. integral_domain r ==> ring_prime r p ==> ring_prime r q ==> ( ring_divides r p q <=> ring_associates r p q ) `, intro THEN splitiff THENL [ intro THEN have `~ring_unit r (p:R)` [ring_prime] THEN have `ring_irreducible r (q:R)` [INTEGRAL_DOMAIN_PRIME_IMP_IRREDUCIBLE] THEN qed[RING_NONUNIT_DIVIDES_IRREDUCIBLE] ; qed[ring_associates] ] );; let ring_divides_product = prove(` !(r:R ring) S t:X f. FINITE S ==> t IN S ==> f t IN ring_carrier r ==> ring_divides r (f t) (ring_product r S f) `, simp[ring_product_delete] THEN rw[ring_divides] THEN qed[RING_PRODUCT;RING_MUL] );; let ring_sum_restrict_subset = prove(` !(r:R ring) S:X->bool U p. S SUBSET U ==> ring_sum r U (\s. if s IN S then p s else ring_0 r) = ring_sum r S p `, intro THEN specialize[ `S:X->bool`; `r:R ring`; `U:X->bool`; `p:X->R` ]RING_SUM_RESTRICT_SET THEN set_fact `S SUBSET U:X->bool ==> {x | x IN U /\ S x} = S` THEN rw[IN] THEN qed[] );; let ring_sum_range_add_1 = prove(` !(r:R ring) n c:num->R. c n IN ring_carrier r ==> ring_sum r (range(n+1)) c = ring_add r (c n) (ring_sum r (range n) c) `, intro THEN have `FINITE(range(n+1))` [finite_range] THEN have `n IN range(n+1)` [range_lt;ARITH_RULE `n < n+1`] THEN specialize[ `r:R ring`; `range(n+1)`; `c:num->R`; `n:num` ]ring_sum_delete2 THEN qed[range_add_1_delete_refl] );; let ring_sum_range_add_1_sub = prove(` !(r:R ring) n c:num->R. c n IN ring_carrier r ==> ring_sum r (range n) c = ring_sub r (ring_sum r (range(n+1)) c) (c n) `, intro THEN simp[ring_sum_range_add_1] THEN have `ring_sum r (range n) c:R IN ring_carrier r` [RING_SUM] THEN RING_TAC );; let ring_exp_sum = prove(` !(a:A ring) (b:B ring) E:A->B f:X->A S:X->bool. FINITE S ==> (!x:A. x IN ring_carrier a ==> E x IN ring_carrier b) ==> E(ring_0 a) = ring_1 b ==> (!x y. E(ring_add a x y) = ring_mul b (E x) (E y)) ==> (!x. x IN S ==> f x IN ring_carrier a) ==> E(ring_sum a S f) = ring_product b S (E o f) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;RING_PRODUCT_CLAUSES] THEN qed[] ; have `x:X IN x INSERT S` [IN_INSERT] THEN have `f(x:X):A IN ring_carrier a` [SUBSET;ring_divides] THEN simp[RING_PRODUCT_CLAUSES] THEN rw[o_THM] THEN simp[RING_SUM_CLAUSES] THEN ASM SET_TAC[RING_SUM] ] );; (* ===== ring_hasQ: contains copy of Q *) let ring_hasQ = new_definition ` ring_hasQ (r:R ring) <=> ( ring_char r = 0 /\ (!n. ~(n = 0) ==> ring_unit r (ring_of_num r n)) ) `;; let ring_hasQ_neg = prove(` !(r:R ring) n. ring_hasQ r ==> ~(n = 0) ==> ring_unit r (ring_neg r (ring_of_num r n)) `, qed[RING_UNIT_NEG;ring_hasQ] );; (* ===== numpreimages *) (* XXX: inherits weirdness of CARD in infinite case *) let numpreimages = new_definition ` numpreimages (f:X->Y) (S:X->bool) y = CARD {x | x IN S /\ f x = y} `;; let image_numpreimages = prove(` !S:X->bool f:X->Y. FINITE S ==> IMAGE f S = {y | ~(numpreimages f S y = 0)} `, rw[extension_z;in_image_vw;IN_ELIM_THM;numpreimages] THEN intro THEN specialize[`S:X->bool`;`\x:X. f x = z:Y`]FINITE_RESTRICT THEN simp[CARD_EQ_0] THEN rw[EXTENSION;EMPTY;IN_ELIM_THM] THEN qed[] );; (* warmup for numpreimages_permutes_o_perm *) let numpreimages_o_permutes = prove(` !S:X->bool f:X->Y g:X->X. FINITE S ==> g permutes S ==> numpreimages (f o g) S = numpreimages f S `, rw[numpreimages;o_THM;fun_eq_thm_v] THEN intro THEN have `!s t:X. s IN {x | x IN S /\ f (g x:X) = v:Y} ==> t IN {x | x IN S /\ f (g x) = v} ==> g s = g t:X ==> s = t` [PERMUTES_INJECTIVE] THEN specialize[ `S:X->bool`; `\x:X. f (g x:X) = v:Y` ]FINITE_RESTRICT THEN specialize[ `g:X->X`; `{x:X | x IN S /\ f (g x:X) = v:Y}` ]CARD_IMAGE_INJ THEN subgoal `IMAGE g {x:X | x IN S /\ f(g x:X) = v:Y} = {x | x IN S /\ f x = v}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM] THEN qed[permutes] ; pass ] THEN qed[] );; let numpreimages_permutes_o_perm = prove(` !A:X->bool f:X->X g:(X->X)->Y. FINITE A ==> f permutes A ==> numpreimages (\i:X->X. g(f o i)) (perm A) = numpreimages g (perm A) `, rw[numpreimages;o_THM;fun_eq_thm_v] THEN intro THEN have `FINITE(perm(A:X->bool))` [finite_perm] THEN subgoal `IMAGE (\i. f o i) {i:X->X | i IN perm A /\ g (f o i) = v:Y} = {i | i IN perm A /\ g i = v}` THENL [ once_rw[in_image_vw] THEN rw[EXTENSION;IN_IMAGE;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THENL [ qed[IN;perm;PERMUTES_COMPOSE] ; simp[] ] ; intro THEN witness `inverse f:X->X o x:X->X` THEN have `inverse f:X->X permutes A` [PERMUTES_INVERSE] THEN have `f o inverse f:X->X = I` [PERMUTES_INVERSES_o] THEN subgoal `x:X->X = (f:X->X) o inverse f o x` THENL [ rw[o_ASSOC] THEN qed[I_O_ID] ; pass ] THEN qed[PERMUTES_COMPOSE;IN;perm] ] ; pass ] THEN subgoal `!x y:X->X. x IN {i | i IN perm A /\ g (f o i) = v:Y} ==> y IN {i | i IN perm A /\ g (f o i) = v} ==> f o x = (f:X->X) o y ==> x = y` THENL [ rw[IN_ELIM_THM] THEN have `inverse f o f:X->X = I` [PERMUTES_INVERSES_o] THEN subgoal `!x:X->X. x = inverse f o (f:X->X) o x` THENL [ rw[o_ASSOC] THEN qed[I_O_ID] ; pass ] THEN qed[] ; pass ] THEN specialize[ `perm(A:X->bool)`; `\i:X->X. g(f:X->X o i) = v:Y` ]FINITE_RESTRICT THEN specialize[ `\i:X->X. f:X->X o i`; `{i | i IN perm A /\ g(f:X->X o i) = v:Y}` ]CARD_IMAGE_INJ THEN qed[] );; let ring_sum_fiber_o = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R y. FINITE {x | x IN S /\ f x = y} ==> g y IN ring_carrier r ==> ring_sum r {x | x IN S /\ f x = y} (g o f) = ring_mul r (ring_of_num r (numpreimages f S y)) (g y) `, intro THEN subgoal `ring_sum r {x:X | x IN S /\ f x = y:Y} (g o f) = ring_sum r {x | x IN S /\ f x = y} (\x. g y):R` THENL [ sufficesby RING_SUM_EQ THEN rw[o_THM;BETA_THM;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[ring_sum_const] THEN rw[numpreimages] );; let ring_product_fiber_o = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R y. FINITE {x | x IN S /\ f x = y} ==> g y IN ring_carrier r ==> ring_product r {x | x IN S /\ f x = y} (g o f) = ring_pow r (g y) (numpreimages f S y) `, intro THEN subgoal `ring_product r {x:X | x IN S /\ f x = y:Y} (g o f) = ring_product r {x | x IN S /\ f x = y} (\x. g y):R` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[o_THM;BETA_THM;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[ring_product_const] THEN rw[numpreimages] );; let ring_sum_o = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R. FINITE S ==> (!y. y IN IMAGE f S ==> g y IN ring_carrier r) ==> ring_sum r S (g o f) = ring_sum r (IMAGE f S) (\y. ring_mul r (ring_of_num r (numpreimages f S y)) (g y)) `, intro THEN specialize[ `r:R ring`; `f:X->Y`; `g:Y->R o f:X->Y`; `S:X->bool` ]RING_SUM_IMAGE_GEN THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN specialize[`S:X->bool`;`\x:X. f x = a:Y`]FINITE_RESTRICT THEN qed[ring_sum_fiber_o] );; let ring_sum_o_v2 = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R. FINITE S ==> (!y. y IN IMAGE f S ==> g y IN ring_carrier r) ==> ring_sum r S (g o f) = ring_sum r {y | ~(numpreimages f S y = 0)} (\y. ring_mul r (ring_of_num r (numpreimages f S y)) (g y)) `, simp[GSYM image_numpreimages] THEN simp[ring_sum_o] );; let ring_sum_o_v3 = prove(` !(r:R ring) S:X->bool f:X->Y U:Y->bool g:Y->R. FINITE S ==> (!y. y IN U ==> g y IN ring_carrier r) ==> IMAGE f S SUBSET U ==> ring_sum r S (g o f) = ring_sum r U (\y. ring_mul r (ring_of_num r (numpreimages f S y)) (g y)) `, intro THEN have `!y:Y. y IN IMAGE (f:X->Y) S ==> g y:R IN ring_carrier r` [SUBSET] THEN simp[ring_sum_o_v2] THEN intro THEN simp[GSYM image_numpreimages] THEN specialize[ `r:R ring`; `IMAGE (f:X->Y) S`; `U:Y->bool`; `\y. ring_mul r (ring_of_num r (numpreimages (f:X->Y) S y)) (g y):R` ](GSYM ring_sum_restrict_subset) THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN case `numpreimages (f:X->Y) S a = 0` THENL [ subgoal `~(a IN IMAGE (f:X->Y) S)` THENL [ simp[image_numpreimages;IN_ELIM_THM] ; pass ] THEN simp[RING_OF_NUM_0;RING_MUL_LZERO] ; pass ] THEN subgoal `a IN IMAGE (f:X->Y) S` THENL [ simp[image_numpreimages;IN_ELIM_THM] ; pass ] THEN qed[] );; let ring_product_o = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R. FINITE S ==> (!y. y IN IMAGE f S ==> g y IN ring_carrier r) ==> ring_product r S (g o f) = ring_product r (IMAGE f S) (\y. ring_pow r (g y) (numpreimages f S y)) `, intro THEN specialize[ `r:R ring`; `f:X->Y`; `g:Y->R o f:X->Y`; `S:X->bool` ]RING_PRODUCT_IMAGE_GEN THEN simp[] THEN sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN specialize[`S:X->bool`;`\x:X. f x = a:Y`]FINITE_RESTRICT THEN qed[ring_product_fiber_o] );; let ring_product_o_v2 = prove(` !(r:R ring) S:X->bool f:X->Y g:Y->R. FINITE S ==> (!y. y IN IMAGE f S ==> g y IN ring_carrier r) ==> ring_product r S (g o f) = ring_product r {y | ~(numpreimages f S y = 0)} (\y. ring_pow r (g y) (numpreimages f S y)) `, simp[GSYM image_numpreimages] THEN simp[ring_product_o] );; (* ===== squarefree elements of rings *) (* XXX: maybe should include a IN ring_carrier r *) (* XXX: maybe should exclude 0, or exclude non-injective *) let ring_squarefree = new_definition ` ring_squarefree(r:R ring) a <=> (!b. b IN ring_carrier r ==> ring_divides r a (ring_mul r b b) ==> ring_divides r a b ) `;; let not_squarefree_if_divisible_by_square = prove(` !(r:R ring) a b. integral_domain r ==> ~(a = ring_0 r) ==> b IN ring_carrier r ==> ~(ring_unit r b) ==> ring_divides r (ring_mul r b b) a ==> ~(ring_squarefree r a) `, intro THEN have `a IN ring_carrier(r:R ring)` [ring_divides] THEN have `b IN ring_carrier(r:R ring)` [ring_unit] THEN choose `q:R` `q IN ring_carrier(r:R ring) /\ a = ring_mul r (ring_mul r b b) q` [ring_divides] THEN have `q IN ring_carrier(r:R ring)` [] THEN have `a = ring_mul(r:R ring) (ring_mul r b b) q` [] THEN recall(RING_RULE `a = ring_mul(r:R ring) (ring_mul r b b) q ==> ring_mul r (ring_mul r b q) (ring_mul r b q) = ring_mul r a q`) THEN have `ring_divides(r:R ring) a (ring_mul r (ring_mul r b q) (ring_mul r b q))` [ring_divides;RING_MUL] THEN have `ring_divides(r:R ring) a (ring_mul r b q)` [ring_squarefree;RING_MUL] THEN choose `u:R` `u IN ring_carrier(r:R ring) /\ ring_mul r b q = ring_mul r a u` [ring_divides] THEN have `u IN ring_carrier(r:R ring)` [] THEN have `ring_mul(r:R ring) b q = ring_mul r a u` [] THEN recall(RING_RULE `a = ring_mul(r:R ring) (ring_mul r b b) q /\ ring_mul r b q = ring_mul r a u ==> ring_mul r a (ring_1 r) = ring_mul r a (ring_mul r b u)`) THEN have `ring_1(r:R ring) = ring_mul r b u` [INTEGRAL_DOMAIN_MUL_LCANCEL;RING_1;RING_MUL] THEN qed[ring_unit] );; let product_coprime_primes_divides_waterfall = prove(` !(r:R ring). !P. FINITE P ==> P SUBSET ring_carrier r ==> !b. b IN ring_carrier r ==> (!p. p IN P ==> ring_prime r p) ==> (!p q. p IN P ==> q IN P ==> ring_divides r p q ==> p = q) ==> (!p. p IN P ==> ring_divides r p b) ==> ring_divides r (ring_product r P I) b `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ simp[RING_PRODUCT_CLAUSES] THEN qed[RING_DIVIDES_1] ; have `(x:R) IN x INSERT P` [IN_INSERT] THEN have `(x:R) IN ring_carrier r` [SUBSET] THEN have `I (x:R) = x` [I_THM] THEN have `I (x:R) IN ring_carrier r` [] THEN have `ring_product(r:R ring) (x INSERT P) I = ring_mul r x (ring_product r P I)` [RING_PRODUCT_CLAUSES] THEN simp[RING_PRODUCT_CLAUSES] THEN have `ring_divides r (x:R) b` [] THEN choose `c:R` `c IN ring_carrier r /\ (b:R) = ring_mul r x c` [ring_divides] THEN subgoal `!q:R. q IN P ==> ring_divides r q c` THENL [ intro THEN have `q:R IN x INSERT P` [IN_INSERT] THEN have `q:R IN ring_carrier r` [SUBSET] THEN have `ring_divides r (q:R) b` [] THEN have `ring_divides r (q:R) (ring_mul r x c)` [] THEN have `~(ring_divides r (q:R) x)` [] THEN qed[ring_prime] ; pass ] THEN have `!q:R. q IN P ==> ring_prime r q` [IN_INSERT] THEN have `!p q:R. p IN P ==> q IN P ==> ring_divides r p q ==> p = q` [IN_INSERT] THEN have `P SUBSET (P:R->bool)` [SUBSET_REFL] THEN have `P SUBSET ring_carrier(r:R ring)` [SUBSET_INSERT;SUBSET_TRANS] THEN have `c:R IN ring_carrier r` [] THEN specialize[`c:R`](UNDISCH(know `P SUBSET ring_carrier(r:R ring) ==> (!b. b IN ring_carrier r ==> (!p. p IN P ==> ring_prime r p) ==> (!p q. p IN P ==> q IN P ==> ring_divides r p q ==> p = q) ==> (!p. p IN P ==> ring_divides r p b) ==> ring_divides r (ring_product r P I) b)`)) THEN qed[RING_DIVIDES_REFL;RING_DIVIDES_MUL2] ] );; let product_coprime_primes_divides = prove(` !(r:R ring) P b. P SUBSET ring_carrier r ==> FINITE P ==> b IN ring_carrier r ==> (!p. p IN P ==> ring_prime r p) ==> (!p q. p IN P ==> q IN P ==> ring_divides r p q ==> p = q) ==> (!p. p IN P ==> ring_divides r p b) ==> ring_divides r (ring_product r P I) b `, intro THEN ASSUME_TAC(ISPECL[`b:R`](UNDISCH_ALL(ISPECL[`r:R ring`;`P:R->bool`]product_coprime_primes_divides_waterfall))) THEN qed[] );; let ring_squarefree_if_product_coprime_primes = prove(` !(r:R ring) P. P SUBSET ring_carrier r ==> FINITE P ==> (!p. p IN P ==> ring_prime r p) ==> (!p q. p IN P ==> q IN P ==> ring_divides r p q ==> p = q) ==> ring_squarefree r (ring_product r P I) `, rw[ring_squarefree] THEN intro THEN subgoal `!p:R. p IN P ==> ring_divides r p (ring_product r P I)` THENL [ intro THEN have `(p:R) IN ring_carrier r` [SUBSET] THEN have `I (p:R) = p` [I_THM] THEN have `I (p:R) IN ring_carrier r` [] THEN have `ring_product(r:R ring) {p} I = p` [RING_PRODUCT_SING] THEN have `FINITE {p:R}` [FINITE_SING] THEN have `{p:R} SUBSET P` [SUBSET;IN_SING] THEN qed[RING_DIVIDES_PRODUCT_SUBSET] ; pass ] THEN have `!p:R. p IN P ==> ring_divides r p (ring_mul r b b)` [RING_DIVIDES_TRANS;RING_MUL] THEN have `!p:R. p IN P ==> ring_divides r p b` [ring_prime] THEN specialize[`r:R ring`;`P:R->bool`;`b:R`]product_coprime_primes_divides THEN qed[] );; let ring_squarefree_if_product_coprime_primes_indexed = prove(` !(r:R ring) S (f:X->R). FINITE S ==> (!s:X. s IN S ==> f s IN ring_carrier r) ==> (!s:X. s IN S ==> ring_prime r (f s)) ==> (!s t. s IN S ==> t IN S ==> ring_divides r (f s) (f t) ==> s = t) ==> ring_squarefree r (ring_product r S f) `, intro THEN def `P:R->bool` `IMAGE (f:X->R) S` THEN have `FINITE (P:R->bool)` [FINITE_IMAGE] THEN havetac `(P:R->bool) SUBSET ring_carrier r` (rw[SUBSET;EXTENSION] THEN qed[IN_IMAGE]) THEN havetac `!p:R. p IN P ==> ring_prime r p` (rw[SUBSET;EXTENSION] THEN qed[IN_IMAGE]) THEN havetac `!p q:R. p IN P ==> q IN P ==> ring_divides r p q ==> p = q` (rw[SUBSET;EXTENSION] THEN qed[IN_IMAGE]) THEN specialize[`r:R ring`;`P:R->bool`]ring_squarefree_if_product_coprime_primes THEN have `!s t:X. s IN S ==> t IN S ==> f s = f t:R ==> s = t` [RING_DIVIDES_REFL] THEN specialize[`r:R ring`;`f:X->R`;`I:R->R`;`S:X->bool`]RING_PRODUCT_IMAGE THEN qed[I_O_ID] );; let ring_squarefree_if_unit = prove(` !(r:R ring) p. ring_unit r p ==> ring_squarefree r p `, rw[ring_squarefree] THEN qed[RING_UNIT_DIVIDES_ANY] );; let ring_squarefree_if_prime = prove(` !(r:R ring) p. ring_prime r p ==> ring_squarefree r p `, intro THEN have `p:R IN ring_carrier r` [ring_prime] THEN have `{p:R} SUBSET ring_carrier r` [SUBSET;IN_SING] THEN have `FINITE {p:R}` [FINITE_SING] THEN have `!q:R. q IN {p} ==> ring_prime r q` [IN_SING] THEN have `!q x:R. q IN {p} ==> x IN {p} ==> ring_divides r q x ==> q = x` [IN_SING] THEN specialize[`r:R ring`;`{p:R}`]ring_squarefree_if_product_coprime_primes THEN qed[RING_PRODUCT_SING;I_THM] );; let ring_coprime_if_unit = prove(` !(r:R ring) a b. ring_unit r a ==> b IN ring_carrier r ==> (ring_coprime r (a,b) /\ ring_coprime r (b,a)) `, rw[ring_coprime] THEN qed[ring_unit;RING_UNIT_DIVISOR] );; let ring_product_divides_factor_by_factor = prove(` !(r:R ring) (f:X->R) (g:X->R) S. FINITE S ==> (!s:X. s IN S ==> ring_divides r (f s) (g s)) ==> ring_divides r (ring_product r S f) (ring_product r S g) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN simp[RING_PRODUCT_CLAUSES] THEN intro THENL [ qed[RING_DIVIDES_1;RING_1] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN have `!s:X. s IN S ==> f s IN ring_carrier(r:R ring)` [ring_divides] THEN have `!s:X. s IN S ==> g s IN ring_carrier(r:R ring)` [ring_divides] THEN have `f(x:X) IN ring_carrier(r:R ring)` [ring_divides] THEN have `g(x:X) IN ring_carrier(r:R ring)` [ring_divides] THEN have `ring_divides (r:R ring) (f(x:X)) (g(x:X))` [] THEN qed[RING_DIVIDES_MUL2] ] );; let ring_sum_delta_delta = prove(` !(r:R ring) S t u a b. FINITE S ==> t IN S ==> u IN S ==> a IN ring_carrier r ==> b IN ring_carrier r ==> ring_sum r S (\s:X. if s = t then a else if s = u then b else ring_0 r) = if t = u then a else ring_add r a b `, intro THEN have `FINITE(S DELETE (t:X))` [FINITE_DELETE] THEN def `f:X->R` `\s:X. if s = t then a else if s = u then b else ring_0 (r:R ring)` THEN have `f(t:X) IN ring_carrier(r:R ring)` [] THEN set_fact `~((t:X) IN (S DELETE t))` THEN specialize[`r:R ring`;`t:X`;`f:X->R`;`S DELETE (t:X)`](CONJUNCT2 RING_SUM_CLAUSES) THEN have `ring_sum(r:R ring) ((t:X) INSERT (S DELETE t)) f = ring_add r a (ring_sum r (S DELETE t) f)` [] THEN set_fact `t IN S ==> (t:X) INSERT (S DELETE t) = S` THEN have `ring_sum(r:R ring) S (f:X->R) = ring_add r a (ring_sum r (S DELETE t) f)` [] THEN case `t = u:X` THENL [ set_fact `t = u ==> !s:X. (s:X) IN S DELETE t ==> ~(s = t) /\ ~(s = u)` THEN have `!s:X. (s:X) IN S DELETE t ==> f(s) = ring_0(r:R ring)` [] THEN specialize[`r:R ring`;`f:X->R`;`\s:X. ring_0(r:R ring)`;`S DELETE (t:X)`]RING_SUM_EQ THEN have `ring_sum(r:R ring) (S DELETE (t:X)) f = ring_sum r (S DELETE t) (\i. ring_0 r)` [] THEN have `ring_sum(r:R ring) (S DELETE (t:X)) f = ring_0 r` [RING_SUM_0] THEN have `ring_sum(r:R ring) S (f:X->R) = ring_add r a (ring_0 r)` [] THEN qed[RING_ADD_RZERO] ; pass ] THEN have `FINITE(S DELETE (t:X) DELETE u)` [FINITE_DELETE] THEN have `f(u:X) IN ring_carrier(r:R ring)` [] THEN set_fact `~((u:X) IN (S DELETE t DELETE u))` THEN specialize[`r:R ring`;`u:X`;`f:X->R`;`S DELETE (t:X) DELETE (u:X)`](CONJUNCT2 RING_SUM_CLAUSES) THEN have `ring_sum(r:R ring) ((u:X) INSERT (S DELETE t DELETE u)) f = ring_add r (f u) (ring_sum r (S DELETE t DELETE u) f)` [] THEN have `f(u:X) = b:R` [] THEN have `ring_sum(r:R ring) ((u:X) INSERT (S DELETE t DELETE u)) f = ring_add r (if t = u then ring_0 r else b) (ring_sum r (S DELETE t DELETE u) f)` [] THEN set_fact `u IN S ==> ~(t = u) ==> (u:X) INSERT (S DELETE t DELETE u) = S DELETE t` THEN have `ring_sum(r:R ring) (S DELETE (t:X)) f = ring_add r (if t = u then ring_0 r else b) (ring_sum r (S DELETE t DELETE u) f)` [] THEN set_fact `!s:X. (s:X) IN S DELETE t DELETE u ==> ~(s = t) /\ ~(s = u)` THEN have `!s:X. ~(s = t) /\ ~(s = u) ==> f(s) = ring_0(r:R ring)` [] THEN have `!s:X. (s:X) IN S DELETE t DELETE u ==> f(s) = ring_0(r:R ring)` [] THEN specialize[`r:R ring`;`f:X->R`;`\s:X. ring_0(r:R ring)`;`S DELETE t DELETE (u:X)`]RING_SUM_EQ THEN have `ring_sum(r:R ring) (S DELETE (t:X) DELETE u) f = ring_sum r (S DELETE t DELETE u) (\i. ring_0 r)` [] THEN have `ring_sum(r:R ring) (S DELETE (t:X) DELETE u) f = ring_0 r` [RING_SUM_0] THEN have `ring_sum(r:R ring) (S DELETE (t:X)) f = ring_add r b (ring_0 r)` [] THEN have `ring_sum(r:R ring) (S DELETE (t:X)) f = b` [RING_ADD_RZERO] THEN qed[] );; let ring_product_delta_delta = prove(` !(r:R ring) S t u a b. FINITE S ==> t IN S ==> u IN S ==> a IN ring_carrier r ==> b IN ring_carrier r ==> ring_product r S (\s:X. if s = t then a else if s = u then b else ring_1 r) = if t = u then a else ring_mul r a b `, intro THEN have `FINITE(S DELETE (t:X))` [FINITE_DELETE] THEN def `f:X->R` `\s:X. if s = t then a else if s = u then b else ring_1 (r:R ring)` THEN have `f(t:X) IN ring_carrier(r:R ring)` [] THEN set_fact `~((t:X) IN (S DELETE t))` THEN specialize[`r:R ring`;`t:X`;`f:X->R`;`S DELETE (t:X)`](CONJUNCT2 RING_PRODUCT_CLAUSES) THEN have `ring_product(r:R ring) ((t:X) INSERT (S DELETE t)) f = ring_mul r a (ring_product r (S DELETE t) f)` [] THEN set_fact `t IN S ==> (t:X) INSERT (S DELETE t) = S` THEN have `ring_product(r:R ring) S (f:X->R) = ring_mul r a (ring_product r (S DELETE t) f)` [] THEN case `t = u:X` THENL [ set_fact `t = u ==> !s:X. (s:X) IN S DELETE t ==> ~(s = t) /\ ~(s = u)` THEN have `!s:X. (s:X) IN S DELETE t ==> f(s) = ring_1(r:R ring)` [] THEN specialize[`r:R ring`;`f:X->R`;`\s:X. ring_1(r:R ring)`;`S DELETE (t:X)`]RING_PRODUCT_EQ THEN have `ring_product(r:R ring) (S DELETE (t:X)) f = ring_product r (S DELETE t) (\i. ring_1 r)` [] THEN have `ring_product(r:R ring) (S DELETE (t:X)) f = ring_1 r` [RING_PRODUCT_1] THEN have `ring_product(r:R ring) S (f:X->R) = ring_mul r a (ring_1 r)` [] THEN qed[RING_MUL_RID] ; pass ] THEN have `FINITE(S DELETE (t:X) DELETE u)` [FINITE_DELETE] THEN have `f(u:X) IN ring_carrier(r:R ring)` [] THEN set_fact `~((u:X) IN (S DELETE t DELETE u))` THEN specialize[`r:R ring`;`u:X`;`f:X->R`;`S DELETE (t:X) DELETE (u:X)`](CONJUNCT2 RING_PRODUCT_CLAUSES) THEN have `ring_product(r:R ring) ((u:X) INSERT (S DELETE t DELETE u)) f = ring_mul r (f u) (ring_product r (S DELETE t DELETE u) f)` [] THEN have `f(u:X) = b:R` [] THEN have `ring_product(r:R ring) ((u:X) INSERT (S DELETE t DELETE u)) f = ring_mul r (if t = u then ring_1 r else b) (ring_product r (S DELETE t DELETE u) f)` [] THEN set_fact `u IN S ==> ~(t = u) ==> (u:X) INSERT (S DELETE t DELETE u) = S DELETE t` THEN have `ring_product(r:R ring) (S DELETE (t:X)) f = ring_mul r (if t = u then ring_1 r else b) (ring_product r (S DELETE t DELETE u) f)` [] THEN set_fact `!s:X. (s:X) IN S DELETE t DELETE u ==> ~(s = t) /\ ~(s = u)` THEN have `!s:X. ~(s = t) /\ ~(s = u) ==> f(s) = ring_1(r:R ring)` [] THEN have `!s:X. (s:X) IN S DELETE t DELETE u ==> f(s) = ring_1(r:R ring)` [] THEN specialize[`r:R ring`;`f:X->R`;`\s:X. ring_1(r:R ring)`;`S DELETE t DELETE (u:X)`]RING_PRODUCT_EQ THEN have `ring_product(r:R ring) (S DELETE (t:X) DELETE u) f = ring_product r (S DELETE t DELETE u) (\i. ring_1 r)` [] THEN have `ring_product(r:R ring) (S DELETE (t:X) DELETE u) f = ring_1 r` [RING_PRODUCT_1] THEN have `ring_product(r:R ring) (S DELETE (t:X)) f = ring_mul r b (ring_1 r)` [] THEN have `ring_product(r:R ring) (S DELETE (t:X)) f = b` [RING_MUL_RID] THEN qed[] );; let square_divides_product_if_factor_divides_factor = prove(` !(r:R ring) S (f:X->R) t u. FINITE S ==> (!s:X. s IN S ==> f s IN ring_carrier r) ==> ~(t = u) ==> t IN S ==> u IN S ==> ring_divides r (f t) (f u) ==> ring_divides r (ring_mul r (f t) (f t)) (ring_product r S f) `, intro THEN def `q:X->R` `\s:X. if s = t then f t else if s = u then f t else ring_1(r:R ring)` THEN subgoal `!s:X. s IN S ==> ring_divides(r:R ring) (q s) (f s)` THENL [ intro THEN have `f(s:X) IN ring_carrier(r:R ring)` [] THEN case `s = t:X` THENL [ have `ring_divides (r:R ring) (f(t:X)) (f(t:X))` [RING_DIVIDES_REFL] THEN qed[] ; case `s = u:X` THENL [ qed[] ; qed[RING_DIVIDES_1] ] ] ; pass ] THEN specialize[`r:R ring`;`q:X->R`;`f:X->R`;`S:X->bool`]ring_product_divides_factor_by_factor THEN have `f(t:X) IN ring_carrier(r:R ring)` [] THEN specialize[`r:R ring`;`S:X->bool`;`t:X`;`u:X`;`f(t:X):R`;`f(t:X):R`]ring_product_delta_delta THEN qed[] );; let prime_divides_prime_and = prove(` !(r:R ring) p q a. UFD r ==> ring_prime r p ==> ring_prime r q ==> ring_divides r p q ==> ring_divides r p a ==> ring_divides r q a `, intro THEN have `~(ring_unit(r:R ring) p)` [ring_prime] THEN have `ring_irreducible(r:R ring) q` [UFD_IRREDUCIBLE_EQ_PRIME] THEN have `ring_associates(r:R ring) p q` [RING_NONUNIT_DIVIDES_IRREDUCIBLE] THEN have `ring_divides(r:R ring) q p` [ring_associates] THEN qed[RING_DIVIDES_TRANS] );; let ring_squarefree_associates = prove(` !(r:R ring) f g. ring_associates r f g ==> ring_squarefree r f ==> ring_squarefree r g `, rw[ring_squarefree] THEN qed[RING_ASSOCIATES_DIVIDES;RING_ASSOCIATES_REFL;RING_MUL] );; (* ===== power series and polynomials *) let poly_neg_subring = prove(` !(r:R ring) S (p:(V->num)->R). poly_neg (subring_generated r S) p = poly_neg r p `, rw[poly_neg;SUBRING_GENERATED] );; let poly_add_subring = prove(` !(r:R ring) S (p:(V->num)->R) q. poly_add (subring_generated r S) p q = poly_add r p q `, rw[poly_add;SUBRING_GENERATED] );; let poly_mul_subring = prove(` !(r:R ring) S (p:(V->num)->R) q. ring_powerseries (subring_generated r S) p ==> ring_powerseries (subring_generated r S) q ==> poly_mul (subring_generated r S) p q = poly_mul r p q `, rw[poly_mul;SUBRING_GENERATED] THEN rw[FUN_EQ_THM] THEN intro THEN subgoal `!(m1:V->num) (m2:V->num). ring_mul r (p m1) (q m2) IN ring_carrier(subring_generated r (S:R->bool))` THENL [ intro THEN have `(p:(V->num)->R) m1 IN ring_carrier(subring_generated(r:R ring) S)` [ring_powerseries] THEN have `(q:(V->num)->R) m2 IN ring_carrier(subring_generated(r:R ring) S)` [ring_powerseries] THEN specialize[`subring_generated(r:R ring) S`;`(p:(V->num)->R) m1`;`(q:(V->num)->R) m2`]RING_MUL THEN qed[SUBRING_GENERATED;CARRIER_SUBRING_GENERATED_SUBRING] ; pass ] THEN sufficesby ring_sum_subring_generated_v2 THEN simp[IN_ELIM_THM;LAMBDA_PAIR] );; let poly_sub_subring = prove(` !(r:R ring) S (p:(V->num)->R) q. poly_sub (subring_generated r S) p q = poly_sub r p q `, rw[poly_sub;RING_SUB_SUBRING_GENERATED] );; let poly_deg_subring = prove(` !(r:R ring) G (p:(V->num)->R). poly_deg (subring_generated r G) p = poly_deg r p `, rw[poly_deg;SUBRING_GENERATED] );; let poly_evaluate_subring = prove(` !(r:R ring) S (p:(V->num)->R) z. S subring_of r ==> ring_powerseries (subring_generated r S) p ==> (!v:V. z v IN S) ==> poly_evaluate r p z = poly_evaluate (subring_generated r S) p z `, rw[poly_evaluate;poly_extend] THEN intro THEN subgoal `!m:V->num. m IN {m | ~(p m = ring_0(r:R ring))} ==> ring_mul (subring_generated r S) (I (p m)) (ring_product (subring_generated r S) (monomial_vars m) (\i. ring_pow (subring_generated r S) (z i) (m i))) IN S` THENL [ rw[IN_ELIM_THM] THEN intro THEN have `ring_product (subring_generated r S) (monomial_vars m) (\i:V. ring_pow (subring_generated r S) (z i) (m i)) IN ring_carrier(subring_generated(r:R ring) S)` [RING_PRODUCT] THEN have `(p:(V->num)->R) m IN ring_carrier(subring_generated r S)` [ring_powerseries] THEN have `I ((p:(V->num)->R) m) IN ring_carrier(subring_generated r S)` [I_DEF] THEN have `ring_mul (subring_generated r S) (I ((p:(V->num)->R) m)) (ring_product (subring_generated r S) (monomial_vars m) (\i:V. ring_pow (subring_generated r S) (z i) (m i))) IN ring_carrier(subring_generated r S)` [RING_MUL] THEN qed[CARRIER_SUBRING_GENERATED_SUBRING] ; pass ] THEN simp[prove(`ring_0(subring_generated(r:R ring) S) = ring_0 r`,qed[SUBRING_GENERATED])] THEN simp[ring_sum_subring_generated] THEN sufficesby RING_SUM_EQ THEN intro THEN simp[SUBRING_GENERATED] THEN sufficesby (prove(`a:R = b ==> ring_mul r c a = ring_mul r c b`,qed[])) THEN simp[RING_POW_SUBRING_GENERATED] THEN sufficesby(GSYM ring_product_subring_generated) THEN simp[] THEN intro THEN qed[RING_POW;RING_POW_SUBRING_GENERATED;CARRIER_SUBRING_GENERATED_SUBRING] );; let poly_eval_subring = prove(` !(r:R ring) S p z. S subring_of r ==> ring_powerseries (subring_generated r S) p ==> z IN S ==> poly_eval r p z = poly_eval (subring_generated r S) p z `, rw[poly_eval] THEN qed[poly_evaluate_subring] );; let ring_powerseries_subring = prove(` !(r:R ring) G (p:(V->num)->R). ring_powerseries(subring_generated r G) p ==> ring_powerseries r p `, rw[ring_powerseries] THEN qed[RING_CARRIER_SUBRING_GENERATED_SUBSET;SUBRING_GENERATED;SUBSET] );; let ring_polynomial_if_subring = prove(` !(r:R ring) (p:(V->num)->R). ring_polynomial(subring_generated r G) p ==> ring_polynomial r p `, rw[ring_polynomial] THEN intro THENL [ qed[ring_powerseries_subring] ; set_fact_assuming `ring_0 (subring_generated r G) = ring_0(r:R ring) ==> {m:V->num | ~(p m = ring_0 r)} = {m | ~(p m = ring_0 (subring_generated r G))}` THEN qed[SUBRING_GENERATED] ] );; let deg_zero_ring = prove(` !(r:R ring) (p:(V->num)->R). ring_1 r = ring_0 r ==> ring_powerseries r p ==> poly_deg r p = 0 `, intro THEN have `!m:V->num. p m IN ring_carrier(r:R ring)` [ring_powerseries] THEN have `trivial_ring(r:R ring)` [TRIVIAL_RING_10] THEN have `!m:V->num. p m IN {ring_0(r:R ring)}` [trivial_ring] THEN set_fact `!m:V->num. p m IN {ring_0(r:R ring)} ==> p m = ring_0 r` THEN have `!m:V->num. p m = ring_0(r:R ring)` [] THEN have `(p:(V->num)->R) = poly_0 r` [POLY_0;FUN_EQ_THM] THEN qed[POLY_DEG_0] );; let poly_const_subring = prove(` !(r:R ring) S c. poly_const (subring_generated r S) c = poly_const r c:(V->num)->R `, rw[poly_const;SUBRING_GENERATED] );; let poly_0_subring = prove(` !(r:R ring) S. poly_0 (subring_generated r S) = poly_0 r:(V->num)->R `, rw[poly_0;poly_const;SUBRING_GENERATED] );; let poly_1_subring = prove(` !(r:R ring) S. poly_1 (subring_generated r S) = poly_1 r:(V->num)->R `, rw[poly_1;poly_const;SUBRING_GENERATED] );; let POWSER_MUL_0 = prove (`(!r (p:(V->num)->A). ring_powerseries r p ==> poly_mul r p (poly_0 r) = poly_0 r) /\ (!r (q:(V->num)->A). ring_powerseries r q ==> poly_mul r (poly_0 r) q = poly_0 r)`, REWRITE_TAC[ring_powerseries; poly_mul; poly_0; poly_const; COND_ID] THEN REPEAT STRIP_TAC THEN GEN_REWRITE_TAC I [FUN_EQ_THM] THEN REWRITE_TAC[] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC RING_SUM_EQ_0 THEN ASM_SIMP_TAC[FORALL_PAIR_THM; RING_MUL_LZERO; RING_MUL_RZERO]);; let powser_ring_carrier = prove(` !r:R ring. !s:V->bool. ring_carrier(powser_ring r s) = {p | ring_powerseries r p /\ poly_vars r p SUBSET s} `, qed[POWSER_RING] );; let poly_ring_carrier = prove(` !r:R ring. !s:V->bool. ring_carrier(poly_ring r s) = {p | ring_polynomial r p /\ poly_vars r p SUBSET s} `, qed[POLY_RING] );; let poly_in_full_ring = prove(` !(r:R ring) p:(V->num)->R. ring_polynomial r p <=> p IN ring_carrier(poly_ring r (:V)) `, intro THEN rw[POLY_RING;IN_ELIM_THM] THEN set_fact `poly_vars(r:R ring) p SUBSET (:V)` THEN qed[] );; let series_in_full_ring = prove(` !(r:R ring) p:(V->num)->R. ring_powerseries r p <=> p IN ring_carrier(powser_ring r (:V)) `, intro THEN rw[POWSER_RING;IN_ELIM_THM] THEN set_fact `poly_vars(r:R ring) p SUBSET (:V)` THEN qed[] );; let poly_sub_ldistrib_lemma = prove(` !(r:R ring) p1 p2 p3. p1 IN ring_carrier r ==> p2 IN ring_carrier r ==> p3 IN ring_carrier r ==> ring_mul r p1 (ring_add r p2 (ring_neg r p3)) = ring_add r (ring_mul r p1 p2) (ring_neg r (ring_mul r p1 p3)) `, RING_TAC );; let poly_sub_ldistrib = prove (`!r p1 p2 (p3:(V->num)->R). ring_powerseries r p1 ==> ring_powerseries r p2 ==> ring_powerseries r p3 ==> poly_mul r p1 (poly_sub r p2 p3) = poly_sub r (poly_mul r p1 p2) (poly_mul r p1 p3)`, intro THEN rw[POLY_SUB] THEN rw[prove(`poly_neg(r:R ring) = ring_neg(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_add(r:R ring) = ring_add(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_mul(r:R ring) = ring_mul(powser_ring r (:V))`,qed[POWSER_RING])] THEN have `p1 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN have `p2 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN have `p3 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN qed[poly_sub_ldistrib_lemma] );; let poly_sub_rdistrib_lemma = prove(` !(r:R ring) p1 p2 p3. p1 IN ring_carrier r ==> p2 IN ring_carrier r ==> p3 IN ring_carrier r ==> ring_mul r (ring_add r p2 (ring_neg r p3)) p1 = ring_add r (ring_mul r p2 p1) (ring_neg r (ring_mul r p3 p1)) `, RING_TAC );; let poly_sub_rdistrib = prove (`!r p1 p2 (p3:(V->num)->R). ring_powerseries r p1 ==> ring_powerseries r p2 ==> ring_powerseries r p3 ==> poly_mul r (poly_sub r p2 p3) p1 = poly_sub r (poly_mul r p2 p1) (poly_mul r p3 p1)`, intro THEN rw[POLY_SUB] THEN rw[prove(`poly_neg(r:R ring) = ring_neg(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_add(r:R ring) = ring_add(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_mul(r:R ring) = ring_mul(powser_ring r (:V))`,qed[POWSER_RING])] THEN have `p1 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN have `p2 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN have `p3 IN ring_carrier(powser_ring(r:R ring) (:V))` [series_in_full_ring] THEN qed[poly_sub_rdistrib_lemma] );; let poly_sub_0 = prove(` !(r:R ring) p:(V->num)->R. ring_powerseries r p ==> poly_sub r p (poly_0 r) = p `, intro THEN rw[POLY_SUB] THEN rw[prove(`poly_neg(r:R ring) = ring_neg(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_add(r:R ring) = ring_add(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_0(r:R ring) = ring_0(powser_ring r (:V))`,qed[POWSER_RING])] THEN set_fact `ring_powerseries r (p:(V->num)->R) ==> p IN {q | ring_powerseries r q /\ poly_vars r q SUBSET (:V)}` THEN specialize[`r:R ring`;`(:V)`](CONJUNCT1 POWSER_RING) THEN have `(p:(V->num)->R) IN ring_carrier(powser_ring(r:R ring) (:V))` [] THEN qed[RING_RULE `ring_add (r:R ring) p (ring_neg r (ring_0 r)) = p`] );; let poly_0_sub = prove(` !(r:R ring) p:(V->num)->R. ring_powerseries r p ==> poly_sub r (poly_0 r) p = poly_neg r p `, intro THEN rw[POLY_SUB] THEN rw[prove(`poly_neg(r:R ring) = ring_neg(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_add(r:R ring) = ring_add(powser_ring r (:V))`,qed[POWSER_RING])] THEN rw[prove(`poly_0(r:R ring) = ring_0(powser_ring r (:V))`,qed[POWSER_RING])] THEN set_fact `ring_powerseries r (p:(V->num)->R) ==> p IN {q | ring_powerseries r q /\ poly_vars r q SUBSET (:V)}` THEN specialize[`r:R ring`;`(:V)`](CONJUNCT1 POWSER_RING) THEN have `(p:(V->num)->R) IN ring_carrier(powser_ring(r:R ring) (:V))` [] THEN qed[RING_RULE `ring_add (r:R ring) (ring_0 r) (ring_neg r p) = ring_neg r p`] );; let poly_vars_empty = prove(` !(r:R ring) p:(V->num)->R. poly_vars r p = {} <=> p = poly_const r (p monomial_1) `, intro THEN splitiff THENL [ rw[poly_vars] THEN rw[EMPTY_UNIONS;monomial_vars;IN_ELIM_THM;poly_const;monomial_1;COND_ID] THEN intro THEN once_rw[FUN_EQ_THM] THEN rw[BETA_THM] THEN intro THEN proven_if `p(x:V->num) = ring_0(r:R ring)` [] THEN have `?m. ~(p m = ring_0(r:R ring)) /\ {i:V | ~(x i = 0)} = {i | ~(m i = 0)}` [] THEN have `{i:V | ~(x i = 0)} = {}` [] THEN subgoal `x = (\i:V. 0)` THENL [ rw[FUN_EQ_THM] THEN ASM SET_TAC[] ; pass ] THEN qed[] ; qed[POLY_VARS_CONST] ] );; let poly_neg_in_poly_ring = prove(` !(r:R ring) p:(V->num)->R S. p IN ring_carrier(poly_ring r S) ==> poly_neg r p IN ring_carrier(poly_ring r S) `, qed[RING_NEG;POLY_RING] );; let poly_add_in_poly_ring = prove(` !(r:R ring) p:(V->num)->R q S. p IN ring_carrier(poly_ring r S) ==> q IN ring_carrier(poly_ring r S) ==> poly_add r p q IN ring_carrier(poly_ring r S) `, qed[RING_ADD;POLY_RING] );; let poly_sub_in_poly_ring = prove(` !(r:R ring) p:(V->num)->R q S. p IN ring_carrier(poly_ring r S) ==> q IN ring_carrier(poly_ring r S) ==> poly_sub r p q IN ring_carrier(poly_ring r S) `, rw[POLY_SUB] THEN qed[poly_neg_in_poly_ring;poly_add_in_poly_ring] );; let poly_mul_in_poly_ring = prove(` !(r:R ring) p:(V->num)->R q S. p IN ring_carrier(poly_ring r S) ==> q IN ring_carrier(poly_ring r S) ==> poly_mul r p q IN ring_carrier(poly_ring r S) `, qed[RING_MUL;POLY_RING] );; let poly_var_o_permutes = prove(` !(r:R ring) v:V S f:V->V m. f permutes S ==> poly_var r v (m o f) = poly_var r (f v) m `, rw[poly_var;monomial_var] THEN rw[FUN_EQ_THM;o_THM] THEN rw[permutes] THEN qed[] );; let poly_neg_o_permutes = prove(` !(r:R ring) p:(V->num)->R (f:V->W) m. poly_neg r p (m o f) = poly_neg r (\m. p (m o f)) m `, rw[poly_neg] );; let poly_add_o_permutes = prove(` !(r:R ring) p:(V->num)->R q (f:V->W) m. poly_add r p q (m o f) = poly_add r (\m. p (m o f)) (\m. q (m o f)) m `, rw[poly_add] );; let poly_sub_o_permutes = prove(` !(r:R ring) p:(V->num)->R q (f:V->W) m. poly_sub r p q (m o f) = poly_sub r (\m. p (m o f)) (\m. q (m o f)) m `, rw[poly_sub] );; let poly_mul_o_permutes = prove(` !(r:R ring) p:(V->num)->R q S f m. f permutes S ==> poly_mul r p q (m o f) = poly_mul r (\m. p (m o f)) (\m. q (m o f)) m `, rw[poly_mul] THEN intro THEN subgoal `{m1,m2 | monomial_mul m1 m2 = m o f} = IMAGE (\(m1,m2). ((m1:V->num) o (f:V->V)),(m2 o f)) {m1,m2 | monomial_mul m1 m2 = m}` THENL [ rw[EXTENSION;IN_IMAGE;FORALL_PAIR_THM;EXISTS_PAIR_THM;IN_ELIM_PAIR_THM;PAIR_EQ] THEN rw[monomial_mul] THEN intro THEN splitiff THENL [ intro THEN witness `(p1:V->num) o inverse(f:V->V)` THEN witness `(p2:V->num) o inverse(f:V->V)` THEN intro THENL [ rw[GSYM o_ASSOC] THEN specialize[`f:V->V`;`S:V->bool`]PERMUTES_INVERSES_o THEN qed[I_O_ID] ; rw[GSYM o_ASSOC] THEN specialize[`f:V->V`;`S:V->bool`]PERMUTES_INVERSES_o THEN qed[I_O_ID] ; rw[FUN_EQ_THM] THEN intro THEN have_rw `x:V = f (inverse f x:V)` [PERMUTES_INVERSES] THEN rw[o_THM] THEN have `inverse f (f(inverse f x:V):V) = inverse f x` [PERMUTES_INVERSES] THEN qed[o_DEF] ] ; intro THEN simp[o_DEF] THEN rw[FUN_EQ_THM] THEN qed[] ] ; pass ] THEN subgoal `ring_sum(r:R ring) (IMAGE (\(m1,m2). m1 o (f:V->V),m2 o f) {m1,m2 | monomial_mul m1 m2 = m:V->num}) (\(m1,m2). ring_mul r (p m1) (q m2)) = ring_sum r {m1,m2 | monomial_mul m1 m2 = m} ((\(m1,m2). ring_mul r (p m1) (q m2)) o (\(m1,m2). m1 o f,m2 o f))` THENL [ subgoal `!x y. x IN {m1,m2 | monomial_mul m1 m2 = m:V->num} ==> y IN {m1,m2 | monomial_mul m1 m2 = m} ==> (\(m1,m2). m1 o (f:V->V),m2 o f) x = (\(m1,m2). m1 o f,m2 o f) y ==> x = y` THENL [ rw[FORALL_PAIR_THM;IN_ELIM_PAIR_THM;PAIR_EQ] THEN qed[o_permutes_cancel] ; pass ] THEN specialize[ `r:R ring`; `(\(m1,m2). (m1:V->num) o (f:V->V),(m2:V->num) o (f:V->V))`; `\(m1,m2). ring_mul(r:R ring) (p(m1:V->num)) (q(m2:V->num))`; `{m1,m2 | monomial_mul m1 m2 = m:V->num}` ]RING_SUM_IMAGE THEN qed[] ; pass ] THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[FORALL_PAIR_THM] THEN intro THEN rw[o_THM] );; let monomial_1_o_permutes = prove(` !f S. f permutes S ==> monomial_1 o f = monomial_1:V->num `, rw[monomial_1;o_DEF] );; let monomial_1_o_permutes_eq = prove(` !m:V->num f S. f permutes S ==> ( m o f = monomial_1 <=> m = monomial_1 ) `, intro THEN splitiff THENL [ intro THEN have `inverse (f:V->V) permutes S` [PERMUTES_INVERSE] THEN have `monomial_1 o inverse (f:V->V) = monomial_1` [monomial_1_o_permutes] THEN have `(m o f) o inverse (f:V->V) = monomial_1` [monomial_1_o_permutes] THEN subgoal `m o (f o inverse (f:V->V)) = monomial_1:V->num` THENL [ rw[o_ASSOC] THEN qed[] ; pass ] THEN qed[I_O_ID;PERMUTES_INVERSES_o] ; qed[monomial_1_o_permutes] ] );; let poly_const_o_permutes = prove(` !(r:R ring) c (f:V->V) S. f permutes S ==> poly_const r c (m o f) = poly_const r c m `, rw[poly_const] THEN qed[monomial_1_o_permutes_eq] );; let poly_0_o_permutes = prove(` !(r:R ring) (f:V->V) S. f permutes S ==> poly_0 r (m o f) = poly_0 r m `, rw[poly_0;poly_const_o_permutes] );; let poly_1_o_permutes = prove(` !(r:R ring) (f:V->V) S. f permutes S ==> poly_1 r (m o f) = poly_1 r m `, rw[poly_1;poly_const_o_permutes] );; let finite_monomial_vars_permutes_lemma = prove(` !m:V->num f:V->V S. f permutes S ==> FINITE(monomial_vars m) ==> FINITE(monomial_vars (m o f)) `, intro THEN subgoal `IMAGE (f:V->V) (monomial_vars (m o f)) SUBSET monomial_vars m` THENL [ rw[monomial_vars;SUBSET;in_image_vw;IN_ELIM_THM;o_THM] THEN qed[] ; pass ] THEN have `FINITE(IMAGE (f:V->V) (monomial_vars (m o f)))` [FINITE_SUBSET] THEN qed[FINITE_IMAGE_INJ_EQ;PERMUTES_INJECTIVE] );; let finite_monomial_vars_permutes = prove(` !m:V->num f:V->V S. f permutes S ==> ( FINITE(monomial_vars m) <=> FINITE(monomial_vars (m o f)) ) `, intro THEN splitiff THENL [ qed[finite_monomial_vars_permutes_lemma] ; subgoal `m:V->num = (m o f) o inverse (f:V->V)` THENL [ rw[GSYM o_ASSOC] THEN specialize[`f:V->V`;`S:V->bool`]PERMUTES_INVERSES_o THEN simp[I_O_ID] ; pass ] THEN have `inverse (f:V->V) permutes S` [PERMUTES_INVERSE] THEN qed[finite_monomial_vars_permutes_lemma] ] );; let infinite_monomial_vars_permutes = prove(` !m:V->num f:V->V S. f permutes S ==> INFINITE(monomial_vars m) ==> INFINITE(monomial_vars (m o f)) `, rw[INFINITE] THEN qed[finite_monomial_vars_permutes] );; let powerseries_o_permutes = prove(` !(r:R ring) p (f:V->V) S. f permutes S ==> ring_powerseries r p ==> ring_powerseries r (\m. p (m o f)) `, rw[ring_powerseries] THEN intro THENL [ qed[] ; qed[infinite_monomial_vars_permutes] ] );; let polynomial_o_permutes = prove(` !(r:R ring) p (f:V->V) S. f permutes S ==> ring_polynomial r p ==> ring_polynomial r (\m. p (m o f)) `, rw[ring_polynomial] THEN intro THENL [ qed[powerseries_o_permutes] ; subgoal `{m:V->num | ~(p (m o (f:V->V)) = ring_0 r)} = IMAGE (\m. m o inverse f) {m | ~(p m = ring_0(r:R ring))}` THENL [ rw[EXTENSION;IN_IMAGE;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN witness `(x:V->num) o (f:V->V)` THEN rw[GSYM o_ASSOC] THEN specialize[ `f:V->V`; `S:V->bool` ]PERMUTES_INVERSES_o THEN simp[I_O_ID] ; intro THEN have `(x:V->num) o (f:V->V) = (x' o inverse f) o f` [] THEN subgoal `(x' o (inverse f)) o (f:V->V) = x':V->num` THENL [ rw[GSYM o_ASSOC] THEN specialize[ `f:V->V`; `S:V->bool` ]PERMUTES_INVERSES_o THEN simp[I_O_ID] ; pass ] THEN qed[] ] ; pass ] THEN qed[FINITE_IMAGE] ] );; let poly_var_subring = prove(` !(r:R ring) G v:V. poly_var(subring_generated r G) v = poly_var r v `, rw[poly_var] THEN rw[SUBRING_GENERATED] );; let poly_vars_subring = prove(` !(r:R ring) G p:(V->num)->R. poly_vars(subring_generated r G) p = poly_vars r p `, rw[poly_vars] THEN rw[SUBRING_GENERATED] );; let ring_powerseries_if_polynomial = prove(` !(r:R ring) p:(V->num)->R. ring_polynomial r p ==> ring_powerseries r p `, qed[ring_polynomial] );; let ring_polynomial_subring_var = prove(` !(r:R ring) G v:V. ring_polynomial(subring_generated r G) (poly_var r v) `, qed[poly_var_subring;RING_POLYNOMIAL_VAR] );; let ring_sum_poly_o_permutes = prove(` !(r:R ring) p:X->(V->num)->R m f U S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> f permutes U ==> ring_sum(poly_ring r (:V)) S p (m o f) = (ring_sum(poly_ring r (:V)) S (\s m. p s (m o f))) m `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;POLY_RING;POLY_0] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(V->num)->R)` [] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `f:V->V`; `U:V->bool` ]polynomial_o_permutes THEN simp[RING_SUM_CLAUSES;GSYM poly_in_full_ring] THEN rw[POLY_RING] THEN simp[poly_add_o_permutes] THEN have `ring_sum (poly_ring r (:V)) S p (m o f) = ring_sum (poly_ring r (:V)) S (\s:X m:V->num. p s (m o f)) m:R` [] THEN rw[poly_add] THEN qed[] ] );; let ring_product_poly_o_permutes_waterfall = prove(` !(r:R ring) p:X->(V->num)->R f U S. FINITE S ==> !m. (!s. s IN S ==> ring_polynomial r (p s)) ==> f permutes U ==> ring_product(poly_ring r (:V)) S p (m o f) = (ring_product(poly_ring r (:V)) S (\s m. p s (m o f))) m `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POLY_RING;poly_1] THEN qed[poly_const_o_permutes] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(V->num)->R)` [] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `f:V->V`; `U:V->bool` ]polynomial_o_permutes THEN simp[RING_PRODUCT_CLAUSES;GSYM poly_in_full_ring] THEN rw[POLY_RING] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_product (poly_ring r (:V)) (S:X->bool) p:(V->num)->R`; `U:V->bool`; `f:V->V`; `m:V->num` ]poly_mul_o_permutes THEN simp[] THEN have `ring_product (poly_ring r (:V)) S p (m o f) = ring_product (poly_ring r (:V)) S (\s:X m:V->num. p s (m o f)) m:R` [] THEN rw[poly_mul] ] );; let ring_product_poly_o_permutes = prove(` !(r:R ring) p:X->(V->num)->R m f U S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> f permutes U ==> ring_product(poly_ring r (:V)) S p (m o f) = (ring_product(poly_ring r (:V)) S (\s m. p s (m o f))) m `, qed[ring_product_poly_o_permutes_waterfall] );; (* ===== x_series(r) = the ring r[[x]] *) (* ===== x_poly(r) = the ring r[x] *) let x_series = new_definition ` x_series (r:R ring) = powser_ring r (:1) `;; let x_poly = new_definition ` x_poly (r:R ring) = poly_ring r (:1) `;; (* XXX: factor more definitions through this *) let series_from_coeffs = new_definition ` series_from_coeffs (f:num->R) = (\m. f(m one)) `;; let x_poly_vars_1 = prove(` !r:R ring. !p:(1->num)->R. poly_vars r p SUBSET (:1) `, rw[poly_vars;UNIONS_SUBSET] THEN SET_TAC[FINITE_MONOMIAL_VARS_1] );; let x_series_carrier = prove(` !r:R ring. ring_carrier(x_series r) = {p | ring_powerseries r p} `, intro THEN rw[x_series] THEN specialize[`r:R ring`;`(:1)`]powser_ring_carrier THEN simp[IN_ELIM_THM] THEN SET_TAC[x_poly_vars_1] );; (* XXX: merge with x_series_use, generalize everything *) let powser_use = prove(` !(r:R ring). poly_0 r = ring_0(powser_ring r (:V)) /\ poly_1 r = ring_1(powser_ring r (:V)) /\ poly_neg r = ring_neg(powser_ring r (:V)) /\ poly_add r = ring_add(powser_ring r (:V)) /\ poly_mul r = ring_mul(powser_ring r (:V)) /\ !p. ring_powerseries r p <=> p IN ring_carrier(powser_ring r (:V)) `, rw[POWSER_RING;IN_ELIM_THM;SUBSET_UNIV] );; let x_series_use = prove(` !(r:R ring). poly_0 r = ring_0(x_series r) /\ poly_1 r = ring_1(x_series r) /\ poly_neg r = ring_neg(x_series r) /\ poly_add r = ring_add(x_series r) /\ poly_mul r = ring_mul(x_series r) /\ !p. ring_powerseries r p <=> p IN ring_carrier(x_series r) `, rw[x_series_carrier;IN_ELIM_THM] THEN qed[POWSER_RING;x_series] );; let x_poly_carrier = prove(` !r:R ring. ring_carrier(x_poly r) = {p | ring_polynomial r p} `, intro THEN rw[x_poly] THEN specialize[`r:R ring`;`(:1)`]poly_ring_carrier THEN simp[IN_ELIM_THM] THEN SET_TAC[x_poly_vars_1] );; (* XXX: merge with x_poly_use, generalize everything *) let poly_use = prove(` !(r:R ring). poly_0 r = ring_0(poly_ring r (:V)) /\ poly_1 r = ring_1(poly_ring r (:V)) /\ poly_neg r = ring_neg(poly_ring r (:V)) /\ poly_add r = ring_add(poly_ring r (:V)) /\ poly_mul r = ring_mul(poly_ring r (:V)) /\ !p. ring_polynomial r p <=> p IN ring_carrier(poly_ring r (:V)) `, rw[POLY_RING;IN_ELIM_THM;SUBSET_UNIV] );; let x_poly_use = prove(` !(r:R ring). poly_0 r = ring_0(x_poly r) /\ poly_1 r = ring_1(x_poly r) /\ poly_neg r = ring_neg(x_poly r) /\ poly_add r = ring_add(x_poly r) /\ poly_mul r = ring_mul(x_poly r) /\ !p. ring_polynomial r p <=> p IN ring_carrier(x_poly r) `, rw[x_poly_carrier;IN_ELIM_THM] THEN qed[POLY_RING;x_poly] );; let x_series_sub_use = prove(` !(r:R ring). poly_sub r = ring_sub(x_series r) `, rw[FUN_EQ_THM] THEN simp[ring_sub;POLY_SUB] THEN qed[x_series_use] );; let x_poly_sub_use = prove(` !(r:R ring). poly_sub r = ring_sub(x_poly r) `, rw[FUN_EQ_THM] THEN simp[ring_sub;POLY_SUB] THEN qed[x_poly_use] );; let x_poly_carrier_in_series_carrier = prove(` !r:R ring. !p. p IN ring_carrier(x_poly r) ==> p IN ring_carrier(x_series r) `, rw[x_series_carrier;x_poly_carrier] THEN rw[IN_ELIM_THM] THEN qed[RING_POLYNOMIAL_IMP_POWERSERIES] );; let x_poly_carrier_subset_series_carrier = prove(` !r:R ring. ring_carrier(x_poly r) SUBSET ring_carrier(x_series r) `, qed[SUBSET;x_poly_carrier_in_series_carrier] );; let x_poly_subring_in_poly_carrier = prove(` !(r:R ring) G p. p IN ring_carrier(x_poly(subring_generated r G)) ==> p IN ring_carrier(x_poly r) `, rw[GSYM x_poly_use] THEN qed[ring_polynomial_if_subring] );; let ring_of_num_x_series = prove(` !(r:R ring) n. ring_of_num(x_series r) n = poly_const r (ring_of_num r n) `, GEN_TAC THEN INDUCT_TAC THENL [ rw[RING_OF_NUM_0] THEN rw[GSYM x_series_use;poly_0] ; rw[ring_of_num] THEN rw[POLY_CONST_ADD] THEN rw[GSYM x_series_use;poly_1] THEN qed[] ] );; let ring_char_x_series = prove(` !(r:R ring). ring_char (x_series r) = ring_char r `, intro THEN rw[RING_CHAR_UNIQUE] THEN rw[ring_of_num_x_series] THEN rw[x_series;POWSER_RING;poly_0;POLY_CONST_EQ] THEN qed[RING_CHAR_UNIQUE] );; (* ===== x_monomial *) let x_monomial = new_definition ` x_monomial (d:num) = \v:1. d `;; let x_monomial_deg = prove(` !d. monomial_deg(x_monomial d) = d `, rw[x_monomial;MONOMIAL_DEG_UNIVARIATE] );; let x_monomial_0_monomial_1 = prove(` x_monomial 0 = monomial_1 `, rw[x_monomial;monomial_1;COND_ID] );; let x_monomial_add = prove(` !a b. monomial_mul (x_monomial a) (x_monomial b) = x_monomial (a+b) `, rw[x_monomial;monomial_mul;FUN_EQ_THM] THEN qed[ARITH_RULE `0+0 = 0`] );; let x_monomial_surjective = prove(` !m. ?d. x_monomial d = m `, rw[x_monomial] THEN qed[ETA_ONE] );; let x_monomial_injective = prove(` !d e. x_monomial d = x_monomial e ==> d = e `, rw[x_monomial] THEN qed[] );; let ring_sum_image_x_monomial = prove(` !(r:R ring) (g:(1->num)->R) s. ring_sum r (IMAGE x_monomial s) g = ring_sum r s (g o x_monomial) `, qed[x_monomial_injective;RING_SUM_IMAGE] );; let x_series_monomial_mul_finite = prove(` !m:1->num. FINITE {(m1,m2) | monomial_mul m1 m2 = m} `, qed[MONOMIAL_FINITE_DIVISORPAIRS;FINITE_MONOMIAL_VARS_1] );; let x_monomial_pair_injective = prove(` !(a,b) (c,d). (\(a,b). x_monomial a,x_monomial b) (a,b) = (\(a,b). x_monomial a,x_monomial b) (c,d) ==> (a,b) = (c,d) `, rw[FORALL_PAIRED_THM] THEN intro THEN have `x_monomial a = x_monomial c` [FST] THEN have `x_monomial b = x_monomial d` [SND] THEN qed[x_monomial_injective] );; let ring_sum_image_x_monomial_pair = prove(` !(r:R ring) (g:(1->num)#(1->num)->R) (s:num#num->bool). ring_sum r (IMAGE (\(a,b). x_monomial a,x_monomial b) s) g = ring_sum r s (g o (\(a,b). x_monomial a,x_monomial b)) `, intro THEN have `!(a,b) (c,d):num#num. (\(a,b). x_monomial a,x_monomial b) (a,b) = (\(a,b). x_monomial a,x_monomial b) (c,d) ==> (a,b) = (c,d)` [x_monomial_pair_injective] THEN ASSUME_TAC(ISPECL[`r:R ring`;`\(a,b):num#num. x_monomial a,x_monomial b`;`g:(1->num)#(1->num)->R`;`s:(num#num)->bool`]ring_sum_image_injective_pair) THEN qed[] );; let x_monomial_factorizations = prove(` !d (m:1->num) (n:1->num). monomial_mul m n = x_monomial d ==> ?a b. a + b = d /\ m = x_monomial a /\ n = x_monomial b `, rw[x_monomial;monomial_mul] THEN intro THEN witness `m one:num` THEN witness `n one:num` THEN qed[ETA_ONE] );; let x_monomial_factorizations_set = prove(` !d. {m,n:1->num | monomial_mul m n = x_monomial d} = IMAGE (\(a,b). x_monomial a,x_monomial b) {a,b | a+b = d} `, rw[image_pair;EXTENSION] THEN intro_genonly THEN splitiff THENL [ rw[IN_ELIM_THM] THEN intro THEN have `?a b. a + b = d /\ m = x_monomial a /\ n = x_monomial b` [x_monomial_factorizations] THEN qed[] ; rw[IN_ELIM_THM] THEN intro THEN witness `x_monomial a` THEN witness `x_monomial b` THEN qed[x_monomial_add] ] );; (* ===== coeff d p = coefficient of x^d in p *) let coeff = new_definition ` coeff (d:num) (p:(1->num)->R) = p(x_monomial d) `;; let eq_coeff = prove(` !(r:R ring) (p:(1->num)->R) q. (!d. coeff d p = coeff d q) ==> p = q `, rw[coeff] THEN qed[x_monomial_surjective;EQ_EXT] );; let coeff_series_in_ring = prove(` !(r:R ring) p. ring_powerseries r p <=> !d. coeff d p IN ring_carrier(r) `, intro THEN splitiff THENL [ rw[coeff] THEN qed[ring_powerseries] ; intro THEN rw[IN_ELIM_THM] THEN rw[ring_powerseries] THEN intro THENL [ qed[x_monomial_surjective;coeff] ; qed[FINITE_MONOMIAL_VARS_1;INFINITE] ] ] );; let coeff_series_from_coeffs = prove(` !(r:R ring) (f:num->R) d. coeff d (series_from_coeffs f) = f d `, rw[coeff;series_from_coeffs;x_monomial] );; let series_from_coeffs_coeff = prove(` !(r:R ring) f:(1->num)->R. series_from_coeffs (\d. coeff d f) = f `, intro THEN sufficesby eq_coeff THEN qed[coeff_series_from_coeffs] );; let series_series_from_coeffs = prove(` !(r:R ring) (f:num->R). ring_powerseries r (series_from_coeffs f) <=> (!d. f d IN ring_carrier r) `, rw[ring_powerseries;series_from_coeffs;INFINITE;FINITE_MONOMIAL_VARS_1] THEN qed[x_monomial_surjective;x_monomial] );; let poly_series_from_coeffs = prove(` !(r:R ring) (f:num->R). ring_polynomial r (series_from_coeffs f) <=> ((!d. f d IN ring_carrier r) /\ FINITE {d | ~(f d = ring_0 r)} ) `, rw[ring_polynomial] THEN rw[ring_powerseries;series_from_coeffs;INFINITE;FINITE_MONOMIAL_VARS_1] THEN intro THEN splitiff THENL [ intro THENL [ simp[prove(`f d = f((\v.d:num) one):R`,qed[])] ; have `!d e. x_monomial d = x_monomial e ==> d = e` [x_monomial_injective] THEN specialize[`x_monomial`;`{m | ~(f(m one:num) = ring_0(r:R ring))}`]FINITE_IMAGE_INJ THEN set_fact `{d | x_monomial d IN {m | ~(f (m one) = ring_0 r)}} = {d | ~(f((x_monomial d) one) = ring_0(r:R ring))}` THEN set_fact_using `{d | ~(f((x_monomial d) one) = ring_0(r:R ring))} = {d | ~(f d = ring_0 r)}` [x_monomial] THEN qed[] ] ; intro THENL [ simp[] ; have `!m n:1->num. m one = n one ==> m = n` [one_Axiom] THEN specialize[`\m:1->num. m one`;`{d:num | ~(f d = ring_0(r:R ring))}`]FINITE_IMAGE_INJ THEN set_fact `{m:1->num | m one IN {d | ~(f d = ring_0 r)}} = {m | ~(f(m one) = ring_0(r:R ring))}` THEN qed[] ] ] );; let poly_coeff = prove(` !(r:R ring) p:(1->num)->R. ring_polynomial r p <=> ((!d. coeff d p IN ring_carrier r) /\ FINITE {d | ~(coeff d p = ring_0 r)} ) `, intro THEN have `series_from_coeffs (\d. coeff d p:R) = p` [series_from_coeffs_coeff] THEN specialize[`r:R ring`;`\d. coeff d p:R`]poly_series_from_coeffs THEN qed[] );; let ring_polynomial_subring_if_coeffs = prove(` !(r:R ring) (p:(1->num)->R). ring_polynomial r p ==> (!d. coeff d p IN ring_carrier(subring_generated r G)) ==> ring_polynomial(subring_generated r G) p `, rw[poly_coeff] THEN REPEAT GEN_TAC THEN DISCH_TAC THEN subgoal `{d | ~(coeff d p = ring_0 (subring_generated r G))} = {d | ~(coeff d p = ring_0(r:R ring))}` THENL [ rw[EXTENSION;IN_ELIM_THM] THEN qed[SUBRING_GENERATED] ; pass ] THEN qed[] );; let coeff_series_carrier_in_ring = prove(` !(r:R ring) p. p IN ring_carrier(x_series r) <=> !d. coeff d p IN ring_carrier(r) `, rw[x_series;POWSER_RING] THEN SET_TAC[coeff_series_in_ring] );; let coeff_poly_in_ring = prove(` !(r:R ring) d p. ring_polynomial r p ==> coeff d p IN ring_carrier(r) `, qed[ring_polynomial;coeff_series_in_ring] );; let coeff_poly_carrier_in_ring = prove(` !(r:R ring) d p. p IN ring_carrier(x_poly r) ==> coeff d p IN ring_carrier(r) `, qed[x_poly_carrier_in_series_carrier;coeff_series_carrier_in_ring] );; let coeff_poly_const = prove(` !(r:R ring) c:R d. coeff d (poly_const r c) = if d = 0 then c else ring_0 r `, rw[coeff;x_monomial;poly_const;monomial_1] THEN qed[] );; let coeff_poly_0 = prove(` !(r:R ring). coeff d (poly_0 r) = ring_0 r `, rw[poly_0;coeff_poly_const] THEN qed[] );; let coeff_poly_1 = prove(` !(r:R ring). coeff d (poly_1 r) = if d = 0 then ring_1 r else ring_0 r `, rw[poly_1;coeff_poly_const] );; let coeff_poly_neg = prove(` !(r:R ring) p d. coeff d (poly_neg r p) = ring_neg r (coeff d p) `, rw[coeff;poly_neg] );; let coeff_poly_add = prove(` !(r:R ring) p q d. coeff d (poly_add r p q) = ring_add r (coeff d p) (coeff d q) `, rw[coeff;poly_add] );; let coeff_series_add = prove(` !(r:R ring) p q d. coeff d (ring_add(x_series r) p q) = ring_add r (coeff d p) (coeff d q) `, rw[x_series;POWSER_RING;coeff_poly_add] );; let coeff_poly_sub = prove(` !(r:R ring) p q d. coeff d (poly_sub r p q) = ring_sub r (coeff d p) (coeff d q) `, rw[coeff;poly_sub] );; let coeff_poly_mul_lemma = prove(` !(r:R ring) p q. ((\(m1,m2). ring_mul r (p m1) (q m2)) o (\(a,b). x_monomial a,x_monomial b)) = (\(a,b). ring_mul r (p (x_monomial a)) (q (x_monomial b))) `, rw[FUN_EQ_THM] THEN intro_genonly THEN choose2 `a:num` `b:num` `x = (a:num),(b:num)` [PAIR_SURJECTIVE] THEN simp[o_THM] );; let coeff_poly_mul = prove(` !(r:R ring) p q d. coeff d (poly_mul r p q) = ring_sum r {a,b | a+b = d} (\(a,b). ring_mul r (coeff a p) (coeff b q)) `, rw[poly_mul;coeff] THEN intro THEN rw[x_monomial_factorizations_set] THEN rw[ring_sum_image_x_monomial_pair] THEN rw[coeff_poly_mul_lemma] );; let coeff_poly_mul_oneindex = prove(` !(r:R ring) d p q. coeff d (poly_mul r p q) = ring_sum(r) (0..d) (\a. ring_mul r (coeff a p) (coeff (d-a) q)) `, intro THEN rw[coeff_poly_mul] THEN def `f:num->(num#num)` `\a:num. a,d-a` THEN have `!a b:num. (f:num->num#num) a = f b ==> a = b` [FST] THEN have `ring_sum r (IMAGE (f:num->num#num) (0..d)) (\(a,b). ring_mul r (coeff a p) (coeff b q)) = ring_sum(r:R ring) (0..d) ((\(a,b). ring_mul r (coeff a p) (coeff b q)) o f)` [RING_SUM_IMAGE] THEN have `IMAGE (f:num->num#num) (0..d) = {a,b | a + b = d}` [image_numseg_antidiagonal] THEN ASSUME_TAC(prove(`(\(a,b). ring_mul(r:R ring) (coeff a p) (coeff b q)) o (\a. a,d-a) = (\a. ring_mul r (coeff a p) (coeff (d-a) q))`,rw[FUN_EQ_THM;o_THM])) THEN qed[] );; (* XXX: use this to prove coeff_poly_const_times *) let poly_const_times = prove(` !(r:R ring) c:R p:(V->num)->R. c IN ring_carrier r ==> ring_powerseries r p ==> poly_mul r (poly_const r c) p = (\m. ring_mul r c (p m)) `, rw[poly_mul] THEN once_rw[fun_eq_thm_v] THEN rw[BETA_THM] THEN intro THEN subgoal `ring_sum(r:R ring) {m1,m2 | monomial_mul m1 m2 = v:V->num} (\(m1,m2). ring_mul r (poly_const r c m1) (p m2)) = ring_sum r {m1,m2 | monomial_mul m1 m2 = v} (\j. if j = (monomial_1,v) then ring_mul r c (p v) else ring_0 r)` THENL [ sufficesby RING_SUM_EQ THEN rw[FORALL_PAIR_THM;IN_ELIM_PAIR_THM;PAIR_EQ;poly_const] THEN intro THEN case `p1 = monomial_1:V->num` THENL [ have `p2 = v:V->num` [MONOMIAL_MUL_LID] THEN simp[] ; pass ] THEN have `p(p2:V->num) IN ring_carrier(r:R ring)` [ring_powerseries] THEN qed[RING_MUL_LZERO] ; pass ] THEN subgoal `monomial_1,v IN {m1,m2 | monomial_mul m1 m2 = v:V->num}` THENL [ rw[IN_ELIM_PAIR_THM] THEN rw[MONOMIAL_MUL_LID] ; pass ] THEN simp[RING_SUM_DELTA] THEN qed[ring_powerseries;RING_MUL] );; let coeff_poly_const_times = prove(` !(r:R ring) c:R p d. c IN ring_carrier r ==> ring_powerseries r p ==> coeff d (poly_mul r (poly_const r c) p) = ring_mul r c (coeff d p) `, rw[coeff_poly_mul_oneindex] THEN simp[coeff_poly_const] THEN simp[prove(`ring_powerseries r p ==> ring_mul (r:R ring) (if a:num = 0 then c else ring_0 r) (coeff (d - a) p) = if a = 0 then ring_mul r c (coeff (d - a) p) else ring_0 r`,qed[RING_MUL_LZERO;coeff_series_in_ring])] THEN rw[RING_SUM_DELTA] THEN intro THEN set_fact_using `0 IN (0..d)` [NUMSEG_LE;ARITH_RULE `0 <= d:num`] THEN have `coeff(d - 0) p IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `ring_mul(r:R ring) c (coeff(d - 0) p) IN ring_carrier r` [RING_MUL] THEN qed[ARITH_RULE `d - 0 = d:num`] );; let coeff_times_poly_const = prove(` !(r:R ring) c:R p d. c IN ring_carrier r ==> ring_powerseries r p ==> coeff d (poly_mul r p (poly_const r c)) = ring_mul r c (coeff d p) `, qed[RING_POWERSERIES_CONST;POLY_MUL_SYM;coeff_poly_const_times] );; let polynomial_if_coeff = prove(` !(r:R ring) p n. ring_powerseries r p ==> (!d. ~(coeff d p = ring_0 r) ==> d <= n) ==> ring_polynomial r p `, intro THEN rw[ring_polynomial] THEN simp[] THEN subgoal `{d | ~(coeff d p = ring_0(r:R ring))} SUBSET {d:num | d <= n}` THENL [ rw[SUBSET;IN_ELIM_THM] THEN qed[] ; pass ] THEN have `FINITE {d | ~(coeff d p = ring_0(r:R ring))}` [FINITE_SUBSET;FINITE_NUMSEG_LE] THEN set_fact_using `{d | ~(coeff d p = ring_0(r:R ring))} = {d | ~(p(x_monomial d) = ring_0(r:R ring))}` [coeff] THEN have `FINITE {d | ~(p(x_monomial d) = ring_0(r:R ring))}` [] THEN recall x_monomial_surjective THEN specialize[`x_monomial`;`\m:1->num. ~(p m = ring_0(r:R ring))`]surjective_finite THEN qed[] );; let deg_le_coeff = prove(` !(r:R ring) p n. ring_powerseries r p ==> (!d. ~(coeff d p = ring_0 r) ==> d <= n) ==> poly_deg r p <= n `, intro THEN sufficesby POLY_DEG_LE THEN have `ring_polynomial(r:R ring) (p:(1->num)->R)` [polynomial_if_coeff] THEN simp[] THEN intro THEN choose `d:num` `x_monomial d = m` [x_monomial_surjective] THEN qed[x_monomial_deg;coeff] );; let deg_coeff = prove(` !(r:R ring) p n. ring_powerseries r p ==> (!d. ~(coeff d p = ring_0 r) ==> d <= n) ==> ~(coeff n p = ring_0 r) ==> poly_deg r p = n `, intro THEN have `ring_polynomial(r:R ring) (p:(1->num)->R)` [polynomial_if_coeff] THEN simp[POLY_DEG_EQ] THEN intro THENL [ choose `d:num` `x_monomial d = m` [x_monomial_surjective] THEN qed[x_monomial_deg;coeff] ; DISJ2_TAC THEN witness `x_monomial n` THEN qed[x_monomial_deg;coeff] ] );; let topcoeff_nonzero = prove(` !(r:R ring) p. ring_polynomial r p ==> (p = poly_0 r <=> coeff (poly_deg r p) p = ring_0 r) `, intro THEN splitiff THENL [ qed[coeff_poly_0] ; rw[coeff;x_monomial] THEN intro THEN have `p IN ring_carrier(poly_ring(r:R ring) (:1))` [x_poly_use;x_poly] THEN have `p = ring_0(poly_ring(r:R ring) (:1))` [POLY_TOP_NONZERO] THEN qed[x_poly_use;x_poly] ] );; let coeff_deg_le = prove(` !(r:R ring) p n d. ring_polynomial r p ==> poly_deg r p <= n ==> ~(coeff d p = ring_0 r) ==> d <= n `, intro THEN have `~(p(x_monomial d) = ring_0(r:R ring))` [coeff] THEN have `monomial_deg(x_monomial d) <= n` [POLY_DEG_LE_EQ] THEN qed[x_monomial_deg] );; let coeff_le_deg = prove(` !(r:R ring) p d. ring_polynomial r p ==> ~(coeff d p = ring_0 r) ==> d <= poly_deg r p `, intro THEN num_linear_fact `poly_deg(r:R ring) (p:(1->num)->R) <= poly_deg r p` THEN qed[coeff_deg_le] );; let finite_coeff = prove(` !(r:R ring) p. ring_polynomial r p ==> FINITE {d | ~(coeff d p = ring_0 r)} `, intro THEN specialize[`poly_deg r (p:(1->num)->R)`]FINITE_NUMSEG_LE THEN have `!d. ~(coeff d p = ring_0 r) ==> d <= poly_deg (r:R ring) p` [coeff_le_deg] THEN set_fact `(!d. ~(coeff d p = ring_0 r) ==> d <= poly_deg (r:R ring) p) ==> {d | ~(coeff d p = ring_0 r)} SUBSET {d | d <= poly_deg (r:R ring) p}` THEN qed[FINITE_SUBSET] );; let poly_if_coeff = prove(` !(r:R ring) p n. ring_powerseries r p ==> (!d. n <= d ==> coeff d p = ring_0 r) ==> ring_polynomial r p `, intro THEN rw[ring_polynomial] THEN subgoal `{m | ~(p m = ring_0(r:R ring))} SUBSET IMAGE x_monomial (0..n)` THENL [ rw[SUBSET;IN_IMAGE;IN_ELIM_THM] THEN intro THEN choose `d:num` `x_monomial d = x` [x_monomial_surjective] THEN witness `d:num` THEN case `n <= d:num` THENL [ have `p(x_monomial d) = ring_0(r:R ring)` [coeff] THEN qed[] ; pass ] THEN num_linear_fact `~(n <= d:num) ==> d <= n` THEN have `d IN 0..n` [IN_NUMSEG_0] THEN qed[] ; pass ] THEN qed[FINITE_IMAGE;FINITE_SUBSET;FINITE_NUMSEG] );; let deg_coeff_from_le = prove(` !(r:R ring) p n. ring_polynomial r p ==> poly_deg r p <= n ==> ~(coeff n p = ring_0 r) ==> poly_deg r p = n `, intro THEN have `!d. ~(coeff d p = ring_0(r:R ring)) ==> d <= n` [coeff_le_deg;LE_TRANS] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN qed[deg_coeff] );; let poly_eval_expand_coeff = prove(` !(r:R ring) p x n. ring_polynomial r p ==> x IN ring_carrier r ==> poly_deg r p <= n ==> poly_eval r p x = ring_sum r (0..n) (\d. ring_mul r (coeff d p) (ring_pow r x d)) `, intro THEN set_fact `ring_polynomial r p ==> p IN {q | ring_polynomial(r:R ring) (q:(1->num)->R)}` THEN have `(p:(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_carrier] THEN have `(p:(1->num)->R) IN ring_carrier(poly_ring r (:1))` [x_poly] THEN simp[POLY_EVAL_EXPAND] THEN simp[ring_sum_numseg_le_expand] THEN sufficesby RING_SUM_EQ THEN simp[GSYM coeff;GSYM x_monomial] THEN qed[RING_POW;RING_MUL_LZERO;coeff_le_deg] );; let deg_mul_const_le = prove(` !(r:R ring) (p:(V->num)->R) c. ring_polynomial r p ==> c IN ring_carrier r ==> poly_deg r (poly_mul r p (poly_const r c)) <= poly_deg r p `, intro THEN have `ring_polynomial r (poly_const r c:(V->num)->R)` [RING_POLYNOMIAL_CONST] THEN have `poly_deg r (poly_const r c:(V->num)->R) = 0` [POLY_DEG_CONST] THEN specialize[`r:R ring`;`p:(V->num)->R`;`poly_const r c:(V->num)->R`]POLY_DEG_MUL_LE THEN num_linear_fact `poly_deg r (p:(V->num)->R) + 0 = poly_deg r (p:(V->num)->R)` THEN qed[] );; let deg_const_mul_le = prove(` !(r:R ring) (p:(V->num)->R) c. ring_polynomial r p ==> c IN ring_carrier r ==> poly_deg r (poly_mul r (poly_const r c) p) <= poly_deg r p `, qed[deg_mul_const_le;RING_POLYNOMIAL_CONST;ring_polynomial;POLY_MUL_SYM] );; let poly_mul_const_const = prove(` !(r:R ring) (p:(V->num)->R) c d. ring_polynomial r p ==> c IN ring_carrier r ==> d IN ring_carrier r ==> poly_mul r (poly_mul r p (poly_const r c)) (poly_const r d) = poly_mul r p (poly_const r (ring_mul r c d)) `, qed[ring_polynomial;RING_POLYNOMIAL_CONST;POLY_MUL_ASSOC;POLY_CONST_MUL] );; let poly_mul_const_const_1 = prove(` !(r:R ring) (p:(V->num)->R) c d. ring_polynomial r p ==> c IN ring_carrier r ==> d IN ring_carrier r ==> ring_mul r c d = ring_1 r ==> poly_mul r (poly_mul r p (poly_const r c)) (poly_const r d) = p `, simp[poly_mul_const_const] THEN rw[GSYM poly_1] THEN intro THEN set_fact `ring_polynomial r (p:(V->num)->R) ==> p IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:V)}` THEN have `(p:(V->num)->R) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:V)}` [] THEN havetac `(p:(V->num)->R) IN ring_carrier(poly_ring r (:V))` (simp[POLY_RING_CLAUSES]) THEN qed[RING_MUL_RID;POLY_RING_CLAUSES] );; let deg_mul_const_const_1 = prove(` !(r:R ring) (p:(V->num)->R) c d. ring_polynomial r p ==> c IN ring_carrier r ==> d IN ring_carrier r ==> ring_mul r c d = ring_1 r ==> poly_deg r (poly_mul r p (poly_const r c)) = poly_deg r p `, intro THEN def `q:(V->num)->R` `poly_mul r (p:(V->num)->R) (poly_const r c)` THEN have `ring_polynomial r (q:(V->num)->R)` [RING_POLYNOMIAL_CONST;RING_POLYNOMIAL_MUL] THEN have `poly_deg r (q:(V->num)->R) <= poly_deg r (p:(V->num)->R)` [deg_mul_const_le] THEN have `(p:(V->num)->R) = poly_mul r q (poly_const r d)` [poly_mul_const_const_1] THEN have `poly_deg r (p:(V->num)->R) <= poly_deg r (q:(V->num)->R)` [deg_mul_const_le] THEN qed[ARITH_RULE `a <= b /\ b <= a ==> a = b:num`] );; let deg_mul_unit_const = prove(` !(r:R ring) (p:(V->num)->R) c. ring_polynomial r p ==> ring_unit r c ==> poly_deg r (poly_mul r p (poly_const r c)) = poly_deg r p `, intro THEN choose `d:R` `d IN ring_carrier(r:R ring) /\ ring_mul r c d = ring_1 r` [ring_unit] THEN have `c IN ring_carrier(r:R ring)` [ring_unit] THEN qed[deg_mul_const_const_1] );; let associates_if_mul_unit_const = prove(` !(r:R ring) (p:(V->num)->R) c. ring_polynomial r p ==> ring_unit r c ==> ring_associates(poly_ring r (:V)) p (poly_mul r p (poly_const r c)) `, intro THEN rw[ring_associates;ring_divides] THEN intro THENL [ qed[RING_POLYNOMIAL] ; qed[ring_unit;RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_CONST;RING_POLYNOMIAL] ; witness `poly_const r c:(V->num)->R` THEN have `c IN ring_carrier(r:R ring)` [ring_unit] THEN qed[RING_POLYNOMIAL_CONST;RING_POLYNOMIAL;POLY_CLAUSES] ; have `c IN ring_carrier(r:R ring)` [ring_unit] THEN qed[RING_POLYNOMIAL;POLY_CLAUSES;RING_POLYNOMIAL_CONST;RING_POLYNOMIAL_MUL] ; qed[POLY_CLAUSES;RING_POLYNOMIAL] ; witness `poly_const r (ring_inv r c):(V->num)->R` THEN have `c IN ring_carrier(r:R ring)` [ring_unit] THEN have `ring_inv r c IN ring_carrier(r:R ring)` [RING_INV] THEN intro THENL [ qed[RING_POLYNOMIAL;RING_POLYNOMIAL_CONST] ; have `ring_mul r c (ring_inv r c) = ring_1(r:R ring)` [ring_div;ring_div_refl] THEN rw[GSYM POLY_CLAUSES] THEN qed[poly_mul_const_const_1] ] ] );; (* ===== cx^d *) let const_x_pow = new_definition ` const_x_pow (r:R ring) (c:R) (d:num) = \m. if m one = d then c else ring_0 r `;; let const_x_pow_0 = prove(` !(r:R ring) c. const_x_pow r c 0 = poly_const r c `, rw[const_x_pow;poly_const;monomial_1] THEN rw[FUN_EQ_THM] THEN qed[one] );; let const_0_x_pow = prove(` !(r:R ring) n. const_x_pow r (ring_0 r) n = poly_0 r `, rw[const_x_pow;poly_0;poly_const;FUN_EQ_THM;COND_ID] );; let const_x_pow_series = prove(` !(r:R ring) c d. c IN ring_carrier r ==> ring_powerseries r (const_x_pow r c d) `, intro THEN rw[const_x_pow;ring_powerseries] THEN qed[RING_0;FINITE_MONOMIAL_VARS_1;INFINITE] );; let const_x_pow_poly_lemma = prove(` !(r:R ring) d:num. {m:1->num | ~((if m one = d then c else ring_0 r) = ring_0 r)} SUBSET {\v. d} `, rw[SUBSET;IN_ELIM_THM;IN_SING] THEN qed[ETA_ONE] );; let const_x_pow_poly_lemma2 = prove(` !(r:R ring) d:num. FINITE {m:1->num | ~((if m one = d then c else ring_0 r) = ring_0 r)} `, qed[const_x_pow_poly_lemma;FINITE_SING;FINITE_SUBSET] );; let const_x_pow_poly = prove(` !(r:R ring) c d. c IN ring_carrier r ==> ring_polynomial r (const_x_pow r c d) `, intro THEN rw[ring_polynomial] THEN conjunction THENL [qed[const_x_pow_series]; pass] THEN rw[const_x_pow] THEN qed[const_x_pow_poly_lemma2] );; let coeff_const_x_pow = prove(` !(r:R ring) c d e. coeff e (const_x_pow r c d) = if e = d then c else ring_0 r `, rw[const_x_pow;coeff;x_monomial] );; let coeff_const_x_pow_times = prove(` !(r:R ring) c d p e. c IN ring_carrier r ==> ring_powerseries r p ==> coeff e (poly_mul r (const_x_pow r c d) p) = if e < d then ring_0 r else ring_mul r c (coeff (e - d) p) `, rw[coeff_poly_mul_oneindex] THEN simp[coeff_const_x_pow] THEN simp[prove(`ring_powerseries r p ==> ring_mul (r:R ring) (if a:num = d then c else ring_0 r) (coeff (e - a) p) = if a = d then ring_mul r c (coeff (e - a) p) else ring_0 r`,qed[RING_MUL_LZERO;coeff_series_in_ring])] THEN rw[RING_SUM_DELTA] THEN intro THEN have `coeff (e - d) p IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `ring_mul r c (coeff (e - d) p) IN ring_carrier(r:R ring)` [RING_MUL] THEN set_fact_using `d IN (0..e) <=> d <= e` [NUMSEG_LE] THEN qed[ARITH_RULE `d:num <= e <=> ~(e < d)`] );; let deg_const_x_pow_le = prove(` !(r:R ring) c d. c IN ring_carrier r ==> poly_deg r (const_x_pow r c d) <= d `, intro THEN have `ring_powerseries(r:R ring) (const_x_pow r c d)` [const_x_pow_series] THEN have `!e. ~(coeff e (const_x_pow(r:R ring) c d) = ring_0 r) ==> e <= d` [coeff_const_x_pow;ARITH_RULE `d <= d:num`] THEN qed[deg_le_coeff] );; let deg_const_x_pow = prove(` !(r:R ring) c d. c IN ring_carrier r ==> ~(c = ring_0 r) ==> poly_deg r (const_x_pow r c d) = d `, intro THEN have `ring_powerseries(r:R ring) (const_x_pow r c d)` [const_x_pow_series] THEN have `!e. ~(coeff e (const_x_pow(r:R ring) c d) = ring_0 r) ==> e <= d` [coeff_const_x_pow;ARITH_RULE `d <= d:num`] THEN have `~(coeff d (const_x_pow r c d) = ring_0(r:R ring))` [coeff_const_x_pow] THEN qed[deg_coeff] );; let eval_const_x_pow = prove(` !(r:R ring) c n x. c IN ring_carrier r ==> x IN ring_carrier r ==> poly_eval r (const_x_pow r c n) x = ring_mul r c (ring_pow r x n) `, intro THEN have `ring_polynomial(r:R ring) (const_x_pow r c n)` [const_x_pow_poly] THEN have `poly_deg(r:R ring) (const_x_pow r c n) <= n` [deg_const_x_pow_le] THEN specialize[`r:R ring`;`const_x_pow(r:R ring) c n`;`x:R`;`n:num`]poly_eval_expand_coeff THEN simp[coeff_const_x_pow] THEN specialize[`r:R ring`;`0..n`;`n:num`;`ring_mul(r:R ring) c (ring_pow r x a)`]RING_SUM_DELTA THEN have `!a. a IN (0..n) ==> ring_mul(r:R ring) (if a = n then c else ring_0 r) (ring_pow r x a) = if a = n then ring_mul r c (ring_pow r x a) else ring_0 r` [RING_MUL_LZERO;RING_POW] THEN specialize[`r:R ring`;`\a. ring_mul(r:R ring) (if a = n then c else ring_0 r) (ring_pow r x a)`;`\a. if a = n then ring_mul r c (ring_pow r x a) else ring_0(r:R ring)`;`0..n`]RING_SUM_EQ THEN have `n IN 0..n` [IN_NUMSEG_0;ARITH_RULE `n <= n:num`] THEN simp[] THEN qed[RING_MUL;RING_POW] );; let subring_const_x_pow = prove(` !(r:R ring) S c n. const_x_pow (subring_generated r S) c n = const_x_pow r c n `, rw[const_x_pow] THEN rw[SUBRING_GENERATED] );; let const_1_x_pow = prove(` !(r:R ring). const_x_pow r (ring_1 r) 0 = poly_1 r `, intro THEN sufficesby eq_coeff THEN rw[coeff_const_x_pow;coeff_poly_1] );; (* ===== x^d *) let x_pow = new_definition ` x_pow (r:R ring) (d:num) = const_x_pow r (ring_1 r) d `;; let x_pow_1 = prove(` !(r:R ring). x_pow r 1 = poly_var r one `, rw[x_pow;const_x_pow;poly_var;monomial_var;FUN_EQ_THM] THEN qed[ETA_ONE;one] );; let x_pow_0 = prove(` !(r:R ring). x_pow r 0 = poly_1 r `, rw[x_pow;const_x_pow_0;poly_1] );; let x_pow_poly = prove(` !(r:R ring) d. ring_polynomial r (x_pow r d) `, qed[x_pow;const_x_pow_poly;RING_1] );; let x_pow_series = prove(` !(r:R ring) d. ring_powerseries r (x_pow r d) `, qed[ring_polynomial;x_pow_poly] );; let coeff_x_pow = prove(` !(r:R ring) d e. coeff e (x_pow r d) = if e = d then ring_1 r else ring_0 r `, rw[x_pow;coeff_const_x_pow] );; let coeff_x_pow_times = prove(` !(r:R ring) d p e. ring_powerseries r p ==> coeff e (poly_mul r (x_pow r d) p) = if e < d then ring_0 r else coeff (e - d) p `, intro THEN have `ring_1 r IN ring_carrier(r:R ring)` [RING_1] THEN rw[x_pow] THEN qed[coeff_const_x_pow_times;RING_MUL_LID;coeff_series_in_ring] );; let x_pow_add = prove(` !(r:R ring) m n. x_pow r (m+n) = poly_mul r (x_pow r m) (x_pow r n) `, intro THEN sufficesby eq_coeff THEN simp[coeff_x_pow_times;x_pow_series] THEN rw[coeff_x_pow] THEN qed[ARITH_RULE `d = m + n <=> (~(d < m) /\ d - m = n:num)`] );; let deg_x_pow_le = prove(` !(r:R ring) d. poly_deg r (x_pow r d) <= d `, qed[x_pow;deg_const_x_pow_le;RING_1] );; let deg_x_pow = prove(` !(r:R ring) d. ~(ring_1 r = ring_0 r) ==> poly_deg r (x_pow r d) = d `, simp[x_pow;deg_const_x_pow;RING_1] );; let eval_x_pow = prove(` !(r:R ring) n x. x IN ring_carrier r ==> poly_eval r (x_pow r n) x = ring_pow r x n `, intro THEN rw[x_pow] THEN have `ring_1(r:R ring) IN ring_carrier r` [RING_1] THEN simp[eval_const_x_pow] THEN qed[RING_POW;RING_MUL_LID] );; let subring_x_pow = prove(` !(r:R ring) S n. x_pow (subring_generated r S) n = x_pow r n `, rw[x_pow;const_x_pow;SUBRING_GENERATED] );; (* ===== x-c *) let x_minus_const = new_definition ` x_minus_const (r:R ring) (c:R) = poly_sub r (x_pow r 1) (poly_const r c) `;; let x_minus_const_poly = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_polynomial r (x_minus_const r c) `, qed[x_minus_const;RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST;x_pow_poly] );; let x_minus_const_series = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_powerseries r (x_minus_const r c) `, qed[ring_polynomial;x_minus_const_poly] );; let coeff_x_minus_const = prove(` !(r:R ring) c e. c IN ring_carrier r ==> coeff e (x_minus_const r c) = if e = 1 then ring_1 r else if e = 0 then ring_neg r c else ring_0 r `, rw[x_minus_const;coeff_poly_sub;coeff_poly_const;coeff_x_pow] THEN intro THEN case `e = 0` THENL [ simp[ARITH_RULE `~(0 = 1)`] ; case `e = 1` THENL [ simp[] ; simp[] ] ] THEN RING_TAC );; let coeff_x_minus_const_times = prove(` !(r:R ring) c p e. c IN ring_carrier r ==> ring_powerseries r p ==> coeff e (poly_mul r (x_minus_const r c) p) = if e = 0 then ring_mul r (ring_neg r c) (coeff 0 p) else ring_sub r (coeff (e-1) p) (ring_mul r c (coeff e p)) `, intro THEN rw[x_minus_const] THEN have `ring_powerseries(r:R ring) (x_pow r 1)` [x_pow_series] THEN have `ring_powerseries(r:R ring) (poly_const r c:(1->num)->R)` [RING_POWERSERIES_CONST] THEN simp[poly_sub_rdistrib] THEN rw[coeff_poly_sub] THEN simp[coeff_x_pow_times] THEN simp[coeff_poly_const_times] THEN case `e = 0` THENL [ simp[ARITH_RULE `0 < 1`] THEN RING_TAC THEN qed[coeff_series_in_ring] ; pass ] THEN simp[ARITH_RULE `~(e:num = 0) ==> ~(e < 1)`] );; let x_minus_const_0 = prove(` !(r:R ring). x_minus_const r (ring_0 r) = x_pow r 1 `, intro THEN sufficesby eq_coeff THEN intro THEN simp[coeff_x_minus_const;coeff_x_pow;RING_0] THEN rw[RING_NEG_0] THEN qed[] );; let deg_x_minus_const_le = prove(` !(r:R ring) c. c IN ring_carrier r ==> poly_deg r (x_minus_const r c) <= 1 `, intro THEN have `ring_powerseries(r:R ring) (x_minus_const r c)` [x_minus_const_series] THEN have `!e. ~(coeff e (x_minus_const(r:R ring) c) = ring_0 r) ==> e <= 1` [coeff_x_minus_const;ARITH_RULE `0 <= 1:num`;ARITH_RULE `1 <= 1:num`] THEN qed[deg_le_coeff] );; let deg_x_minus_const = prove(` !(r:R ring) c. c IN ring_carrier r ==> ~(ring_1 r = ring_0 r) ==> poly_deg r (x_minus_const r c) = 1 `, intro THEN have `ring_powerseries(r:R ring) (x_minus_const r c)` [x_minus_const_series] THEN have `!e. ~(coeff e (x_minus_const(r:R ring) c) = ring_0 r) ==> e <= 1` [coeff_x_minus_const;ARITH_RULE `0 <= 1:num`;ARITH_RULE `1 <= 1:num`] THEN have `~(coeff 1 (x_minus_const r c) = ring_0(r:R ring))` [coeff_x_minus_const] THEN qed[deg_coeff] );; let x_minus_const_nonzero = prove(` !(r:R ring) c. c IN ring_carrier r ==> ~(ring_1 r = ring_0 r) ==> ~(x_minus_const r c = poly_0 r) `, qed[deg_x_minus_const;POLY_DEG_0;ARITH_RULE `~(0 = 1)`] );; let eval_x_minus_const = prove(` !(r:R ring) c x. c IN ring_carrier r ==> x IN ring_carrier r ==> poly_eval r (x_minus_const r c) x = ring_sub r x c `, intro THEN rw[x_minus_const] THEN have `ring_polynomial(r:R ring) (x_pow r 1)` [x_pow_poly] THEN have `ring_polynomial(r:R ring) (poly_const r c:(1->num)->R)` [RING_POLYNOMIAL_CONST] THEN simp[POLY_EVAL_SUB] THEN simp[POLY_EVAL_CONST] THEN simp[eval_x_pow] THEN simp[RING_POW_1] );; let eval_x_minus_const_refl = prove(` !(r:R ring) c. c IN ring_carrier r ==> poly_eval r (x_minus_const r c) c = ring_0 r `, simp[eval_x_minus_const] THEN rw[RING_SUB_REFL] );; let x_minus_const_not_unit = prove(` !(r:R ring) z. integral_domain r ==> z IN ring_carrier r ==> ~ring_unit(x_poly r) (x_minus_const r z) `, intro THEN specialize_assuming[`r:R ring`;`(:1)`;`x_minus_const (r:R ring) z`]RING_UNIT_POLY_DOMAIN THEN choose `c:R` `ring_unit(r:R ring) c /\ x_minus_const r z = poly_const r c` [x_poly] THEN have `poly_deg (r:R ring) (x_minus_const r z) = 1` [deg_x_minus_const;integral_domain] THEN have `poly_deg (r:R ring) (poly_const r c:(1->num)->R) = 0` [POLY_DEG_CONST] THEN qed[ARITH_RULE `~(1 = 0)`] );; (* ===== 1-cx *) let one_minus_constx = new_definition ` one_minus_constx (r:R ring) (c:R) = poly_sub r (x_pow r 0) (const_x_pow r c 1) `;; let one_minus_constx_poly = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_polynomial r (one_minus_constx r c) `, qed[one_minus_constx;RING_POLYNOMIAL_SUB;x_pow_poly;const_x_pow_poly] );; let one_minus_constx_series = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_powerseries r (one_minus_constx r c) `, qed[ring_polynomial;one_minus_constx_poly] );; let coeff_one_minus_constx = prove(` !(r:R ring) c e. c IN ring_carrier r ==> coeff e (one_minus_constx r c) = if e = 0 then ring_1 r else if e = 1 then ring_neg r c else ring_0 r `, rw[one_minus_constx;coeff_poly_sub;coeff_poly_const;coeff_x_pow;coeff_const_x_pow] THEN intro THEN case `e = 1` THENL [ simp[ARITH_RULE `~(0 = 1)`] ; case `e = 0` THENL [ simp[] ; simp[] ] ] THEN RING_TAC );; let coeff_one_minus_constx_times = prove(` !(r:R ring) c p e. c IN ring_carrier r ==> ring_powerseries r p ==> coeff e (poly_mul r (one_minus_constx r c) p) = if e = 0 then coeff 0 p else ring_sub r (coeff e p) (ring_mul r c (coeff (e-1) p)) `, intro THEN rw[one_minus_constx] THEN have `ring_powerseries(r:R ring) (x_pow r 0)` [x_pow_series] THEN have `ring_powerseries(r:R ring) (const_x_pow r c 1)` [const_x_pow_series] THEN simp[poly_sub_rdistrib] THEN rw[coeff_poly_sub] THEN simp[coeff_x_pow_times] THEN simp[coeff_const_x_pow_times] THEN simp[ARITH_RULE `~(e:num < 0)`] THEN case `e = 0` THENL [ simp[ARITH_RULE `0 - 0 = 0:num`] THEN simp[ARITH_RULE `0 < 1:num`] THEN qed[RING_SUB_RZERO;coeff_series_in_ring] ; simp[ARITH_RULE `~(e:num = 0) ==> ~(e < 1)`] THEN qed[ARITH_RULE `e - 0 = e:num`] ] );; let deg_one_minus_constx_le = prove(` !(r:R ring) c. c IN ring_carrier r ==> poly_deg r (one_minus_constx r c) <= 1 `, intro THEN have `ring_powerseries(r:R ring) (one_minus_constx r c)` [one_minus_constx_series] THEN have `!e. ~(coeff e (one_minus_constx(r:R ring) c) = ring_0 r) ==> e <= 1` [coeff_one_minus_constx;ARITH_RULE `0 <= 1:num`;ARITH_RULE `1 <= 1:num`] THEN qed[deg_le_coeff] );; let eval_one_minus_constx = prove(` !(r:R ring) c x. c IN ring_carrier r ==> x IN ring_carrier r ==> poly_eval r (one_minus_constx r c) x = ring_sub r (ring_1 r) (ring_mul r c x) `, intro THEN rw[one_minus_constx] THEN have `ring_polynomial(r:R ring) (x_pow r 0)` [x_pow_poly] THEN have `ring_polynomial(r:R ring) (const_x_pow r c 1)` [const_x_pow_poly] THEN simp[POLY_EVAL_SUB] THEN simp[eval_const_x_pow;eval_x_pow] THEN simp[RING_POW_1;RING_POW_0] );; let nonzero_one_minus_constx = prove(` !(r:R ring) c. ~(ring_1 r = ring_0 r) ==> c IN ring_carrier r ==> ~(one_minus_constx r c = poly_0 r) `, intro THEN have `coeff 0 (one_minus_constx(r:R ring) c) = ring_1 r` [coeff_one_minus_constx] THEN have `coeff 0 (poly_0 r:(1->num)->R) = ring_0 r` [coeff_poly_0] THEN qed[] );; (* ===== infinite geometric series *) (* in r[[x]]; different topology from SUMS_GP etc. *) (* infinite_geometric_series r c = sum c^n x^n *) let infinite_geometric_series = new_definition ` infinite_geometric_series (r:R ring) (c) = \n. (ring_pow r c (n one)) `;; let infinite_geometric_series_powerseries = prove(` !(r:R ring) c. c IN ring_carrier r ==> ring_powerseries r (infinite_geometric_series r c) `, intro THEN rw[infinite_geometric_series;ring_powerseries] THEN qed[RING_POW;FINITE_MONOMIAL_VARS_1;INFINITE] );; (* sum c^n x^n is in r[[x]] *) let infinite_geometric_series_x_series = prove(` !(r:R ring) c. c IN ring_carrier r ==> infinite_geometric_series r c IN ring_carrier(x_series r) `, rw[x_series;POWSER_RING;poly_vars;UNIONS_SUBSET] THEN SET_TAC[monomial;infinite_geometric_series_powerseries] );; let coeff_infinite_geometric_series = prove(` !(r:R ring) c e. coeff e (infinite_geometric_series r c) = ring_pow r c e `, rw[infinite_geometric_series;coeff;x_monomial] );; (* (1-cx) sum c^n x^n = 1 *) let infinite_geometric_series_inverse = prove(` !(r:R ring) c. c IN ring_carrier r ==> poly_mul r (one_minus_constx r c) (infinite_geometric_series r c) = poly_1 r `, intro THEN have `ring_powerseries(r:R ring) (infinite_geometric_series r c)` [infinite_geometric_series_powerseries] THEN sufficesby eq_coeff THEN simp[coeff_one_minus_constx_times] THEN rw[coeff_infinite_geometric_series] THEN rw[coeff_poly_1] THEN rw[RING_POW_0] THEN intro THEN have `ring_pow(r:R ring) c d IN ring_carrier r` [RING_POW] THEN have `ring_sub(r:R ring) (ring_pow r c d) (ring_pow r c d) = ring_0 r` [RING_SUB_REFL] THEN qed[ring_pow_sub1] );; (* ===== poly_pow *) let poly_pow = new_definition ` poly_pow (r:R ring) = ring_pow (powser_ring r (:V)) `;; let poly_pow_series = prove(` !(r:R ring) p:(V->num)->R n. ring_powerseries r p ==> ring_powerseries r (poly_pow r p n) `, rw[poly_pow] THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ intro THEN rw[RING_POW_0;POWSER_RING] THEN qed[RING_POWERSERIES_1] ; intro THEN rw[ring_pow;POWSER_RING] THEN qed[RING_POWERSERIES_MUL] ] );; let poly_pow_poly = prove(` !(r:R ring) p:(V->num)->R n. ring_polynomial r p ==> ring_polynomial r (poly_pow r p n) `, rw[poly_pow] THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ intro THEN rw[RING_POW_0;POWSER_RING] THEN qed[RING_POLYNOMIAL_1] ; intro THEN rw[ring_pow;POWSER_RING] THEN qed[RING_POLYNOMIAL_MUL] ] );; let poly_pow_in_poly_ring = prove(` !(r:R ring) p:(V->num)->R n S. p IN ring_carrier(poly_ring r S) ==> poly_pow r p n IN ring_carrier(poly_ring r S) `, rw[poly_pow] THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ intro THEN rw[RING_POW_0;POWSER_RING] THEN qed[POLY_RING;RING_1] ; intro THEN rw[ring_pow;POWSER_RING] THEN qed[poly_mul_in_poly_ring] ] );; let x_series_use_pow = prove(` !(r:R ring). poly_pow r = ring_pow (x_series r) `, qed[poly_pow;x_series] );; let x_poly_use_pow = prove(` !(r:R ring). poly_pow r = ring_pow (x_poly r) `, rw[x_series_use_pow;x_poly;x_series] THEN qed[POLY_RING_AS_SUBRING;RING_POW_SUBRING_GENERATED] );; let poly_ring_use_pow = prove(` !(r:R ring). poly_pow r = ring_pow (poly_ring r (:V)) `, rw[poly_pow] THEN qed[POLY_RING_AS_SUBRING;RING_POW_SUBRING_GENERATED] );; let poly_pow_0 = prove(` !(r:R ring) p:(V->num)->R. poly_pow r p 0 = poly_1 r:(V->num)->R `, once_rw[poly_pow] THEN qed[POWSER_RING;RING_POW_0] );; let poly_pow_1 = prove(` !(r:R ring) p:(V->num)->R. ring_powerseries r p ==> poly_pow r p 1 = p `, intro THEN once_rw[poly_pow] THEN subgoal `(p:(V->num)->R) IN ring_carrier(powser_ring r (:V))` THENL [ rw[POWSER_RING;IN_ELIM_THM;SUBSET_UNIV] THEN simp[] ; pass ] THEN qed[RING_POW_1] );; let poly_1_pow = prove(` !(r:R ring) n. poly_pow r (poly_1 r) n = poly_1 r:(V->num)->R `, once_rw[poly_pow] THEN qed[POWSER_RING;RING_POW_ONE] );; let poly_pow_add = prove(` !(r:R ring) (p:(V->num)->R) m n. ring_powerseries r p ==> poly_pow r p (m+n) = poly_mul r (poly_pow r p m) (poly_pow r p n) `, once_rw[poly_pow] THEN intro THEN have `(p:(V->num)->R) IN ring_carrier(powser_ring r (:V))` [series_in_full_ring] THEN simp[RING_POW_ADD] THEN qed[POWSER_RING] );; let x_pow_pow = prove(` !(r:R ring) m n. poly_pow r (x_pow r m) n = x_pow r (m*n) `, GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;x_pow_0;ARITH_RULE `m * 0 = 0`] ; rw[ARITH_RULE `SUC n = n+1`;ARITH_RULE `m*(n+1) = m*n+m`] THEN simp[poly_pow_add;poly_pow_1;x_pow_series] THEN qed[x_pow_add] ] );; let poly_pow_mul = prove(` !(r:R ring) (p:(V->num)->R) m n. ring_powerseries r p ==> poly_pow r (poly_pow r p m) n = poly_pow r p (m*n) `, once_rw[poly_pow] THEN intro THEN have `(p:(V->num)->R) IN ring_carrier(powser_ring r (:V))` [series_in_full_ring] THEN simp[RING_POW_MUL] );; let poly_mul_pow = prove(` !(r:R ring) (p:(V->num)->R) q n. ring_powerseries r p ==> ring_powerseries r q ==> poly_pow r (poly_mul r p q) n = poly_mul r (poly_pow r p n) (poly_pow r q n) `, once_rw[poly_pow] THEN intro THEN have `(p:(V->num)->R) IN ring_carrier(powser_ring r (:V))` [series_in_full_ring] THEN have `(q:(V->num)->R) IN ring_carrier(powser_ring r (:V))` [series_in_full_ring] THEN have `poly_mul(r:R ring) = ring_mul(powser_ring r (:V))` [POWSER_RING] THEN simp[RING_MUL_POW] );; let poly_deg_pow_le = prove(` !(r:R ring) (p:(V->num)->R) n. ring_polynomial r p ==> poly_deg r (poly_pow r p n) <= n * poly_deg r p `, GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0] THEN rw[POLY_DEG_1] THEN ARITH_TAC ; intro THEN have `ring_polynomial r (poly_pow r (p:(V->num)->R) n)` [poly_pow_poly] THEN rw[poly_pow] THEN rw[ring_pow] THEN rw[GSYM poly_pow;POWSER_RING] THEN have `poly_deg r (poly_mul r (p:(V->num)->R) (poly_pow r p n)) <= poly_deg r p + poly_deg r (poly_pow r p n)` [POLY_DEG_MUL_LE] THEN have `poly_deg r (poly_pow r (p:(V->num)->R) n) <= n * poly_deg r p` [] THEN ASM_ARITH_TAC ] );; let coeff_pow_infinite_geometric_series = prove(` !(r:R ring) c e n. c IN ring_carrier r ==> coeff n (poly_pow r (infinite_geometric_series r c) (e+1)) = ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r c n) `, rw[x_series_use_pow] THEN GEN_TAC THEN GEN_TAC THEN sufficesby num_WF THEN intro THEN case `e = 0` THENL [ simp[ARITH_RULE `0 + 1 = 1`;RING_POW_1;infinite_geometric_series_x_series] THEN simp[binom;RING_OF_NUM_1;RING_MUL_LID;RING_POW] THEN qed[coeff_infinite_geometric_series] ; pass ] THEN rw[ARITH_RULE `e+1 = SUC e`;ring_pow] THEN rw[GSYM x_series_use;coeff_poly_mul_oneindex] THEN rw[coeff_infinite_geometric_series] THEN subgoal `ring_sum(r:R ring) (0..n) (\a. ring_mul r (ring_pow r c a) (coeff (n - a) (ring_pow (x_series r) (infinite_geometric_series r c) e))) = ring_sum r (0..n) (\a. ring_mul r (ring_pow r c a) (ring_mul r (ring_of_num r (binom((n-a)+(e-1),(e-1)))) (ring_pow r c (n-a))))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN num_linear_fact `~(e = 0) ==> (e-1) < e` THEN have `coeff (n-a) (ring_pow (x_series(r:R ring)) (infinite_geometric_series r c) ((e-1) + 1)) = ring_mul r (ring_of_num r (binom ((n-a) + (e-1),e-1))) (ring_pow r c (n-a))` [] THEN num_linear_fact `~(e = 0) ==> (e-1)+1 = e` THEN have `coeff (n-a) (ring_pow (x_series(r:R ring)) (infinite_geometric_series r c) e) = ring_mul r (ring_of_num r (binom ((n-a) + (e-1),e-1))) (ring_pow r c (n-a))` [] THEN qed[] ; pass ] THEN subgoal `ring_sum(r:R ring) (0..n) (\a. ring_mul r (ring_pow r c a) (ring_mul r (ring_of_num r (binom((n-a)+(e-1),(e-1)))) (ring_pow r c (n-a)))) = ring_sum r (0..n) (\a. ring_mul r (ring_of_num r (binom((n-a)+(e-1),(e-1)))) (ring_pow r c n))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN have `a <= n:num` [IN_NUMSEG_0] THEN num_linear_fact `a <= n:num ==> a + (n-a) = n` THEN have `ring_mul(r:R ring) (ring_pow r c a) (ring_pow r c (n-a)) = ring_pow r c n` [RING_POW_ADD] THEN qed[RING_RULE `ring_mul(r:R ring) A M = N ==> ring_mul r A (ring_mul r B M) = ring_mul r B N`;RING_POW;RING_OF_NUM] ; pass ] THEN subgoal `ring_sum(r:R ring) (0..n) (\a. ring_mul r (ring_of_num r (binom (n - a + e - 1,e - 1))) (ring_pow r c n)) = ring_mul r (ring_sum r (0..n) (\a. ring_of_num r (binom (n - a + e - 1,e - 1)))) (ring_pow r c n)` THENL [ sufficesby RING_SUM_RMUL THEN qed[RING_POW;FINITE_NUMSEG;RING_OF_NUM] ; pass ] THEN subgoal `ring_sum(r:R ring) (0..n) (\a. ring_of_num r (binom (n - a + e - 1,e - 1))) = ring_of_num r (binom (n + e,e))` THENL [ specialize[`e:num`;`n:num`]binom_reverse_stair_sum THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN specialize[`r:R ring`;`\a. binom(n-a+e-1,e-1)`;`0..n`]ring_sum_num THEN qed[] ; pass ] THEN simp[] );; let pow_infinite_geometric_series = prove(` !(r:R ring) c e. c IN ring_carrier r ==> poly_pow r (infinite_geometric_series r c) (e+1) = series_from_coeffs (\n. ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r c n)) `, qed[coeff_pow_infinite_geometric_series;eq_coeff;coeff_series_from_coeffs] );; let pow_infinite_geometric_series_inverse_lemma = prove(` !(r:R ring) c e. c IN ring_carrier r ==> poly_mul r (poly_pow r (one_minus_constx r c) (e+1)) (poly_pow r (infinite_geometric_series r c) (e+1)) = poly_1 r `, intro THEN have `ring_powerseries(r:R ring) (one_minus_constx r c)` [one_minus_constx_series] THEN have `ring_powerseries(r:R ring) (infinite_geometric_series r c)` [infinite_geometric_series_powerseries] THEN simp[GSYM poly_mul_pow] THEN simp[infinite_geometric_series_inverse] THEN qed[poly_1_pow] );; (* (1-cx)^(e+1) sum binom(n+e,e) c^n x^n = 1 *) let pow_infinite_geometric_series_inverse = prove(` !(r:R ring) c e. c IN ring_carrier r ==> poly_mul r (poly_pow r (one_minus_constx r c) (e+1)) (series_from_coeffs (\n. ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r c n))) = poly_1 r `, intro THEN qed[pow_infinite_geometric_series_inverse_lemma;pow_infinite_geometric_series] );; let poly_pow_o_permutes = prove(` !(r:R ring) p:(V->num)->R S f n m. f permutes S ==> ring_powerseries r p ==> poly_pow r p n (m o f) = poly_pow r (\m. p (m o f)) n m `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0] THEN qed[poly_1_o_permutes] ; rw[ARITH_RULE `SUC n = 1 + n`] THEN intro THEN have `ring_powerseries(r:R ring) (\m:V->num. p(m o (f:V->V)))` [powerseries_o_permutes] THEN simp[poly_pow_1;poly_pow_add] THEN specialize[ `r:R ring`; `p:(V->num)->R`; `poly_pow r p n:(V->num)->R`; `S:V->bool`; `f:V->V` ]poly_mul_o_permutes THEN subgoal `poly_pow(r:R ring) (\m. p(m o f)) n = \m. poly_pow r p n (m o (f:V->V))` THENL [ rw[FUN_EQ_THM] THEN qed[] ; pass ] THEN qed[] ] );; let eval_poly_pow_multi = prove(` !(r:R ring) (p:(V->num)->R) c U n. p IN ring_carrier(poly_ring r U) ==> (!v. v IN U ==> c v IN ring_carrier r) ==> poly_evaluate r (poly_pow r p n) c = ring_pow r (poly_evaluate r p c) n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;RING_POW_0;POLY_EVALUATE_1] ; rw[ARITH_RULE `SUC n = 1+n`] THEN intro THEN have `ring_polynomial r (p:(V->num)->R)` [POLY_RING;IN_ELIM_THM] THEN have `ring_powerseries r (p:(V->num)->R)` [ring_polynomial] THEN simp[poly_pow_add;poly_pow_1;RING_POW_ADD;RING_POW_1;POLY_EVALUATE] THEN have `ring_polynomial r (poly_pow r (p:(V->num)->R) n)` [poly_pow_poly] THEN subgoal `!v:V. v IN poly_vars(r:R ring) p UNION poly_vars r (poly_pow r p n) ==> c v IN ring_carrier r` THENL [ rw[IN_UNION] THEN subgoal `poly_vars r (p:(V->num)->R) SUBSET U` THENL [ have `(p:(V->num)->R) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET U}` [POLY_RING] THEN set_fact `(p:(V->num)->R) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET U} ==> poly_vars r p SUBSET U` THEN qed[] ; pass ] THEN subgoal `poly_vars r (poly_pow r (p:(V->num)->R) n) SUBSET U` THENL [ have `poly_pow r (p:(V->num)->R) n IN {q | ring_polynomial r q /\ poly_vars r q SUBSET U}` [poly_pow_in_poly_ring;POLY_RING] THEN set_fact `poly_pow r (p:(V->num)->R) n IN {q | ring_polynomial r q /\ poly_vars r q SUBSET U} ==> poly_vars r (poly_pow r p n) SUBSET U` THEN qed[] ; pass ] THEN intro THEN have `v:V IN U` [SUBSET] THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `p:(V->num)->R`; `poly_pow r p n:(V->num)->R`; `c:V->R` ]POLY_EVALUATE_MUL THEN rw[know `poly_evaluate(r:R ring) (poly_mul r p (poly_pow r p n)) (c:V->R) = ring_mul r (poly_evaluate r p c) (poly_evaluate r (poly_pow r p n) c)`] THEN qed[] ] );; let poly_const_pow = prove(` !(r:R ring) c n. c IN ring_carrier r ==> poly_const r (ring_pow r c n) = poly_pow r (poly_const r c) n:(V->num)->R `, GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[RING_POW_0;poly_pow_0;poly_1] ; rw[ring_pow] THEN rw[ARITH_RULE `SUC n = 1+n`] THEN intro THEN have `ring_powerseries r (poly_const r c:(V->num)->R)` [RING_POWERSERIES_CONST] THEN simp[poly_pow_add;poly_pow_1] THEN simp[POLY_CONST_MUL;RING_POW] ] );; (* ===== x_monomial_shift: x^d |-> x^(d+1) *) let x_monomial_shift = new_definition ` x_monomial_shift (m:1->num) = (\v. m v + 1) `;; let x_monomial_shift_eq_x_monomial = prove(` !m d. x_monomial_shift m = x_monomial d <=> (?c. d = c + 1 /\ m = x_monomial c) `, rw[x_monomial_shift;x_monomial;FUN_EQ_THM] THEN intro THEN splitiff THENL [ intro THEN witness `m one:num` THEN qed[ETA_ONE;one] ; intro THEN qed[] ] );; let x_monomial_shift_is_not_monomial_1 = prove(` !m. ~(x_monomial_shift m = monomial_1) `, rw[x_monomial_shift;monomial_1;FUN_EQ_THM] THEN ARITH_TAC );; let x_monomial_shift_mul = prove(` !m n. monomial_mul (x_monomial_shift m) n = x_monomial_shift (monomial_mul m n) `, rw[FUN_EQ_THM] THEN rw[monomial_mul;x_monomial_shift] THEN ARITH_TAC );; let x_monomial_mul_shift = prove(` !m n. monomial_mul m (x_monomial_shift n) = x_monomial_shift (monomial_mul m n) `, rw[FUN_EQ_THM] THEN rw[monomial_mul;x_monomial_shift] THEN ARITH_TAC );; let x_monomial_shift_injective = prove(` !m n. x_monomial_shift m = x_monomial_shift n ==> m = n `, rw[x_monomial_shift;FUN_EQ_THM] THEN qed[ARITH_RULE `m + 1 = n + 1 ==> m = n`] );; (* ===== univariate derivatives *) let x_derivative = new_definition ` x_derivative (r:R ring) (p:(1->num)->R) = (\m. ring_mul r (ring_of_num r (m one + 1)) (p (x_monomial_shift m))) `;; let x_derivative_series = prove(` !(r:R ring) p. ring_powerseries r p ==> ring_powerseries r (x_derivative r p) `, rw[ring_powerseries;x_derivative] THEN qed[RING_OF_NUM;RING_MUL;FINITE_MONOMIAL_VARS_1;INFINITE] );; let x_derivative_polynomial = prove(` !(r:R ring) p. ring_polynomial r p ==> ring_polynomial r (x_derivative r p) `, rw[ring_polynomial] THEN intro THENL [ qed[x_derivative_series] ; rw[x_derivative] THEN have `!m. ~(ring_mul(r:R ring) (ring_of_num r (m one + 1)) (p (x_monomial_shift m)) = ring_0 r) ==> ~(p (x_monomial_shift m) = ring_0 r)` [RING_MUL_RZERO;RING_OF_NUM] THEN set_fact `(!m. ~(ring_mul(r:R ring) (ring_of_num r (m one + 1)) (p (x_monomial_shift m)) = ring_0 r) ==> ~(p (x_monomial_shift m) = ring_0 r)) ==> {m | ~(ring_mul(r:R ring) (ring_of_num r (m one + 1)) (p (x_monomial_shift m)) = ring_0 r)} SUBSET {m | ~(p (x_monomial_shift m) = ring_0 r)}` THEN have `{m | ~(ring_mul(r:R ring) (ring_of_num r (m one + 1)) (p (x_monomial_shift m)) = ring_0 r)} SUBSET {m | ~(p (x_monomial_shift m) = ring_0 r)}` [] THEN specialize_assuming[`x_monomial_shift`;`{m:1->num | ~(p m = ring_0(r:R ring))}`]FINITE_IMAGE_INJ THEN have `FINITE {m | x_monomial_shift m IN {m | ~(p m = ring_0(r:R ring))}}` [x_monomial_shift_injective] THEN set_fact `{m | x_monomial_shift m IN {m | ~(p m = ring_0(r:R ring))}} = {m | ~(p (x_monomial_shift m) = ring_0(r:R ring))}` THEN have `FINITE {m | ~(p (x_monomial_shift m) = ring_0(r:R ring))}` [] THEN qed[FINITE_SUBSET] ] );; let coeff_x_derivative = prove(` !(r:R ring) p d. coeff d (x_derivative r p) = ring_mul r (ring_of_num r (d+1)) (coeff (d+1) p) `, intro THEN rw[x_derivative;coeff] THEN have `x_monomial_shift (x_monomial d) = x_monomial (d+1)` [x_monomial_shift_eq_x_monomial] THEN simp[] THEN rw[x_monomial] );; let x_derivative_add_series = prove(` !(r:R ring) p q. ring_powerseries r p ==> ring_powerseries r q ==> x_derivative r (poly_add r p q) = poly_add r (x_derivative r p) (x_derivative r q) `, intro THEN have `!m:1->num. p m IN ring_carrier(r:R ring)` [ring_powerseries] THEN have `!m:1->num. q m IN ring_carrier(r:R ring)` [ring_powerseries] THEN rw[x_derivative;poly_add] THEN once_rw[FUN_EQ_THM] THEN simp[RING_ADD_LDISTRIB;RING_OF_NUM] );; let x_derivative_neg_series = prove(` !(r:R ring) p. ring_powerseries r p ==> x_derivative r (poly_neg r p) = poly_neg r (x_derivative r p) `, intro THEN have `!m:1->num. p m IN ring_carrier(r:R ring)` [ring_powerseries] THEN rw[x_derivative;poly_neg] THEN once_rw[FUN_EQ_THM] THEN simp[RING_MUL_RNEG;RING_OF_NUM] );; let x_derivative_sub_series = prove(` !(r:R ring) p q. ring_powerseries r p ==> ring_powerseries r q ==> x_derivative r (poly_sub r p q) = poly_sub r (x_derivative r p) (x_derivative r q) `, rw[POLY_SUB] THEN qed[RING_POWERSERIES_NEG;x_derivative_add_series;x_derivative_neg_series] );; let x_derivative_poly_const = prove(` !(r:R ring) c. x_derivative r (poly_const r c) = poly_0 r `, rw[x_derivative;poly_0;poly_const;x_monomial_shift_is_not_monomial_1;COND_ID] THEN once_rw[FUN_EQ_THM] THEN qed[FUN_EQ_THM;RING_OF_NUM;RING_MUL_RZERO] );; let x_derivative_poly_0 = prove(` !(r:R ring). x_derivative r (poly_0 r) = poly_0 r `, qed[poly_0;x_derivative_poly_const] );; let x_derivative_poly_1 = prove(` !(r:R ring). x_derivative r (poly_1 r) = poly_0 r `, qed[poly_1;x_derivative_poly_const] );; let x_derivative_poly_const_mul_series = prove(` !(r:R ring) c p. c IN ring_carrier r ==> ring_powerseries r p ==> x_derivative r (poly_mul r (poly_const r c) p) = poly_mul r (poly_const r c) (x_derivative r p) `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_x_derivative] THEN have `ring_powerseries(r:R ring) (x_derivative r p)` [x_derivative_series] THEN simp[coeff_poly_const_times] THEN rw[coeff_x_derivative] THEN have `coeff (d+1) p IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `ring_of_num r (d+1) IN ring_carrier(r:R ring)` [RING_OF_NUM] THEN RING_TAC );; let x_derivative_const_x_pow = prove(` !(r:R ring) e. c IN ring_carrier r ==> x_derivative r (const_x_pow r c e) = const_x_pow r (ring_mul r (ring_of_num r e) c) (e-1) `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_x_derivative] THEN rw[coeff_const_x_pow] THEN case `d+1 = e:num` THENL [ num_linear_fact `d+1 = e:num ==> d = e-1` THEN qed[] ; have `ring_of_num(r:R ring) (d+1) IN ring_carrier r` [RING_OF_NUM] THEN case `e = 0` THENL [ have `ring_of_num(r:R ring) e = ring_0 r` [RING_OF_NUM_0] THEN qed[RING_MUL_RZERO;RING_MUL_LZERO] ; num_linear_fact `~(d+1 = e:num) /\ ~(e = 0) ==> ~(d = e-1)` THEN qed[RING_MUL_RZERO] ] ] );; let x_derivative_x_pow = prove(` !(r:R ring) e. x_derivative r (x_pow r e) = const_x_pow r (ring_of_num r e) (e-1) `, intro THEN rw[x_pow] THEN have `ring_1(r:R ring) IN ring_carrier r` [RING_1] THEN qed[x_derivative_const_x_pow;RING_MUL_RID;RING_OF_NUM] );; let coeff_x_derivative_poly_mul = prove(` !(r:R ring) p q d. ring_powerseries r p ==> ring_powerseries r q ==> coeff d (x_derivative r (poly_mul r p q)) = ring_add r (coeff d (poly_mul r (x_derivative r p) q)) (coeff d (poly_mul r p (x_derivative r q))) `, rw[coeff_x_derivative;coeff_poly_mul_oneindex] THEN intro THEN have `!i. coeff i p IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `!i. coeff i q IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `ring_0(r:R ring) = ring_mul r (ring_mul r (ring_of_num r 0) (coeff 0 p)) (coeff ((d + 1) - 0) q)` [RING_OF_NUM_0;RING_MUL_LZERO] THEN specialize[`r:R ring`;`\a. ring_mul(r:R ring) (ring_mul r (ring_of_num r a) (coeff a p)) (coeff ((d+1)-a) q)`;`d:num`](GSYM ring_sum_shift1) THEN specialize[`0`;`d+1`]FINITE_NUMSEG THEN have `ring_of_num(r:R ring) (d+1) IN ring_carrier r` [RING_OF_NUM] THEN have `!a. a IN (0..d+1) ==> ring_mul r (coeff a p) (coeff ((d + 1) - a) q) IN ring_carrier(r:R ring)` [RING_MUL] THEN set_fact_using `!a:num. a IN (0..d) ==> a <= d` [NUMSEG_LE] THEN num_linear_fact `!a:num. a <= d ==> d - a = (d+1)-(a+1)` THEN have `!a:num. a IN (0..d) ==> d - a = (d+1)-(a+1)` [] THEN specialize[`r:R ring`;`\a. ring_mul (r:R ring) (coeff a p) (coeff ((d+1)-a) q)`;`ring_of_num(r:R ring) (d+1)`;`0..d+1`](GSYM RING_SUM_LMUL) THEN num_linear_fact `!a:num. a <= d ==> d-a+1 = (d+1)-a` THEN have `!a:num. a IN (0..d) ==> d-a+1 = (d+1)-a` [] THEN have `ring_of_num r ((d+1)-(d+1)) = ring_0(r:R ring)` [RING_OF_NUM_0;ARITH_RULE `(d+1)-(d+1)=0`] THEN have `ring_0(r:R ring) = ring_mul r (coeff (d+1) p) (ring_mul r (ring_of_num r ((d+1)-(d+1))) (coeff ((d+1)-(d+1)) q))` [RING_OF_NUM;RING_MUL_LZERO;RING_MUL_RZERO] THEN specialize[`r:R ring`;`\a. ring_mul(r:R ring) (coeff a p) (ring_mul r (ring_of_num r ((d+1)-a)) (coeff ((d+1)-a) q))`;`d:num`](GSYM ring_sum_insert_top) THEN simp[] THEN have `!a. ring_mul r (ring_mul r (ring_of_num r a) (coeff a p)) (coeff ((d + 1) - a) q) IN ring_carrier(r:R ring)` [RING_OF_NUM;RING_MUL] THEN have `!a. ring_mul r (coeff a p) (ring_mul r (ring_of_num r ((d + 1) - a)) (coeff ((d + 1) - a) q)) IN ring_carrier(r:R ring)` [RING_OF_NUM;RING_MUL] THEN simp[GSYM RING_SUM_ADD] THEN have `!a. ring_add(r:R ring) (ring_mul r (ring_mul r (ring_of_num r a) (coeff a p)) (coeff ((d+1)-a) q)) (ring_mul r (coeff a p) (ring_mul r (ring_of_num r ((d+1)-a)) (coeff ((d+1)-a) q))) = ring_mul r (ring_add r (ring_of_num r a) (ring_of_num r ((d+1)-a))) (ring_mul r (coeff a p) (coeff ((d+1)-a) q))` [RING_RULE `ring_add(r:R ring) (ring_mul r (ring_mul r (ring_of_num r a) (coeff a p)) (coeff ((d+1)-a) q)) (ring_mul r (coeff a p) (ring_mul r (ring_of_num r ((d+1)-a)) (coeff ((d+1)-a) q))) = ring_mul r (ring_add r (ring_of_num r a) (ring_of_num r ((d+1)-a))) (ring_mul r (coeff a p) (coeff ((d+1)-a) q))`] THEN simp[GSYM RING_OF_NUM_ADD] THEN num_linear_fact `!a. a <= d + 1 ==> a + (d+1) - a = d+1` THEN set_fact_using `!a. a IN (0..d+1) ==> a <= d+1` [NUMSEG_LE] THEN have `!a. a IN (0..d+1) ==> a+(d+1)-a = d+1` [NUMSEG_LE] THEN simp[] );; let x_derivative_mul = prove(` !(r:R ring) p q. ring_powerseries r p ==> ring_powerseries r q ==> x_derivative r (poly_mul r p q) = poly_add r (poly_mul r (x_derivative r p) q) (poly_mul r p (x_derivative r q)) `, intro THEN sufficesby eq_coeff THEN intro THEN simp[coeff_x_derivative_poly_mul] THEN rw[coeff_poly_add] );; let x_derivative_mul_const = prove(` !(r:R ring) c q. c IN ring_carrier r ==> ring_powerseries r q ==> x_derivative r (poly_mul r (poly_const r c) q) = poly_mul r (poly_const r c) (x_derivative r q) `, intro THEN have `ring_powerseries(r:R ring) ((poly_const r c):(1->num)->R)` [RING_POWERSERIES_CONST] THEN have `ring_powerseries(r:R ring) (x_derivative r q)` [x_derivative_series] THEN have `ring_powerseries(r:R ring) (poly_mul r (poly_const r c) (x_derivative r q))` [RING_POWERSERIES_MUL] THEN simp[x_derivative_mul] THEN simp[x_derivative_poly_const] THEN simp[POLY_MUL_0;RING_MUL_LZERO] THEN simp[POWSER_MUL_0] THEN qed[POLY_ADD_LZERO] );; let x_derivative_x_minus_const = prove(` !(r:R ring) c. c IN ring_carrier r ==> x_derivative r (x_minus_const r c) = poly_1 r `, intro THEN rw[x_minus_const] THEN have `ring_powerseries(r:R ring) (x_pow r 1)` [x_pow_series] THEN have `ring_powerseries(r:R ring) ((poly_const r c):(1->num)->R)` [RING_POWERSERIES_CONST] THEN simp[x_derivative_sub_series] THEN simp[x_derivative_x_pow] THEN simp[x_derivative_poly_const] THEN rw[RING_OF_NUM_1] THEN rw[ARITH_RULE `1-1 = 0:num`] THEN rw[const_x_pow_0;poly_1] THEN qed[poly_sub_0;POLY_CONST;RING_1;RING_POWERSERIES_CONST] );; let x_derivative_one_minus_constx = prove(` !(r:R ring) c. c IN ring_carrier r ==> x_derivative r (one_minus_constx r c) = poly_const r (ring_neg r c) `, intro THEN rw[one_minus_constx] THEN have `ring_powerseries(r:R ring) (x_pow r 0)` [x_pow_series] THEN have `ring_powerseries(r:R ring) (const_x_pow r c 1)` [const_x_pow_series] THEN simp[x_derivative_sub_series] THEN simp[x_derivative_x_pow] THEN simp[x_derivative_const_x_pow] THEN rw[ARITH_RULE `0-1 = 0:num`] THEN rw[ARITH_RULE `1-1 = 0:num`] THEN rw[RING_OF_NUM_0] THEN rw[RING_OF_NUM_1] THEN rw[const_x_pow_0] THEN simp[RING_MUL_LID] THEN rw[GSYM poly_0] THEN have `ring_powerseries(r:R ring) ((poly_const r c):(1->num)->R)` [RING_POWERSERIES_CONST] THEN qed[poly_0_sub;POLY_CONST_NEG] );; let x_derivative_subring = prove(` !(r:R ring) G p. x_derivative (subring_generated r G) p = x_derivative r p `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_x_derivative] THEN rw[SUBRING_GENERATED] THEN rw[RING_OF_NUM_SUBRING_GENERATED] );; (* basically: derivative of sp/sq is derivative of p/q *) let x_derivative_ratio_scaling = prove(` !(r:R ring) p q s. ring_powerseries r p ==> ring_powerseries r q ==> ring_powerseries r s ==> poly_sub r ( poly_mul r ( x_derivative r (poly_mul r s p) ) ( poly_mul r s q ) ) ( poly_mul r ( poly_mul r s p ) ( x_derivative r (poly_mul r s q) ) ) = poly_mul r ( poly_pow r s 2 ) ( poly_sub r ( poly_mul r (x_derivative r p) (q) ) ( poly_mul r (p) (x_derivative r q) ) ) `, intro THEN simp[x_derivative_mul] THEN simp[x_series_use;x_series_sub_use;x_series_use_pow] THEN specialize_assuming[`x_series(r:R ring)` ;`x_derivative(r:R ring) p` ;`x_derivative(r:R ring) q` ;`x_derivative(r:R ring) s` ;`p:(1->num)->R` ;`q:(1->num)->R` ;`s:(1->num)->R` ](GENL[ `r:R ring`;`P:R`;`Q:R`;`S:R`;`p:R`;`q:R`;`s:R` ](RING_RULE `ring_sub(r:R ring) (ring_mul r (ring_add r (ring_mul r (S) p) (ring_mul r s (P))) (ring_mul r s q)) (ring_mul r (ring_mul r s p) (ring_add r (ring_mul r (S) q) (ring_mul r s (Q)))) = ring_mul r (ring_pow r s 2) (ring_sub r (ring_mul r (P) q) (ring_mul r p (Q)))` )) THEN qed[x_derivative_series;x_series_use] );; (* ===== sums of finite sequences of power series *) let poly_sum = new_definition ` poly_sum (r:R ring) (S:X->bool) (p:X->(1->num)->R) = ring_sum(x_series r) S p `;; let poly_sum_empty = prove(` !(r:R ring) p:X->(1->num)->R. poly_sum r {} p = poly_0 r `, rw[poly_sum] THEN rw[RING_SUM_CLAUSES] THEN qed[x_series_use] );; let poly_sum_insert = prove(` !(r:R ring) p:X->(1->num)->R S t. FINITE S ==> poly_sum r (t INSERT S) p = (if ring_powerseries r (p t) ==> t IN S then poly_sum r S p else poly_add r (p t) (poly_sum r S p)) `, rw[poly_sum] THEN rw[x_series_use] THEN qed[RING_SUM_CLAUSES] );; let poly_sum_delete2 = prove(` !(r:R ring) S p:X->(1->num)->R t. FINITE S ==> t IN S ==> ring_powerseries r (p t) ==> poly_sum r S p = poly_add r (p t) (poly_sum r (S DELETE t) p) `, intro THEN rw[poly_sum;x_series_use] THEN have `p(t:X) IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN qed[ring_sum_delete2] );; let poly_sum_series = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> ring_powerseries r (poly_sum r S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty] THEN qed[RING_POWERSERIES_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [] THEN simp[poly_sum_insert] THEN qed[RING_POWERSERIES_ADD] ] );; (* XXX: merge into poly_sum_series *) let series_sum_series_multi = prove(` !(r:R ring) p:X->(V->num)->R S. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> ring_powerseries r (ring_sum(powser_ring r (:V)) S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;POWSER_RING;RING_POWERSERIES_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(V->num)->R)` [] THEN have `p(x:X):(V->num)->R IN ring_carrier(powser_ring r (:V))` [series_in_full_ring] THEN simp[RING_SUM_CLAUSES;POWSER_RING] THEN qed[RING_POWERSERIES_ADD] ] );; let poly_sum_ring_sum_x_poly = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> poly_sum r S p = ring_sum(x_poly r) S p `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty;RING_SUM_CLAUSES] THEN qed[x_poly_use] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [ring_polynomial] THEN have `(p(x:X):(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN simp[poly_sum_insert;RING_SUM_CLAUSES] THEN rw[x_poly_use] ] );; let poly_sum_poly = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> ring_polynomial r (poly_sum r S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty] THEN qed[RING_POLYNOMIAL_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(1->num)->R)` [] THEN simp[poly_sum_insert] THEN qed[RING_POLYNOMIAL_ADD] ] );; let poly_sum_eq = prove(` !(r:R ring) S:X->bool p:X->(1->num)->R q:X->(1->num)->R. (!s. s IN S ==> p s = q s) ==> poly_sum r S p = poly_sum r S q `, qed[poly_sum;RING_SUM_EQ] );; let coeff_poly_sum = prove(` !(r:R ring) p:X->(1->num)->R d:num S:X->bool. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> coeff d (poly_sum r S p) = ring_sum r S (\s. coeff d (p s)) `, rw[poly_sum] THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;GSYM x_series_use;coeff_poly_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [] THEN have `(p(x:X):(1->num)->R) IN ring_carrier(x_series r)` [x_series_use] THEN have `coeff d (p(x:X)) IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN simp[RING_SUM_CLAUSES] THEN rw[coeff_series_add] THEN qed[] ] );; let x_derivative_sum = prove(` !(r:R ring) p (S:X->bool). FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> x_derivative r (poly_sum r S p) = poly_sum r S (\s. x_derivative r (p s)) `, intro THEN sufficesby eq_coeff THEN intro THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (x_derivative r (p s))` [x_derivative_series] THEN simp[coeff_poly_sum] THEN rw[coeff_x_derivative] THEN simp[coeff_poly_sum] THEN have `ring_of_num(r:R ring) (d+1) IN ring_carrier r` [RING_OF_NUM] THEN have `!s:X. s IN S ==> coeff (d+1) (p s) IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN simp[RING_SUM_LMUL] );; let poly_deg_sum_le = prove(` !(r:R ring) (p:X->(1->num)->R) n S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> poly_deg r (p s) <= n) ==> poly_deg r (poly_sum r S p) <= n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty;POLY_DEG_0;ARITH_RULE `0 <= n`] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_sum_insert] THEN have `ring_polynomial(r:R ring) (poly_sum r (S:X->bool) p)` [poly_sum_poly] THEN qed[POLY_DEG_ADD_LE;ARITH_RULE `d <= n /\ e <= n /\ c <= MAX d e ==> c <= n:num`] ] );; let poly_sum_lmul = prove(` !(r:R ring) (p:X->(1->num)->R) c S. ring_powerseries r c ==> FINITE S ==> (!s:X. s IN S ==> ring_powerseries r (p s)) ==> poly_sum r S (\s. poly_mul r c (p s)) = poly_mul r c (poly_sum r S p) `, intro THEN rw[poly_sum;x_series_use] THEN have `c IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN have `!s:X. s IN S ==> p s IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN specialize[`x_series(r:R ring)`;`p:X->(1->num)->R`;`c:(1->num)->R`;`S:X->bool`]RING_SUM_LMUL THEN qed[] );; let poly_sum_const = prove(` !(r:R ring) (p:(1->num)->R) S. ring_powerseries r p ==> FINITE S ==> poly_sum r S (\s:X. p) = poly_mul r (poly_const r (ring_of_num r (CARD S))) p `, intro THEN rw[poly_sum] THEN have `p IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN simp[ring_sum_const;GSYM x_series_use] THEN qed[ring_of_num_x_series] );; (* XXX: merge into poly_sum_poly *) let poly_sum_poly_multi = prove(` !(r:R ring) p:X->(V->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> ring_polynomial r (ring_sum(poly_ring r (:V)) S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES] THEN simp[GSYM poly_use] THEN qed[RING_POLYNOMIAL_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(V->num)->R)` [] THEN simp[RING_SUM_CLAUSES] THEN simp[GSYM poly_use] THEN qed[RING_POLYNOMIAL_ADD] ] );; (* XXX: merge *) let poly_sum_subring_multi = prove(` !(r:R ring) G (p:X->(V->num)->R) S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial(subring_generated r G) (p s)) ==> ring_sum(poly_ring(subring_generated r G) (:V)) S p = ring_sum(poly_ring r (:V)) S p `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES] THEN rw[POLY_RING] THEN rw[poly_0;poly_const;SUBRING_GENERATED] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(subring_generated r G) (p(x:X):(V->num)->R)` [] THEN have `ring_polynomial r (p(x:X):(V->num)->R)` [ring_polynomial_if_subring] THEN have `!s:X. s IN S ==> ring_polynomial(subring_generated(r:R ring) G) (p s:(V->num)->R)` [] THEN simp[RING_SUM_CLAUSES;GSYM poly_use] THEN have `ring_polynomial (subring_generated r G) (ring_sum(poly_ring r (:V)) S (p:X->(V->num)->R))` [poly_sum_poly_multi] THEN qed[poly_add_subring;ring_polynomial] ] );; let ring_polynomial_subring_sum = prove(` !(r:R ring) G p:X->(V->num)->R S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial(subring_generated r G) (p s)) ==> ring_polynomial(subring_generated r G) (ring_sum(poly_ring r (:V)) S p) `, simp[GSYM poly_sum_subring_multi] THEN qed[RING_SUM;poly_in_full_ring] );; let poly_sum_restrict_subset = prove(` !(r:R ring) S:X->bool U p. S SUBSET U ==> poly_sum r U (\s. if s IN S then p s else poly_0 r) = poly_sum r S p `, rw[poly_sum;x_series_use] THEN intro THEN specialize[ `S:X->bool`; `x_series(r:R ring)`; `U:X->bool`; `p:X->(1->num)->R` ]RING_SUM_RESTRICT_SET THEN set_fact `S SUBSET U:X->bool ==> {x | x IN U /\ S x} = S` THEN rw[IN] THEN qed[] );; let poly_const_sum = prove(` !(r:R ring) f:X->R S. FINITE S ==> (!s. s IN S ==> f s IN ring_carrier r) ==> poly_const r (ring_sum r S f) = poly_sum r S (\s. poly_const r (f s)) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;poly_sum_empty] THEN rw[POLY_CONST_0] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN have `ring_powerseries r (poly_const r (f(x:X)):(1->num)->R)` [RING_POWERSERIES_CONST] THEN simp[RING_SUM_CLAUSES;poly_sum_insert] THEN qed[POLY_CONST_ADD;RING_SUM] ] );; (* ===== products of finite sequences of power series *) let poly_product = new_definition ` poly_product (r:R ring) (S:X->bool) (p:X->(1->num)->R) = ring_product(x_series r) S p `;; let poly_product_empty = prove(` !(r:R ring) p:X->(1->num)->R. poly_product r {} p = poly_1 r `, rw[poly_product] THEN rw[RING_PRODUCT_CLAUSES] THEN qed[x_series_use] );; let poly_product_insert = prove(` !(r:R ring) p:X->(1->num)->R S t. FINITE S ==> poly_product r (t INSERT S) p = (if ring_powerseries r (p t) ==> t IN S then poly_product r S p else poly_mul r (p t) (poly_product r S p)) `, rw[poly_product] THEN rw[x_series_use] THEN qed[RING_PRODUCT_CLAUSES] );; let poly_product_1 = prove(` !(r:R ring) S. poly_product r S (\s:X. poly_1 r) = poly_1 r `, rw[poly_product] THEN rw[x_series_use] THEN rw[RING_PRODUCT_1] );; let poly_product_eq = prove(` !(r:R ring) p:X->(1->num)->R q S. (!s. s IN S ==> p s = q s) ==> poly_product r S p = poly_product r S q `, rw[poly_product] THEN qed[RING_PRODUCT_EQ] );; (* XXX: merge into poly_product_series *) let poly_product_series_multi = prove(` !(r:R ring) p:X->(V->num)->R S. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> ring_powerseries r (ring_product(powser_ring r (:V)) S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POWSER_RING] THEN qed[RING_POWERSERIES_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(V->num)->R)` [] THEN simp[RING_PRODUCT_CLAUSES;POWSER_RING] THEN qed[RING_POWERSERIES_MUL] ] );; (* XXX: merge into poly_product_poly *) let poly_product_poly_multi = prove(` !(r:R ring) p:X->(V->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> ring_polynomial r (ring_product(poly_ring r (:V)) S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN simp[GSYM poly_use] THEN qed[RING_POLYNOMIAL_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(V->num)->R)` [] THEN simp[RING_PRODUCT_CLAUSES] THEN simp[GSYM poly_use] THEN qed[RING_POLYNOMIAL_MUL] ] );; let poly_product_series = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> ring_powerseries r (poly_product r S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty] THEN qed[RING_POWERSERIES_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [] THEN simp[poly_product_insert] THEN qed[RING_POWERSERIES_MUL] ] );; let poly_product_poly = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> ring_polynomial r (poly_product r S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty] THEN qed[RING_POLYNOMIAL_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(1->num)->R)` [] THEN simp[poly_product_insert] THEN qed[RING_POLYNOMIAL_MUL] ] );; let poly_product_ring_product_x_poly = prove(` !(r:R ring) p:X->(1->num)->R S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> poly_product r S p = ring_product(x_poly r) S p `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;RING_PRODUCT_CLAUSES] THEN qed[x_poly;POLY_RING] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(1->num)->R)` [] THEN have `(p(x:X):(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_product_insert;RING_PRODUCT_CLAUSES] THEN qed[x_poly;POLY_RING] ] );; let poly_product_pow = prove(` !(r:R ring) S (p:X->(1->num)->R) n. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> poly_pow r (poly_product r S p) n = poly_product r S (\s:X. poly_pow r (p s) n) `, intro THEN have `!s:X. s IN S ==> p s IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN have `!s:X. s IN S ==> p s IN ring_carrier(powser_ring (r:R ring) (:1))` [x_series] THEN rw[poly_pow;poly_product;x_series] THEN specialize[`powser_ring(r:R ring) (:1)`;`p:X->(1->num)->R`;`n:num`;`S:X->bool`]ring_pow_product THEN qed[] );; let poly_mul_sum_mul_delete = prove(` !(r:R ring) S:X->bool f:X->(1->num)->R g:X->(1->num)->R x. ~(x IN S) ==> FINITE S ==> ring_powerseries r (g x) ==> (!s. s IN S ==> ring_powerseries r (g s)) ==> (!s. s IN S ==> ring_powerseries r (f s)) ==> poly_mul r (g x) (poly_sum r S (\s. poly_mul r (f s) (poly_product r (S DELETE s) g))) = poly_sum r S (\s. poly_mul r (f s) (poly_product r ((x INSERT S) DELETE s) g)) `, rw[x_series_use] THEN rw[poly_sum] THEN rw[poly_product] THEN GEN_TAC THEN specialize_assuming[`x_series(r:R ring)`]ring_mul_sum_mul_delete THEN qed[] );; let x_derivative_product = prove(` !(r:R ring) p (S:X->bool). FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> x_derivative r (poly_product r S p) = poly_sum r S (\s. poly_mul r (x_derivative r (p s)) (poly_product r (S DELETE s) p)) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;poly_sum_empty] THEN rw[x_derivative_poly_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `~(x IN S) ==> ((x:X) INSERT S) DELETE x = S` THEN have `((x:X) INSERT S) DELETE x = S` [] THEN have `ring_powerseries(r:R ring) (x_derivative r (p(x:X)))` [x_derivative_series] THEN have `ring_powerseries(r:R ring) (poly_product r (S:X->bool) p)` [poly_product_series] THEN have `ring_powerseries(r:R ring) (poly_mul r (x_derivative r (p(x:X))) (poly_product r (S:X->bool) p))` [RING_POWERSERIES_MUL] THEN simp[poly_product_insert;poly_sum_insert] THEN simp[x_derivative_mul] THEN set_fact `!s:X. s IN S ==> ~(x IN S) ==> (x INSERT S) DELETE s = x INSERT (S DELETE s)` THEN have `!s:X. s IN S ==> ring_powerseries (r:R ring) ((p:X->(1->num)->R) s)` [IN_INSERT] THEN have `!s:X. FINITE(S DELETE s)` [FINITE_DELETE] THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) ((\s. x_derivative r (p s)) s)` [x_derivative_series] THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [] THEN specialize_assuming[`r:R ring`;`S:X->bool`;`\s:X. x_derivative(r:R ring) (p s)`;`p:X->(1->num)->R`;`x:X`]poly_mul_sum_mul_delete THEN qed[] ] );; let poly_product_const = prove(` !(r:R ring) (p:(1->num)->R) S. ring_powerseries r p ==> FINITE S ==> poly_product r S (\s:X. p) = poly_pow r p (CARD S) `, intro THEN rw[poly_pow;poly_product] THEN have `p IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN qed[ring_product_const;x_series] );; let poly_pow_is_product = prove(` !(r:R ring) (p:(1->num)->R) n. ring_powerseries r p ==> poly_pow r p n = poly_product r (1..n) (\i. p) `, intro THEN rw[poly_pow;poly_product] THEN have `p IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN qed[ring_pow_is_product;x_series] );; let x_derivative_pow = prove(` !(r:R ring) p n. ring_powerseries r p ==> x_derivative r (poly_pow r p n) = poly_mul r ( poly_const r (ring_of_num r n) ) ( poly_mul r ( x_derivative r p ) ( poly_pow r p (n-1) ) ) `, intro THEN case `n = 0` THENL [ simp[poly_pow_0;RING_OF_NUM_0] THEN rw[poly_1;x_derivative_poly_const] THEN rw[GSYM poly_0] THEN qed[POWSER_MUL_0;x_derivative_series;poly_pow_series;RING_POWERSERIES_MUL] ; pass ] THEN simp[poly_pow_is_product] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!i. i IN 1..n ==> ring_powerseries r (p:(1->num)->R)` [] THEN specialize[`r:R ring`;`\i:num. p:(1->num)->R`;`1..n`]x_derivative_product THEN subgoal `poly_sum(r:R ring) (1..n) (\s. poly_mul r (x_derivative r p) (poly_product r ((1..n) DELETE s) (\i. p))) = poly_sum r (1..n) (\s. poly_mul r (x_derivative r p) (poly_pow r p (n-1)))` THENL [ sufficesby poly_sum_eq THEN intro THEN rw[BETA_THM] THEN have `CARD(1..n) = n` [CARD_NUMSEG_1] THEN have `CARD((1..n) DELETE s) = CARD(1..n) - 1` [CARD_DELETE] THEN have `CARD((1..n) DELETE s) = n - 1` [] THEN have `FINITE((1..n) DELETE s)` [FINITE_DELETE] THEN simp[poly_product_const] ; pass ] THEN simp[poly_product_const;FINITE_NUMSEG;CARD_NUMSEG_1] THEN have `ring_powerseries(r:R ring) (poly_mul r (x_derivative r p) (poly_pow r p (n - 1)))` [RING_POWERSERIES_MUL;x_derivative_series;poly_pow_series] THEN specialize[`r:R ring`;`poly_mul(r:R ring) (x_derivative r p) (poly_pow r p (n - 1))`;`1..n`]poly_sum_const THEN have `CARD(1..n) = n` [CARD_NUMSEG_1] THEN qed[] );; let eval_poly_product = prove(` !(r:R ring) p:(X->(1->num)->R) z S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> z IN ring_carrier r ==> poly_eval r (poly_product r S p) z = ring_product r S (\s. poly_eval r (p s) z) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;RING_PRODUCT_CLAUSES] THEN qed[POLY_EVAL_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial r ((p:(X->(1->num)->R)) x)` [] THEN have `ring_powerseries r ((p:(X->(1->num)->R)) x)` [ring_polynomial] THEN have `poly_eval r ((p:(X->(1->num)->R)) x) z IN ring_carrier r` [POLY_EVAL] THEN simp[poly_product_insert;RING_PRODUCT_CLAUSES] THEN have `ring_polynomial r (poly_product r S (p:(X->(1->num)->R)))` [poly_product_poly] THEN simp[POLY_EVAL_MUL] ] );; (* XXX: merge into poly_product_subring; generalize other poly_product theorems *) let poly_product_subring_multi = prove(` !(r:R ring) G (p:X->(V->num)->R) S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial(subring_generated r G) (p s)) ==> ring_product(poly_ring(subring_generated r G) (:V)) S p = ring_product(poly_ring r (:V)) S p `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN rw[POLY_RING] THEN rw[poly_1;poly_const;SUBRING_GENERATED] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(subring_generated r G) (p(x:X):(V->num)->R)` [] THEN have `ring_polynomial r (p(x:X):(V->num)->R)` [ring_polynomial_if_subring] THEN have `!s:X. s IN S ==> ring_polynomial(subring_generated(r:R ring) G) (p s:(V->num)->R)` [] THEN simp[RING_PRODUCT_CLAUSES;GSYM poly_use] THEN have `ring_polynomial (subring_generated r G) (ring_product(poly_ring r (:V)) S (p:X->(V->num)->R))` [poly_product_poly_multi] THEN qed[poly_mul_subring;ring_polynomial] ] );; let poly_product_subring = prove(` !(r:R ring) G (p:X->(1->num)->R) S. FINITE S ==> (!s:X. s IN S ==> ring_powerseries(subring_generated r G) (p s)) ==> poly_product (subring_generated r G) S p = poly_product r S p `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty] THEN rw[poly_1;poly_const;SUBRING_GENERATED] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(subring_generated r G) (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_powerseries_subring] THEN have `!s:X. s IN S ==> ring_powerseries(subring_generated(r:R ring) G) (p s:(1->num)->R)` [] THEN simp[poly_product_insert] THEN have `ring_powerseries (subring_generated r G) (poly_product r S (p:X->(1->num)->R))` [poly_product_series] THEN qed[poly_mul_subring] ] );; let ring_polynomial_subring_product = prove(` !(r:R ring) G p:X->(V->num)->R S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial(subring_generated r G) (p s)) ==> ring_polynomial(subring_generated r G) (ring_product(poly_ring r (:V)) S p) `, simp[GSYM poly_product_subring_multi] THEN qed[RING_PRODUCT;poly_in_full_ring] );; let poly_product_image = prove(` !(r:R ring) S (f:X->Y) (g:Y->(1->num)->R). (!x y. x IN S ==> y IN S ==> f x = f y ==> x = y) ==> poly_product r (IMAGE f S) g = poly_product r S (g o f) `, rw[poly_product] THEN intro THEN specialize[`x_series(r:R ring)`;`f:X->Y`;`g:Y->(1->num)->R`;`S:X->bool`]RING_PRODUCT_IMAGE THEN qed[] );; let poly_deg_product_le = prove(` !(r:R ring) (p:X->(1->num)->R) n S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> poly_deg r (p s) <= n s) ==> poly_deg r (poly_product r S p) <= nsum S n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;NSUM_CLAUSES;POLY_DEG_1] THEN ARITH_TAC ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_product_insert;NSUM_CLAUSES] THEN simp[CARD_CLAUSES] THEN have `ring_polynomial(r:R ring) (poly_product r (S:X->bool) p)` [poly_product_poly] THEN have `poly_deg(r:R ring) (poly_product r (S:X->bool) p) <= nsum S n` [poly_product_poly] THEN qed[POLY_DEG_MUL_LE;ARITH_RULE `d <= dx + ds ==> dx <= nx ==> ds <= ns ==> d <= nx + ns:num`] ] );; (* should factor via poly_deg_product_le *) let poly_deg_product_each_le = prove(` !(r:R ring) (p:X->(1->num)->R) n S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> poly_deg r (p s) <= n) ==> poly_deg r (poly_product r S p) <= (CARD S) * n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;POLY_DEG_1] THEN ARITH_TAC ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_product_insert] THEN simp[CARD_CLAUSES] THEN have `ring_polynomial(r:R ring) (poly_product r (S:X->bool) p)` [poly_product_poly] THEN have `poly_deg(r:R ring) (poly_product r (S:X->bool) p) <= CARD S * n` [poly_product_poly] THEN qed[POLY_DEG_MUL_LE;ARITH_RULE `a <= b + c ==> b <= n ==> c <= S * n ==> a <= SUC(S) * n`] ] );; let poly_product_delete = prove(` !(r:R ring) S t f:X->((1->num)->R). FINITE S ==> t IN S ==> ring_powerseries r (f t) ==> poly_product r S f = poly_mul r (f t) (poly_product r (S DELETE t) f) `, intro THEN have `f(t:X) IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN have `f(t:X) IN ring_carrier(powser_ring(r:R ring) (:1))` [x_series] THEN rw[poly_product;x_series_use;x_series] THEN qed[ring_product_delete] );; let poly_pow_subring = prove(` !(r:R ring) G (p:(V->num)->R) n. ring_powerseries(subring_generated r G) p ==> poly_pow (subring_generated r G) p n = poly_pow r p n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;poly_1;poly_const] THEN simp[SUBRING_GENERATED] ; rw[ARITH_RULE `SUC n = 1+n`] THEN intro THEN specialize[`r:R ring`;`G:R->bool`;`p:(V->num)->R`]ring_powerseries_subring THEN simp[poly_pow_add;poly_pow_1] THEN qed[poly_mul_subring;poly_pow_series] ] );; let poly_product_botcoeff1 = prove(` !(r:R ring) (p:X->(1->num)->R) S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> coeff 0 (p s) = ring_1 r) ==> coeff 0 (poly_product r S p) = ring_1 r `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;coeff_poly_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_product_insert] THEN rw[coeff_poly_mul_oneindex] THEN simp[RING_SUM_CLAUSES_NUMSEG;ARITH_RULE `0-0=0`;RING_1;RING_MUL;RING_MUL_LID] ] );; let poly_product_in_poly_ring = prove(` !(r:R ring) (p:X->(V->num)->R) E S. FINITE S ==> (!s:X. s IN S ==> p s IN ring_carrier(poly_ring r E)) ==> ring_product(poly_ring r (:V)) S p IN ring_carrier(poly_ring r E) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN qed[RING_1;POLY_RING] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `(p(x:X):(V->num)->R) IN ring_carrier(poly_ring r E)` [] THEN have `(p(x:X):(V->num)->R) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET E}` [POLY_RING] THEN set_fact `(p(x:X):(V->num)->R) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET E} ==> ring_polynomial r (p x)` THEN have `(p(x:X):(V->num)->R) IN ring_carrier(poly_ring r (:V))` [RING_POLYNOMIAL] THEN simp[RING_PRODUCT_CLAUSES] THEN have `ring_product (poly_ring(r:R ring) (:V)) (S:X->bool) p IN ring_carrier (poly_ring r E)` [] THEN rw[CONJUNCT2 POLY_RING] THEN qed[poly_mul_in_poly_ring] ] );; let powser_product_o_permutes = prove(` !(r:R ring) p:X->(V->num)->R U f S. FINITE S ==> !m. f permutes U ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> ring_product(powser_ring r (:V)) S p (m o f) = ring_product(powser_ring r (:V)) S (\s m. (p s) (m o f)) m `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN rw[POWSER_RING] THEN qed[poly_1_o_permutes] ; intro THEN set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES] THEN have `ring_powerseries(r:R ring) (p(x:X):(V->num)->R)` [] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `f:V->V`; `U:V->bool` ]powerseries_o_permutes THEN have `p(x:X) IN ring_carrier(powser_ring(r:R ring) (:V))` [powser_use] THEN have `(\m:V->num. p(x:X) (m o (f:V->V))) IN ring_carrier(powser_ring(r:R ring) (:V))` [powser_use] THEN simp[] THEN rw[POWSER_RING] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_product (powser_ring(r:R ring) (:V)) (S:X->bool) p`; `U:V->bool`; `f:V->V`; `m:V->num` ]poly_mul_o_permutes THEN simp[] THEN subgoal `(\m:V->num. ring_product (powser_ring(r:R ring) (:V)) S (\s:X n:V->num. p s (n o (f:V->V))) m) = ring_product (powser_ring r (:V)) S (\s n. p s (n o f))` THENL [ rw[FUN_EQ_THM] ; pass ] THEN simp[] ] );; let poly_product_o_permutes = prove(` !(r:R ring) p:X->(V->num)->R U f S. FINITE S ==> !m. f permutes U ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> ring_product(poly_ring r (:V)) S p (m o f) = ring_product(poly_ring r (:V)) S (\s m. (p s) (m o f)) m `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES] THEN rw[POLY_RING] THEN qed[poly_1_o_permutes] ; intro THEN set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES] THEN have `ring_polynomial(r:R ring) (p(x:X):(V->num)->R)` [] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `f:V->V`; `U:V->bool` ]polynomial_o_permutes THEN have `p(x:X) IN ring_carrier(poly_ring(r:R ring) (:V))` [poly_use] THEN have `(\m:V->num. p(x:X) (m o (f:V->V))) IN ring_carrier(poly_ring(r:R ring) (:V))` [poly_use] THEN simp[] THEN rw[POLY_RING] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_product (poly_ring(r:R ring) (:V)) (S:X->bool) p`; `U:V->bool`; `f:V->V`; `m:V->num` ]poly_mul_o_permutes THEN simp[] THEN subgoal `(\m:V->num. ring_product (poly_ring(r:R ring) (:V)) S (\s:X n:V->num. p s (n o (f:V->V))) m) = ring_product (poly_ring r (:V)) S (\s n. p s (n o f))` THENL [ rw[FUN_EQ_THM] ; pass ] THEN simp[] ] );; (* XXX: merge into eval_poly_product *) let eval_poly_product_multi = prove(` !(r:R ring) (p:X->(V->num)->R) c U S. FINITE S ==> (!s:X. s IN S ==> p s IN ring_carrier(poly_ring r U)) ==> (!v. v IN U ==> c v IN ring_carrier r) ==> poly_evaluate r (ring_product(poly_ring r (:V)) S p) c = ring_product r S (\s. poly_evaluate r (p s) c) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POLY_RING;POLY_EVALUATE_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN specialize[`r:R ring`;`U:V->bool`](CONJUNCT1 POLY_RING) THEN have `p(x:X):(V->num)->R IN {q | ring_polynomial(r:R ring) q /\ poly_vars r q SUBSET U}` [] THEN set_fact `p(x:X):(V->num)->R IN {q | ring_polynomial(r:R ring) q /\ poly_vars r q SUBSET U} ==> ring_polynomial r (p x)` THEN have `p(x:X) IN ring_carrier(poly_ring(r:R ring) (:V))` [poly_use] THEN simp[RING_PRODUCT_CLAUSES;POLY_EVALUATE;POLY_RING] THEN have `ring_polynomial(r:R ring) (ring_product (poly_ring r (:V)) (S:X->bool) p)` [RING_PRODUCT;poly_use] THEN subgoal `!i. i IN poly_vars(r:R ring) (p(x:X)) UNION poly_vars r (ring_product (poly_ring r (:V)) S p) ==> c i IN ring_carrier r` THENL [ rw[IN_UNION] THEN set_fact `p(x:X):(V->num)->R IN {q | ring_polynomial(r:R ring) q /\ poly_vars r q SUBSET U} ==> poly_vars r (p x) SUBSET U` THEN subgoal `ring_product (poly_ring(r:R ring) (:V)) (S:X->bool) p IN ring_carrier(poly_ring r U)` THENL [ have `!s:X. s IN S ==> p s:(V->num)->R IN ring_carrier (poly_ring r U)` [] THEN specialize[ `r:R ring`; `p:X->(V->num)->R`; `U:V->bool`; `S:X->bool` ]poly_product_in_poly_ring THEN qed[] ; pass ] THEN have `ring_product (poly_ring r (:V)) (S:X->bool) p IN {q | ring_polynomial(r:R ring) q /\ poly_vars r q SUBSET U}` [] THEN set_fact `ring_product (poly_ring r (:V)) (S:X->bool) p IN {q | ring_polynomial(r:R ring) q /\ poly_vars r q SUBSET U} ==> poly_vars r (ring_product (poly_ring r (:V)) S p) SUBSET U` THEN GEN_TAC THEN DISCH_TAC THEN have `i:V IN U` [SUBSET] THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_product (poly_ring(r:R ring) (:V)) (S:X->bool) p`; `c:V->R` ]POLY_EVALUATE_MUL THEN qed[] ] );; let poly_product_restrict_subset = prove(` !(r:R ring) S:X->bool U p. S SUBSET U ==> poly_product r U (\s. if s IN S then p s else poly_1 r) = poly_product r S p `, rw[poly_product;x_series_use] THEN intro THEN specialize[ `S:X->bool`; `x_series(r:R ring)`; `U:X->bool`; `p:X->(1->num)->R` ]RING_PRODUCT_RESTRICT_SET THEN set_fact `S SUBSET U:X->bool ==> {x | x IN U /\ S x} = S` THEN rw[IN] THEN qed[] );; let poly_const_product = prove(` !(r:R ring) f:X->R S. FINITE S ==> (!t. t IN S ==> f t IN ring_carrier r) ==> poly_const r (ring_product r S f) = poly_product r S (\t. poly_const r (f t)) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;poly_product_empty] THEN rw[POLY_CONST_1] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN have `ring_powerseries r (poly_const r (f(x:X)):(1->num)->R)` [RING_POWERSERIES_CONST] THEN simp[RING_PRODUCT_CLAUSES;poly_product_insert] THEN qed[POLY_CONST_MUL;RING_PRODUCT] ] );; (* ===== monic_vanishing_at: prod_i (x-c_i) *) let monic_vanishing_at = new_definition ` monic_vanishing_at (r:R ring) (S:X->bool) (c:X->R) = poly_product r S (\s. x_minus_const r (c s)) `;; let monic_vanishing_at_empty = prove(` !(r:R ring) c:X->R. monic_vanishing_at r {} c = poly_1 r `, qed[monic_vanishing_at;poly_product_empty] );; let monic_vanishing_at_insert = prove(` !(r:R ring) S:X->bool c t. (c t) IN ring_carrier r ==> FINITE S ==> monic_vanishing_at r (t INSERT S) c = if t IN S then monic_vanishing_at r S c else poly_mul r (x_minus_const r (c t)) (monic_vanishing_at r S c) `, rw[monic_vanishing_at] THEN intro THEN have `ring_powerseries(r:R ring) (x_minus_const r (c(t:X)))` [x_minus_const_series] THEN simp[poly_product_insert] );; let monic_vanishing_at_plus1 = prove(` !(r:R ring) c n. (c n) IN ring_carrier r ==> monic_vanishing_at r {i:num | i < n+1} c = poly_mul r (x_minus_const r (c n)) (monic_vanishing_at r {i:num | i < n} c) `, intro THEN have `FINITE {i:num | i < n}` [FINITE_NUMSEG_LT] THEN subgoal `{i:num | i < n+1} = n INSERT {i:num | i < n}` THENL [ rw[EXTENSION;IN_INSERT;IN_ELIM_THM] THEN qed[ARITH_RULE `i < n+1 <=> i = n \/ i < n`] ; num_linear_fact `~(n < n:num)` THEN set_fact `~(n < n) ==> ~(n IN {i:num | i < n})` THEN specialize[`r:R ring`;`{i:num | i < n}`;`c:num->R`;`n:num`]monic_vanishing_at_insert THEN qed[] ] );; let monic_vanishing_at_eq = prove(` !(r:R ring) S:X->bool c d. (!s:X. s IN S ==> c s = d s) ==> monic_vanishing_at r S c = monic_vanishing_at r S d `, rw[monic_vanishing_at] THEN qed[poly_product_eq] );; let monic_vanishing_at_poly = prove(` !(r:R ring) S c:X->R. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_polynomial r (monic_vanishing_at r S c) `, intro THEN have `(!s. s IN S ==> ring_polynomial(r:R ring) (x_minus_const r (c(s:X))))` [x_minus_const_poly] THEN rw[monic_vanishing_at] THEN qed[poly_product_poly] );; let monic_vanishing_at_series = prove(` !(r:R ring) S c:X->R. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_powerseries r (monic_vanishing_at r S c) `, qed[monic_vanishing_at_poly;ring_polynomial] );; let x_derivative_monic_vanishing_at = prove(` !(r:R ring) S c:X->R. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> x_derivative r (monic_vanishing_at r S c) = poly_sum r S (\s. (monic_vanishing_at r (S DELETE s) c)) `, intro THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (x_minus_const r (c s))` [x_minus_const_series] THEN specialize[`r:R ring`;`\s:X. x_minus_const(r:R ring) (c s)`;`S:X->bool`]x_derivative_product THEN rw[monic_vanishing_at] THEN simp[] THEN sufficesby poly_sum_eq THEN intro THEN simp[x_derivative_x_minus_const] THEN have `FINITE(S DELETE s:X)` [FINITE_DELETE] THEN set_fact `!t:X. t IN S DELETE s ==> t IN S` THEN have `!t:X. t IN S DELETE s ==> ring_powerseries(r:R ring) (x_minus_const r (c t))` [] THEN specialize[`r:R ring`;`\s:X. x_minus_const(r:R ring) (c s)`;`(S:X->bool) DELETE s`]poly_product_series THEN qed[POLY_MUL_LID] );; let eval_monic_vanishing_at = prove(` !(r:R ring) S (c:X->R) z. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> z IN ring_carrier r ==> poly_eval r (monic_vanishing_at r S c) z = ring_product r S (\s. ring_sub r z (c s)) `, intro THEN rw[monic_vanishing_at] THEN have `!s:X. s IN S ==> ring_polynomial(r:R ring) (x_minus_const r (c s))` [x_minus_const_poly] THEN simp[eval_poly_product] THEN sufficesby RING_PRODUCT_EQ THEN qed[eval_x_minus_const] );; (* compare INTEGRAL_DOMAIN_PRODUCT_EQ_0 *) let ring_product_eq_0 = prove(` !(r:R ring) S (f:X->R). FINITE S /\ (!s. s IN S ==> f s IN ring_carrier r) /\ (?t. t IN S /\ f t = ring_0 r) ==> ring_product r S f = ring_0 r `, intro THEN set_fact `t:X IN S ==> S = t INSERT (S DELETE t)` THEN set_fact `~(t:X IN (S DELETE t))` THEN have `FINITE (S DELETE (t:X))` [FINITE_DELETE] THEN have `ring_product(r:R ring) (t INSERT (S DELETE (t:X))) f = ring_mul r (f t) (ring_product r (S DELETE t) f)` [RING_PRODUCT_CLAUSES] THEN have `ring_product(r:R ring) S f = ring_mul r (f t) (ring_product r (S DELETE(t:X)) f)` [] THEN qed[RING_MUL_LZERO;RING_PRODUCT] );; let eval_monic_vanishing_at_refl = prove(` !(r:R ring) S (c:X->R) t. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> t IN S ==> poly_eval r (monic_vanishing_at r S c) (c t) = ring_0 r `, intro THEN have `c(t:X) IN ring_carrier(r:R ring)` [] THEN simp[eval_monic_vanishing_at] THEN sufficesby ring_product_eq_0 THEN qed[RING_SUB_REFL;RING_SUB] );; let monic_vanishing_at_image = prove(` !(r:R ring) S (f:X->Y) (g:Y->R). (!x y. x IN S ==> y IN S ==> f x = f y ==> x = y) ==> monic_vanishing_at r (IMAGE f S) g = monic_vanishing_at r S (g o f) `, rw[monic_vanishing_at] THEN intro THEN specialize[`r:R ring`;`S:X->bool`;`f:X->Y`;`\s:Y. x_minus_const(r:R ring) (g s)`]poly_product_image THEN simp[] THEN sufficesby poly_product_eq THEN qed[FUN_EQ_THM;o_THM] );; (* ===== monic polynomials *) let monic = new_definition ` monic (r:R ring) (p:(1->num)->R) <=> coeff (poly_deg r p) p = ring_1 r `;; let monic_zero_ring = prove(` !(r:R ring) p. ring_1 r = ring_0 r ==> ring_powerseries r p ==> monic r p `, intro THEN have `trivial_ring(r:R ring)` [TRIVIAL_RING_10] THEN have `coeff (poly_deg r p) p IN ring_carrier(r:R ring)` [coeff_series_in_ring] THEN have `coeff (poly_deg r p) p IN {ring_0(r:R ring)}` [trivial_ring] THEN set_fact `!x. x IN {ring_0(r:R ring)} ==> x = ring_0 r` THEN qed[monic] );; let monic_poly_0 = prove(` !(r:R ring). monic r (poly_0 r) <=> ring_1 r = ring_0 r `, rw[monic;POLY_DEG_0;coeff_poly_0] THEN qed[] );; let monic_poly_1 = prove(` !(r:R ring). monic r (poly_1 r) `, rw[monic;POLY_DEG_1;coeff_poly_1] );; let poly_1_if_monic_deg_0 = prove(` !(r:R ring) p. ring_polynomial r p ==> poly_deg r p = 0 ==> monic r p ==> p = poly_1 r `, intro THEN choose `c:R` `c IN ring_carrier r /\ p = poly_const r c:(1->num)->R` [POLY_DEG_EQ_0] THEN have `coeff 0 p = c:R` [coeff_poly_const] THEN have `coeff 0 p = ring_1(r:R ring)` [monic] THEN qed[poly_1] );; let monic_x_pow = prove(` !(r:R ring) n. monic r (x_pow r n) `, intro THEN case `ring_1 r = ring_0(r:R ring)` THENL [ qed[x_pow_series;monic_zero_ring] ; have `poly_deg(r:R ring) (x_pow r n) = n` [deg_x_pow] THEN simp[monic;coeff_x_pow] ] );; let monic_x_minus_const = prove(` !(r:R ring) c. c IN ring_carrier r ==> monic r (x_minus_const r c) `, intro THEN case `ring_1 r = ring_0(r:R ring)` THENL [ qed[x_minus_const_series;monic_zero_ring] ; have `poly_deg(r:R ring) (x_minus_const r c) = 1` [deg_x_minus_const] THEN simp[monic;coeff_x_minus_const] ] );; let topcoeff_monic_poly_mul = prove(` !(r:R ring) p q. ring_polynomial r p ==> ring_polynomial r q ==> monic r p ==> monic r q ==> coeff (poly_deg r p + poly_deg r q) (poly_mul r p q) = ring_1 r `, rw[monic] THEN intro THEN rw[coeff_poly_mul_oneindex] THEN subgoal `ring_sum(r:R ring) (0..poly_deg r p + poly_deg r q) (\a. ring_mul r (coeff a p) (coeff ((poly_deg r p + poly_deg r q) - a) q)) = ring_sum r (0..poly_deg r p + poly_deg r q) (\a. if a = poly_deg r p then ring_mul r (coeff a p) (coeff ((poly_deg r p + poly_deg r q) - a) q) else ring_0 r)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN simp[] THEN case `a = poly_deg r (p:(1->num)->R)` THENL [ simp[] ; pass ] THEN simp[] THEN case `a < poly_deg r (p:(1->num)->R)` THENL [ num_linear_fact `a < poly_deg r (p:(1->num)->R) ==> ~((poly_deg r p + poly_deg r (q:(1->num)->R)) - a <= poly_deg r q)` THEN num_linear_fact `poly_deg r (q:(1->num)->R) <= poly_deg r q` THEN have `coeff ((poly_deg r (p:(1->num)->R) + poly_deg r (q:(1->num)->R)) - a) q = ring_0 r` [coeff_deg_le] THEN qed[RING_MUL_RZERO;coeff_poly_in_ring] ; num_linear_fact `~(a = poly_deg r (p:(1->num)->R)) /\ ~(a < poly_deg r p) ==> ~(a <= poly_deg r p)` THEN num_linear_fact `poly_deg r (p:(1->num)->R) <= poly_deg r p` THEN have `coeff a (p:(1->num)->R) = ring_0 r` [coeff_deg_le] THEN qed[RING_MUL_LZERO;coeff_poly_in_ring] ] ; pass ] THEN have `poly_deg r (p:(1->num)->R) IN 0..poly_deg r p + poly_deg r (q:(1->num)->R)` [IN_NUMSEG_0;ARITH_RULE `poly_deg r (p:(1->num)->R) <= poly_deg r p + poly_deg r (q:(1->num)->R)`] THEN simp[RING_SUM_DELTA;RING_MUL_LID;RING_1;RING_MUL;coeff_poly_in_ring;ARITH_RULE `(poly_deg r (p:(1->num)->R) + poly_deg r (q:(1->num)->R)) - poly_deg r p = poly_deg r q`] );; let deg_monic_poly_mul = prove(` !(r:R ring) p q. ring_polynomial r p ==> ring_polynomial r q ==> monic r p ==> monic r q ==> poly_deg r (poly_mul r p q) = poly_deg r p + poly_deg r q `, intro THEN case `ring_1(r:R ring) = ring_0 r` THENL [ have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (q:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (poly_mul r (p:(1->num)->R) (q:(1->num)->R))` [RING_POWERSERIES_MUL] THEN simp[deg_zero_ring] THEN qed[ARITH_RULE `0+0 = 0`] ; pass ] THEN have `ring_polynomial r (poly_mul r (p:(1->num)->R) (q:(1->num)->R))` [RING_POLYNOMIAL_MUL] THEN have `poly_deg r (poly_mul r p q) <= poly_deg r (p:(1->num)->R) + poly_deg r (q:(1->num)->R)` [POLY_DEG_MUL_LE] THEN have `coeff (poly_deg r (p:(1->num)->R) + poly_deg r (q:(1->num)->R)) (poly_mul r p q) = ring_1 r` [topcoeff_monic_poly_mul] THEN have `~(coeff (poly_deg r (p:(1->num)->R) + poly_deg r (q:(1->num)->R)) (poly_mul r p q) = ring_0 r)` [] THEN qed[deg_coeff_from_le] );; let monic_poly_mul = prove(` !(r:R ring) p q. ring_polynomial r p ==> ring_polynomial r q ==> monic r p ==> monic r q ==> monic r (poly_mul r p q) `, intro THEN have `poly_deg r (poly_mul r (p:(1->num)->R) (q:(1->num)->R)) = poly_deg r p + poly_deg r q` [deg_monic_poly_mul] THEN simp[monic] THEN qed[topcoeff_monic_poly_mul] );; let monic_poly_product = prove(` !(r:R ring) p (S:X->bool). FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> (!s. s IN S ==> monic r (p s)) ==> monic r (poly_product r S p) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty] THEN qed[monic_poly_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN S ==> ring_polynomial r (p(s:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_polynomial] THEN have `ring_polynomial r (p(x:X):(1->num)->R)` [] THEN have `ring_polynomial (r:R ring) (poly_product r (S:X->bool) p)` [poly_product_poly] THEN have `monic (r:R ring) (poly_product r (S:X->bool) p)` [] THEN have `monic r (p(x:X):(1->num)->R)` [] THEN simp[poly_product_insert] THEN qed[monic_poly_mul] ] );; let monic_vanishing_at_monic = prove(` !(r:R ring) S:X->bool c. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> monic r (monic_vanishing_at r S c) `, rw[monic_vanishing_at] THEN intro THEN have `!s:X. s IN S ==> ring_polynomial(r:R ring) (x_minus_const r (c s))` [x_minus_const_poly] THEN have `!s:X. s IN S ==> monic(r:R ring) (x_minus_const r (c s))` [monic_x_minus_const] THEN simp[monic_poly_product] );; let monic_subring = prove(` !(r:R ring) G p. monic (subring_generated r G) p <=> monic r p `, rw[monic;SUBRING_GENERATED;POLY_DEG_SUBRING_GENERATED] );; (* ===== r[x] when r is a field *) let PID_x_poly_field = prove(` !r:R ring. field r ==> PID(x_poly r) `, qed[PID_POLY_RING;x_poly] );; let UFD_x_poly_field = prove(` !r:R ring. field r ==> UFD(x_poly r) `, qed[PID_x_poly_field;PID_IMP_UFD] );; let integral_domain_x_poly_field = prove(` !r:R ring. field r ==> integral_domain(x_poly r) `, qed[PID_x_poly_field;PID_IMP_INTEGRAL_DOMAIN] );; let prime_iff_irreducible_over_field = prove(` !(r:R ring) p. field r ==> (ring_irreducible(x_poly r) p <=> ring_prime(x_poly r) p) `, qed[PID_x_poly_field;PID_IRREDUCIBLE_EQ_PRIME] );; let squarefree_if_irreducible_over_field = prove(` !(r:R ring) p. field r ==> ring_irreducible(x_poly r) p ==> ring_squarefree(x_poly r) p `, qed[prime_iff_irreducible_over_field;ring_squarefree_if_prime] );; let x_poly_field_monic_associate = prove(` !(r:R ring) (p:(1->num)->R). field r ==> ring_polynomial r p ==> ~(p = poly_0 r) ==> ?q. (ring_polynomial r q /\ monic r q /\ ring_associates(x_poly r) p q) `, intro THEN def `n:num` `poly_deg r (p:(1->num)->R)` THEN def `pn:R` `coeff n (p:(1->num)->R)` THEN have `pn IN ring_carrier(r:R ring)` [coeff_poly_in_ring] THEN have `~(pn = ring_0(r:R ring))` [topcoeff_nonzero] THEN have `ring_unit(r:R ring) pn` [FIELD_UNIT] THEN have `ring_inv(r:R ring) pn IN ring_carrier r` [RING_INV] THEN have `ring_mul(r:R ring) pn (ring_inv r pn) = ring_1 r` [ring_div_refl;ring_div] THEN have `ring_mul(r:R ring) (ring_inv r pn) pn = ring_1 r` [RING_MUL_SYM] THEN def `q:(1->num)->R` `poly_mul r (p:(1->num)->R) (poly_const r (ring_inv r pn))` THEN have `poly_deg r (q:(1->num)->R) = poly_deg r (p:(1->num)->R)` [deg_mul_const_const_1] THEN witness `q:(1->num)->R` THEN rw[monic] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_polynomial r (q:(1->num)->R)` [RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_CONST] THEN have `ring_powerseries r (q:(1->num)->R)` [ring_polynomial] THEN specialize[`r:R ring`;`ring_inv(r:R ring) pn`;`p:(1->num)->R`;`poly_deg r (p:(1->num)->R)`]coeff_times_poly_const THEN have `coeff (poly_deg r (p:(1->num)->R)) (q:(1->num)->R) = ring_1 r` [coeff_times_poly_const] THEN have `coeff (poly_deg r (q:(1->num)->R)) (q:(1->num)->R) = ring_1 r` [] THEN rw[x_poly] THEN have `ring_unit(r:R ring) (ring_inv r pn)` [RING_UNIT_INV] THEN qed[associates_if_mul_unit_const;x_poly] );; let mul_unit_const_if_associates = prove(` !(r:R ring) (p:(V->num)->R) q. field r ==> ring_associates(poly_ring r (:V)) p q ==> ?c. ring_unit r c /\ q = poly_mul r p (poly_const r c) `, intro THEN have `integral_domain(r:R ring)` [FIELD_IMP_INTEGRAL_DOMAIN] THEN have `integral_domain(poly_ring (r:R ring) (:V))` [INTEGRAL_DOMAIN_POLY_RING] THEN choose `u:(V->num)->R` `ring_unit(poly_ring (r:R ring) (:V)) u /\ ring_mul(poly_ring (r:R ring) (:V)) p u = q` [INTEGRAL_DOMAIN_ASSOCIATES] THEN choose `c:R` `ring_unit(r:R ring) c /\ u = poly_const r c:(V->num)->R` [RING_UNIT_POLY_DOMAIN] THEN witness `c:R` THEN qed[POLY_RING_CLAUSES] );; let monic_associates = prove(` !(r:R ring) p q. field r ==> monic r p ==> monic r q ==> ring_associates(x_poly r) p q ==> p = q `, intro THEN choose `c:R` `ring_unit(r:R ring) c /\ q = poly_mul r p (poly_const r c:(1->num)->R)` [mul_unit_const_if_associates;x_poly] THEN have `integral_domain(x_poly(r:R ring))` [integral_domain_x_poly_field] THEN have `p IN ring_carrier(x_poly(r:R ring))` [INTEGRAL_DOMAIN_ASSOCIATES] THEN have `q IN ring_carrier(x_poly(r:R ring))` [INTEGRAL_DOMAIN_ASSOCIATES] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN have `ring_polynomial r (q:(1->num)->R)` [x_poly_use] THEN have `ring_unit(r:R ring) c` [] THEN specialize[`r:R ring`;`p:(1->num)->R`;`c:R`]deg_mul_unit_const THEN have `poly_deg r (q:(1->num)->R) = poly_deg r (p:(1->num)->R)` [deg_mul_unit_const;x_poly_use] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `c IN ring_carrier(r:R ring)` [ring_unit] THEN specialize[`r:R ring`;`c:R`;`p:(1->num)->R`;`poly_deg r (p:(1->num)->R)`]coeff_times_poly_const THEN have `coeff (poly_deg r (p:(1->num)->R)) p = ring_1 r` [monic] THEN have `coeff (poly_deg r (p:(1->num)->R)) q = ring_1 r` [monic] THEN have `ring_1 r = ring_mul r c (ring_1(r:R ring))` [] THEN have `c = ring_1(r:R ring)` [RING_MUL_RID] THEN have `poly_const r c = ring_1 (x_poly(r:R ring))` [poly_1;x_poly_use] THEN qed[RING_MUL_RID;x_poly_use] );; let no_square_divisor_if_coprime_derivative_lemma1 = prove(` !(r:R ring) q (u:(V->num)->R). ring_polynomial r q ==> ring_polynomial r u ==> poly_mul r (poly_mul r q q) u = poly_mul r q (poly_mul r q u) `, intro THEN have `ring_powerseries(r:R ring) (q:(V->num)->R)` [ring_polynomial] THEN have `ring_powerseries(r:R ring) (u:(V->num)->R)` [ring_polynomial] THEN qed[POLY_MUL_ASSOC] );; let no_square_divisor_if_coprime_derivative_lemma2 = prove(` !(r:R ring) q u. ring_polynomial r q ==> ring_polynomial r u ==> poly_add r (poly_mul r (poly_add r (poly_mul r (x_derivative r q) q) (poly_mul r q (x_derivative r q))) u) (poly_mul r (poly_mul r q q) (x_derivative r u)) = poly_mul r q (poly_add r (poly_mul r (poly_add r (x_derivative r q) (x_derivative r q)) u) (poly_mul r q (x_derivative r u))) `, intro THEN have `ring_polynomial(r:R ring) (x_derivative r q)` [x_derivative_polynomial] THEN have `ring_polynomial(r:R ring) (x_derivative r u)` [x_derivative_polynomial] THEN have `q IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `u IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `x_derivative r q IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `x_derivative r u IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN rw[x_poly_use] THEN specialize[ `x_poly(r:R ring)`;`q:(1->num)->R` ;`x_derivative r q:(1->num)->R` ;`u:(1->num)->R` ;`x_derivative r u:(1->num)->R`] (GENL[`r:R ring`;`q:R`;`Q:R`;`u:R`;`U:R`]( RING_RULE `ring_add(r:R ring) (ring_mul r (ring_add r (ring_mul r Q q) (ring_mul r q Q)) u) (ring_mul r (ring_mul r q q) U) = ring_mul r q (ring_add r (ring_mul r (ring_add r Q Q) u) (ring_mul r q U))` )) THEN qed[] );; let no_square_divisor_if_coprime_derivative = prove(` !(r:R ring) p q. field r ==> ring_polynomial r q ==> ring_coprime(x_poly r) (p,x_derivative r p) ==> ring_divides(x_poly r) (poly_mul r q q) p ==> ring_unit(x_poly r) q `, intro THEN choose `u:(1->num)->R` `u IN ring_carrier(x_poly r) /\ p:(1->num)->R = ring_mul(x_poly r) (poly_mul r q q) u` [ring_divides] THEN have `p:(1->num)->R = poly_mul r (poly_mul r q q) u` [x_poly_use] THEN have `(p:(1->num)->R) IN ring_carrier(x_poly r)` [ring_coprime] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (q:(1->num)->R)` [ring_polynomial] THEN specialize[`r:R ring`;`q:(1->num)->R`;`q:(1->num)->R`]x_derivative_mul THEN have `ring_polynomial r (u:(1->num)->R)` [x_poly_use] THEN have `ring_powerseries r (u:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (poly_mul r q (q:(1->num)->R))` [RING_POWERSERIES_MUL] THEN specialize[`r:R ring`;`poly_mul r q (q:(1->num)->R)`;`u:(1->num)->R`]x_derivative_mul THEN have `x_derivative(r:R ring) p = poly_add r (poly_mul r (poly_add r (poly_mul r (x_derivative r q) q) (poly_mul r q (x_derivative r q))) u) (poly_mul r (poly_mul r q q) (x_derivative r u))` [] THEN have `x_derivative(r:R ring) p = poly_mul r q (poly_add r (poly_mul r (poly_add r (x_derivative r q) (x_derivative r q)) u) (poly_mul r q (x_derivative r u)))` [no_square_divisor_if_coprime_derivative_lemma2] THEN have `ring_polynomial r (x_derivative r (p:(1->num)->R))` [x_derivative_polynomial] THEN have `ring_polynomial r (x_derivative r (q:(1->num)->R))` [x_derivative_polynomial] THEN have `ring_polynomial r (x_derivative r (u:(1->num)->R))` [x_derivative_polynomial] THEN have `ring_polynomial r (poly_add r (x_derivative r (q:(1->num)->R)) (x_derivative r q))` [RING_POLYNOMIAL_ADD] THEN have `ring_polynomial r (poly_mul r (poly_add r (x_derivative r (q:(1->num)->R)) (x_derivative r q)) u)` [RING_POLYNOMIAL_MUL] THEN have `ring_polynomial r (poly_mul r (q:(1->num)->R) (x_derivative r u))` [RING_POLYNOMIAL_MUL] THEN have `ring_polynomial r (poly_add r (poly_mul r (poly_add r (x_derivative r (q:(1->num)->R)) (x_derivative r q)) u) (poly_mul r q (x_derivative r u)))` [RING_POLYNOMIAL_ADD] THEN have `poly_add r (poly_mul r (poly_add r (x_derivative r (q:(1->num)->R)) (x_derivative r q)) u) (poly_mul r q (x_derivative r u)) IN ring_carrier(x_poly r)` [x_poly_use] THEN have `(q:(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN have `(p:(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN have `x_derivative r (p:(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN subgoal `ring_divides(x_poly(r:R ring)) q (x_derivative r p)` THENL [ rw[ring_divides] THEN intro THENL [qed[]; qed[]; pass] THEN witness `poly_add r (poly_mul r (poly_add r (x_derivative r (q:(1->num)->R)) (x_derivative r q)) u) (poly_mul r q (x_derivative r u))` THEN intro THENL [ qed[] ; rw[GSYM x_poly_use] THEN qed[] ] ; pass ] THEN subgoal `ring_divides(x_poly(r:R ring)) q p` THENL [ rw[ring_divides] THEN intro THENL [qed[]; qed[]; pass] THEN witness `poly_mul r q u:(1->num)->R` THEN have `ring_polynomial r (poly_mul r q u:(1->num)->R)` [RING_POLYNOMIAL_MUL] THEN have `(poly_mul r q u:(1->num)->R) IN ring_carrier(x_poly r)` [x_poly_use] THEN have `p:(1->num)->R = poly_mul r q (poly_mul r q u)` [no_square_divisor_if_coprime_derivative_lemma1] THEN qed[x_poly_use] ; pass ] THEN qed[ring_coprime] );; let nonzero_if_coprime_derivative = prove(` !(r:R ring) p. field r ==> ring_coprime(x_poly r) (p,x_derivative r p) ==> ~(p = poly_0 r) `, intro THEN have `PID (x_poly(r:R ring))` [PID_x_poly_field] THEN have `ring_0(x_poly r):(1->num)->R = poly_0 r` [x_poly_use] THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_coprime] THEN have `x_derivative r p = ring_0(x_poly r):(1->num)->R` [x_derivative_poly_0] THEN have `ring_gcd(x_poly r) (p,x_derivative r p) = ring_0(x_poly r):(1->num)->R` [RING_GCD_00] THEN have `x_derivative r p IN ring_carrier(x_poly(r:R ring))` [ring_coprime] THEN specialize[`x_poly(r:R ring)`;`p:(1->num)->R`;`x_derivative (r:R ring) p`]RING_GCD_EQ_1 THEN have `ring_gcd(x_poly r) (p,x_derivative r p) = ring_1(x_poly r):(1->num)->R` [RING_GCD_EQ_1] THEN qed[PID_IMP_INTEGRAL_DOMAIN;integral_domain] );; let squarefree_if_coprime_derivative = prove(` !(r:R ring) p. field r ==> ring_coprime(x_poly r) (p,x_derivative r p) ==> ring_squarefree(x_poly r) p `, intro THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_coprime] THEN have `~(p = ring_0(x_poly(r:R ring)))` [nonzero_if_coprime_derivative;x_poly_use] THEN proven_if `ring_unit(x_poly(r:R ring)) p` [ring_squarefree_if_unit] THEN have `PID (x_poly(r:R ring))` [PID_x_poly_field] THEN have `UFD (x_poly(r:R ring))` [PID_IMP_UFD] THEN ASSUME_TAC(UNDISCH_ALL (ISPECL[`p:(1->num)->R`] (CONJUNCT2 (UNDISCH_ALL (fst (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL [`x_poly(r:R ring)`]UFD_EQ_PRIMEFACT_NONUNIT))))))))) THEN choose2 `n:num` `q:num->(1->num)->R` `1 <= n /\ (!i. 1 <= i /\ i <= n ==> ring_prime(x_poly(r:R ring)) (q i)) /\ ring_product(x_poly r) (1..n) q = p` [] THEN subgoal `!i j. 1 <= i /\ i <= n /\ 1 <= j /\ j <= n /\ ring_divides(x_poly(r:R ring)) (q i) (q j) /\ ~(i = j) ==> F` THENL [ intro THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!i:num. i IN (1..n) ==> q i IN ring_carrier(x_poly(r:R ring))` [ring_prime;IN_NUMSEG] THEN have `i IN (1..n)` [IN_NUMSEG] THEN have `j IN (1..n)` [IN_NUMSEG] THEN specialize[`x_poly(r:R ring)`;`1..n`;`q:num->(1->num)->R`;`i:num`;`j:num`]square_divides_product_if_factor_divides_factor THEN have `ring_divides(x_poly(r:R ring)) (poly_mul r (q(i:num)) (q(i))) p` [x_poly_use] THEN have `ring_polynomial r (q(i:num):(1->num)->R)` [ring_prime;x_poly_use] THEN specialize[`r:R ring`;`p:(1->num)->R`;`q(i:num):(1->num)->R`]no_square_divisor_if_coprime_derivative THEN qed[ring_prime] ; pass ] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!i. i IN 1..n ==> ring_prime(x_poly(r:R ring)) (q i)` [IN_NUMSEG] THEN have `!i. i IN 1..n ==> q i IN ring_carrier(x_poly(r:R ring))` [ring_prime] THEN have `!i j. i IN 1..n ==> j IN 1..n ==> ring_divides(x_poly(r:R ring)) (q i) (q j) ==> i = j` [IN_NUMSEG] THEN specialize[`x_poly(r:R ring)`;`1..n`;`q:num->(1->num)->R`]ring_squarefree_if_product_coprime_primes_indexed THEN qed[] );; let deg_x_derivative_lemma = prove(` !(r:R ring) p d. ring_polynomial r p ==> ~(coeff d (x_derivative r p) = ring_0(r:R ring)) ==> d <= poly_deg r p - 1 `, rw[coeff_x_derivative] THEN intro THEN have `~(coeff (d+1) p = ring_0(r:R ring))` [RING_MUL_RZERO;RING_OF_NUM] THEN have `d+1 <= poly_deg r (p:(1->num)->R)` [coeff_le_deg] THEN num_linear_fact `d+1 <= poly_deg r (p:(1->num)->R) ==> d <= poly_deg r p - 1` THEN qed[] );; let deg_x_derivative_le = prove(` !(r:R ring) p. ring_polynomial r p ==> poly_deg r (x_derivative r p) <= poly_deg r p - 1 `, intro THEN have `!d. ~(coeff d (x_derivative r p) = ring_0(r:R ring)) ==> d <= poly_deg r p - 1` [deg_x_derivative_lemma] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_polynomial r (x_derivative r (p:(1->num)->R))` [x_derivative_polynomial] THEN qed[deg_le_coeff;ring_polynomial] );; (* warning: 0 - 1 = 0 *) let deg_x_derivative = prove(` !(r:R ring) p. integral_domain r ==> ring_char r = 0 ==> ring_polynomial r p ==> poly_deg r (x_derivative r p) = poly_deg r p - 1 `, intro THEN have `!d. ~(coeff d (x_derivative r p) = ring_0(r:R ring)) ==> d <= poly_deg r p - 1` [deg_x_derivative_lemma] THEN case `poly_deg r (p:(1->num)->R) = 0` THENL [ choose `c:R` `c IN ring_carrier r /\ (p:(1->num)->R) = poly_const r c` [POLY_DEG_EQ_0] THEN have `x_derivative r (p:(1->num)->R) = poly_0 r` [x_derivative_poly_const] THEN qed[POLY_DEG_0;ARITH_RULE `0 - 1 = 0`] ; pass ] THEN subgoal `~(coeff (poly_deg r p - 1) (x_derivative r p) = ring_0(r:R ring))` THENL [ have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN rw[coeff_x_derivative] THEN num_linear_fact `~(poly_deg r (p:(1->num)->R) - 1 + 1 = 0)` THEN have `~(ring_of_num r (poly_deg r (p:(1->num)->R) - 1 + 1) = ring_0 r)` [RING_CHAR_EQ_0] THEN num_linear_fact `~(poly_deg r (p:(1->num)->R) = 0) ==> poly_deg r p - 1 + 1 = poly_deg r p` THEN have `~((p:(1->num)->R) = poly_0 r)` [POLY_DEG_0;x_poly_use] THEN have `~(coeff (poly_deg r p - 1 + 1) (p:(1->num)->R) = ring_0 r)` [topcoeff_nonzero] THEN qed[integral_domain;RING_OF_NUM;coeff_poly_in_ring] ; pass ] THEN have `ring_polynomial r (x_derivative r (p:(1->num)->R))` [x_derivative_polynomial] THEN qed[deg_coeff;ring_polynomial] );; let x_derivative_nonzero = prove(` !(r:R ring) p. integral_domain r ==> ring_char r = 0 ==> ring_polynomial r p ==> ~(poly_deg r p = 0) ==> ~(x_derivative r p = poly_0 r) `, intro THEN num_linear_fact `~(poly_deg r (p:(1->num)->R) = 0) ==> poly_deg r p - 1 + 1 = poly_deg r p` THEN subgoal `~(coeff (poly_deg r p - 1) (x_derivative r p) = ring_0(r:R ring))` THENL [ have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN rw[coeff_x_derivative] THEN num_linear_fact `~(poly_deg r (p:(1->num)->R) - 1 + 1 = 0)` THEN have `~(ring_of_num r (poly_deg r (p:(1->num)->R) - 1 + 1) = ring_0 r)` [RING_CHAR_EQ_0] THEN have `~((p:(1->num)->R) = poly_0 r)` [POLY_DEG_0;x_poly_use] THEN have `~(coeff (poly_deg r p - 1 + 1) (p:(1->num)->R) = ring_0 r)` [topcoeff_nonzero] THEN qed[integral_domain;RING_OF_NUM;coeff_poly_in_ring] ; pass ] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `coeff (poly_deg r (p:(1->num)->R) - 1) (x_derivative r p) = ring_mul r (ring_of_num r (poly_deg r p - 1 + 1)) (coeff (poly_deg r p - 1 + 1) p)` [coeff_x_derivative] THEN have `coeff (poly_deg r (p:(1->num)->R) - 1) (x_derivative r p) = ring_mul r (ring_of_num r (poly_deg r p)) (coeff (poly_deg r p) p)` [coeff_x_derivative] THEN have `~((p:(1->num)->R) = poly_0 r)` [POLY_DEG_0;x_poly_use] THEN have `~(coeff (poly_deg r p) (p:(1->num)->R) = ring_0 r)` [topcoeff_nonzero] THEN have `~(ring_of_num r (poly_deg r (p:(1->num)->R)) = ring_0 r)` [RING_CHAR_EQ_0] THEN have `~(coeff (poly_deg r (p:(1->num)->R) - 1) (x_derivative r p) = ring_0 r)` [integral_domain;RING_OF_NUM;coeff_poly_in_ring] THEN qed[coeff_poly_0] );; let deg_divides = prove(` !(r:R ring) (S:V->bool) p q. integral_domain r ==> ring_divides(poly_ring r S) p q ==> ~(q = poly_0 r) ==> poly_deg r p <= poly_deg r q `, intro THEN have `(p:(V->num)->R) IN ring_carrier(poly_ring r S)` [ring_divides] THEN have `(p:(V->num)->R) IN {f | ring_polynomial r f /\ poly_vars r f SUBSET S}` [POLY_RING] THEN set_fact `(p:(V->num)->R) IN {f | ring_polynomial r f /\ poly_vars r f SUBSET S} ==> ring_polynomial r p` THEN choose `u:(V->num)->R` `(u:(V->num)->R) IN ring_carrier(poly_ring r S) /\ q = ring_mul(poly_ring r S) p u` [ring_divides] THEN have `(u:(V->num)->R) IN {f | ring_polynomial r f /\ poly_vars r f SUBSET S}` [POLY_RING] THEN set_fact `(u:(V->num)->R) IN {f | ring_polynomial r f /\ poly_vars r f SUBSET S} ==> ring_polynomial r u` THEN have `~((p:(V->num)->R) = poly_0 r)` [POLY_MUL_0;POLY_RING] THEN have `~((u:(V->num)->R) = poly_0 r)` [POLY_MUL_0;POLY_RING] THEN have `poly_deg r (poly_mul r p (u:(V->num)->R)) = poly_deg r p + poly_deg r u` [POLY_DEG_MUL] THEN num_linear_fact `poly_deg r (poly_mul r p (u:(V->num)->R)) = poly_deg r p + poly_deg r u ==> poly_deg r p <= poly_deg r (poly_mul r p u)` THEN qed[POLY_RING] );; let poly_deg_product = prove(` !(r:R ring) p n S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> ~(p s = poly_0 r)) ==> (!s:X. s IN S ==> poly_deg r (p s) = n s) ==> field r ==> poly_deg r (poly_product r S p) = nsum S n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;NSUM_CLAUSES;POLY_DEG_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial r (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries r (p(x:X):(1->num)->R)` [ring_polynomial] THEN simp[poly_product_insert;NSUM_CLAUSES] THEN simp[CARD_CLAUSES] THEN have `ring_polynomial(r:R ring) (poly_product r (S:X->bool) p)` [poly_product_poly] THEN have `poly_deg(r:R ring) (poly_product r (S:X->bool) p) = nsum S n` [poly_product_poly] THEN subgoal `~(poly_product r (S:X->bool) p = poly_0(r:R ring))` THENL [ simp[poly_product_ring_product_x_poly;x_poly_use] THEN have `integral_domain(x_poly(r:R ring))` [integral_domain_x_poly_field] THEN specialize[`x_poly(r:R ring)`;`S:X->bool`;`p:X->(1->num)->R`]INTEGRAL_DOMAIN_PRODUCT_EQ_0 THEN qed[x_poly_use] ; pass ] THEN specialize_assuming[`r:R ring`;`p(x:X):(1->num)->R`;`poly_product r (S:X->bool) p:(1->num)->R`]POLY_DEG_MUL THEN qed[FIELD_IMP_INTEGRAL_DOMAIN] ] );; let deg_prime = prove(` !(r:R ring) p. field r ==> ring_prime(x_poly r) p ==> ~(poly_deg r p = 0) `, intro THEN have `(p:(1->num)->R) IN ring_carrier(x_poly r)` [ring_prime] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN choose `c:R` `c IN ring_carrier r /\ (p:(1->num)->R) = poly_const r c` [POLY_DEG_EQ_0] THEN case `c = ring_0(r:R ring)` THENL [ have `p = ring_0(x_poly(r:R ring))` [x_poly_use;poly_0] THEN qed[ring_prime] ; choose `d:R` `d IN ring_carrier(r:R ring) /\ ring_mul r c d = ring_1 r` [field] THEN have `poly_mul r (poly_const r c) (poly_const r d) = poly_const r (ring_mul r c d):(1->num)->R` [POLY_CONST_MUL] THEN have `poly_mul r (poly_const r c) (poly_const r d) = poly_1 r:(1->num)->R` [poly_1] THEN have `poly_mul r p (poly_const r d) = poly_1 r:(1->num)->R` [] THEN have `ring_mul(x_poly r) p (poly_const r d) = ring_1(x_poly(r:R ring))` [x_poly_use] THEN have `poly_const r d IN ring_carrier(x_poly(r:R ring))` [POLY_CONST;x_poly] THEN have `ring_unit(x_poly(r:R ring)) p` [ring_unit] THEN qed[ring_prime] ] );; let deg_0_if_unit = prove(` !(r:R ring) p. field r ==> ring_unit(x_poly r) p ==> poly_deg r p = 0 `, intro THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_unit] THEN choose `q:(1->num)->R` `q IN ring_carrier(x_poly(r:R ring)) /\ ring_mul(x_poly r) p q = ring_1(x_poly r)` [ring_unit] THEN have `integral_domain(r:R ring)` [FIELD_IMP_INTEGRAL_DOMAIN] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN have `ring_polynomial r (q:(1->num)->R)` [x_poly_use] THEN have `poly_mul r p q = poly_1 r:(1->num)->R` [x_poly_use] THEN have `~(poly_1 r = poly_0 r:(1->num)->R)` [coeff_poly_1;coeff_poly_0;integral_domain] THEN have `~(p = poly_0 r:(1->num)->R)` [POLY_MUL_0] THEN have `~(q = poly_0 r:(1->num)->R)` [POLY_MUL_0] THEN have `poly_deg r (poly_mul r p q:(1->num)->R) = poly_deg r p + poly_deg r q` [POLY_DEG_MUL] THEN have `poly_deg r (poly_1 r:(1->num)->R) = 0` [POLY_DEG_1] THEN qed[ARITH_RULE `0 = d + e ==> d = 0`] );; let prime_if_deg_1 = prove(` !(r:R ring) p. field r ==> ring_polynomial r p ==> poly_deg r p = 1 ==> ring_prime(x_poly r) p `, intro THEN have `~ring_unit(x_poly(r:R ring)) p` [deg_0_if_unit;ARITH_RULE `~(1 = 0)`] THEN have `p IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `~(p = poly_0(r:R ring):(1->num)->R)` [POLY_DEG_0;ARITH_RULE `~(1 = 0)`] THEN have `~(p = ring_0(x_poly(r:R ring)))` [x_poly_use] THEN have `PID (x_poly(r:R ring))` [PID_x_poly_field] THEN have `UFD (x_poly(r:R ring))` [PID_IMP_UFD] THEN ASSUME_TAC(UNDISCH_ALL (ISPECL[`p:(1->num)->R`] (CONJUNCT2 (UNDISCH_ALL (fst (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL [`x_poly(r:R ring)`]UFD_EQ_PRIMEFACT_NONUNIT))))))))) THEN choose2 `n:num` `q:num->((1->num)->R)` `1 <= n /\ (!i. 1 <= i ==> i <= n ==> ring_prime (x_poly(r:R ring)) (q i)) /\ ring_product (x_poly r) (1..n) q = p` [] THEN have `!i. 1 <= i ==> i <= n ==> ring_polynomial r (q i:(1->num)->R)` [ring_prime;x_poly_use] THEN have `!i. 1 <= i ==> i <= n ==> ~(q i = poly_0 r:(1->num)->R)` [deg_prime;POLY_DEG_0] THEN have `!i. i IN 1..n ==> 1 <= i /\ i <= n` [IN_NUMSEG] THEN have `!i. i IN 1..n ==> ring_polynomial r (q i:(1->num)->R)` [ring_prime;x_poly_use] THEN have `!i. i IN 1..n ==> ~(q i = poly_0 r:(1->num)->R)` [deg_prime;POLY_DEG_0] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN specialize[`r:R ring`;`q:num->(1->num)->R`;`\s:num. poly_deg r (q s:(1->num)->R)`;`1..n`]poly_deg_product THEN have `p = poly_product(r:R ring) (1..n) q` [poly_product_ring_product_x_poly] THEN have `nsum(1..n) (\s. poly_deg r (q s:(1->num)->R)) = 1` [] THEN have `!s. s IN 1..n ==> 1 <= poly_deg r (q s:(1->num)->R)` [deg_prime;ARITH_RULE `~(x = 0) ==> 1 <= x`] THEN have `nsum(1..n) (\s. 1) <= nsum(1..n) (\s. poly_deg r (q s:(1->num)->R))` [NSUM_LE] THEN have `CARD(1..n) <= nsum(1..n) (\s. poly_deg r (q s:(1->num)->R))` [CARD_EQ_NSUM] THEN have `n <= nsum(1..n) (\s. poly_deg r (q s:(1->num)->R))` [CARD_NUMSEG_1] THEN have `n <= 1` [] THEN have `1 <= n` [] THEN num_linear_fact `1 <= n ==> n <= 1 ==> n = 1` THEN have `p = ring_product(x_poly(r:R ring)) {1} q` [NUMSEG_SING] THEN have `p = q 1:(1->num)->R` [RING_PRODUCT_SING;x_poly_use] THEN qed[] );; let coprime_prime_derivative = prove(` !(r:R ring) p. field r ==> ring_char r = 0 ==> ring_prime(x_poly r) p ==> ring_coprime(x_poly r) (p,x_derivative r p) `, rw[ring_coprime] THEN intro THENL [ qed[ring_prime] ; qed[ring_prime;x_poly_use;x_derivative_polynomial] ; have `integral_domain(x_poly(r:R ring))` [integral_domain_x_poly_field] THEN have `ring_irreducible(x_poly(r:R ring)) p` [INTEGRAL_DOMAIN_PRIME_IMP_IRREDUCIBLE] THEN proven_if `ring_unit (x_poly(r:R ring)) d` [] THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_prime] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN have `ring_polynomial r (x_derivative r (p:(1->num)->R))` [x_derivative_polynomial] THEN have `x_derivative r p IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `ring_associates(x_poly(r:R ring)) d p` [RING_NONUNIT_DIVIDES_IRREDUCIBLE] THEN have `ring_divides(x_poly(r:R ring)) p (x_derivative r p)` [RING_ASSOCIATES_REFL;RING_ASSOCIATES_DIVIDES] THEN have `integral_domain(r:R ring)` [FIELD_IMP_INTEGRAL_DOMAIN] THEN have `~(poly_deg r (p:(1->num)->R) = 0)` [deg_prime] THEN have `~(x_derivative r p = poly_0(r:R ring))` [x_derivative_nonzero] THEN have `poly_deg r (p:(1->num)->R) <= poly_deg r (x_derivative r p)` [deg_divides;x_poly] THEN have `poly_deg r (x_derivative r p) = poly_deg r (p:(1->num)->R) - 1` [deg_x_derivative] THEN num_linear_fact `poly_deg r (p:(1->num)->R) <= poly_deg r (x_derivative r p) /\ poly_deg r (x_derivative r p) = poly_deg r (p:(1->num)->R) - 1 ==> poly_deg r p = 0` THEN qed[] ] );; let square_divides_if_also_divides_derivative = prove(` !(r:R ring) p f. field r ==> ring_char r = 0 ==> ring_prime(x_poly r) p ==> ring_divides(x_poly r) p f ==> ring_divides(x_poly r) p (x_derivative r f) ==> ring_divides(x_poly r) (poly_mul r p p) f `, intro THEN choose `q:(1->num)->R` `q IN ring_carrier(x_poly(r:R ring)) /\ f = ring_mul(x_poly r) p q` [ring_divides] THEN choose `Q:(1->num)->R` `Q IN ring_carrier(x_poly(r:R ring)) /\ x_derivative r f = ring_mul(x_poly r) p Q` [ring_divides] THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_divides] THEN have `ring_polynomial r (p:(1->num)->R)` [x_poly_use] THEN have `ring_polynomial r (q:(1->num)->R)` [x_poly_use] THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (q:(1->num)->R)` [ring_polynomial] THEN specialize[`r:R ring`;`p:(1->num)->R`;`q:(1->num)->R`]x_derivative_mul THEN have `x_derivative(r:R ring) (ring_mul(x_poly r) p q) = poly_add r (poly_mul r (x_derivative r p) q) (poly_mul r p (x_derivative r q))` [x_poly_use] THEN have `x_derivative(r:R ring) (ring_mul(x_poly r) p q) = ring_add(x_poly r) (poly_mul r (x_derivative r p) q) (poly_mul r p (x_derivative r q))` [x_poly_use] THEN have `x_derivative(r:R ring) (ring_mul(x_poly r) p q) = ring_add(x_poly r) (ring_mul(x_poly r) (x_derivative r p) q) (poly_mul r p (x_derivative r q))` [x_poly_use] THEN have `x_derivative(r:R ring) (ring_mul(x_poly r) p q) = ring_add(x_poly r) (ring_mul(x_poly r) (x_derivative r p) q) (ring_mul(x_poly r) p (x_derivative r q))` [x_poly_use] THEN have `ring_mul(x_poly(r:R ring)) p Q = ring_add(x_poly r) (ring_mul(x_poly r) (x_derivative r p) q) (ring_mul(x_poly r) p (x_derivative r q))` [] THEN have `ring_coprime(x_poly(r:R ring)) (p,x_derivative r p)` [coprime_prime_derivative] THEN have `PID(x_poly(r:R ring))` [PID_x_poly_field] THEN have `bezout_ring(x_poly(r:R ring))` [PID_IMP_BEZOUT_RING] THEN have `ring_polynomial r (x_derivative r (p:(1->num)->R))` [x_derivative_polynomial] THEN have `x_derivative r p IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN specialize[`x_poly(r:R ring)`;`p:(1->num)->R`;`x_derivative(r:R ring) p`]BEZOUT_RING_COPRIME THEN choose2 `x:(1->num)->R` `y:(1->num)->R` `x IN ring_carrier(x_poly(r:R ring)) /\ y IN ring_carrier(x_poly r) /\ ring_add (x_poly r) (ring_mul (x_poly r) p x) (ring_mul (x_poly r) (x_derivative r p) y) = ring_1 (x_poly r)` [] THEN have `Q IN ring_carrier(x_poly(r:R ring))` [] THEN have `q IN ring_carrier(x_poly(r:R ring))` [] THEN have `x_derivative r q IN ring_carrier(x_poly(r:R ring))` [x_poly_use;x_derivative_polynomial] THEN have `x IN ring_carrier(x_poly(r:R ring))` [] THEN have `y IN ring_carrier(x_poly(r:R ring))` [] THEN have `ring_mul (x_poly(r:R ring)) p Q = ring_add (x_poly r) (ring_mul (x_poly r) (x_derivative r p) q) (ring_mul (x_poly r) p (x_derivative r q))` [] THEN have `ring_add (x_poly(r:R ring)) (ring_mul (x_poly r) p x) (ring_mul (x_poly r) (x_derivative r p) y) = ring_1 (x_poly r)` [] THEN specialize[ `x_poly(r:R ring)` ;`Q:(1->num)->R` ;`p:(1->num)->R` ;`x_derivative r p:(1->num)->R` ;`q:(1->num)->R` ;`x_derivative r q:(1->num)->R` ;`x:(1->num)->R` ;`y:(1->num)->R` ](GENL[`r:R ring`;`Q:R`;`p:R`;`p':R`;`q:R`;`q':R`;`x:R`;`y:R`]( RING_RULE `ring_mul(r:R ring) p Q = ring_add r (ring_mul r p' q) (ring_mul r p q') /\ ring_add r (ring_mul r p x) (ring_mul r p' y) = ring_1 r ==> ring_mul r p q = ring_mul r (ring_mul r p p) (ring_sub r (ring_add r (ring_mul r Q y) (ring_mul r x q)) (ring_mul r q' y))` )) THEN subgoal `ring_divides(x_poly(r:R ring)) (ring_mul(x_poly r) p p) f` THENL [ rw[ring_divides] THEN intro THENL [ qed[RING_MUL] ; qed[RING_MUL] ; witness `ring_sub (x_poly(r:R ring)) (ring_add (x_poly r) (ring_mul (x_poly r) Q y) (ring_mul (x_poly r) x q)) (ring_mul (x_poly r) (x_derivative r q) y)` THEN intro THENL [ have `ring_add (x_poly r) (ring_mul (x_poly r) Q y) (ring_mul (x_poly r) x q) IN ring_carrier(x_poly(r:R ring))` [RING_ADD;RING_MUL] THEN qed[RING_SUB;RING_MUL] ; qed[] ] ] ; qed[x_poly_use] ] );; let poly_divides_product = prove(` !(r:R ring) p (f:X->(1->num)->R) S t. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (f s)) ==> t IN S ==> ring_divides(x_poly r) p (f t) ==> ring_divides(x_poly r) p (poly_product r S f) `, intro THEN have `FINITE (S DELETE (t:X))` [FINITE_DELETE] THEN specialize[`r:R ring`;`f:X->(1->num)->R`;`S DELETE (t:X)`;`t:X`]poly_product_insert THEN set_fact `(t:X) IN S ==> t INSERT (S DELETE t) = S` THEN have `ring_powerseries(r:R ring) (f(t:X):(1->num)->R)` [ring_polynomial] THEN set_fact `~(t:X IN (S DELETE t))` THEN have `poly_product(r:R ring) S f = poly_mul r (f(t:X)) (poly_product r (S DELETE t) f)` [] THEN have `poly_product(r:R ring) S f = ring_mul(x_poly r) (f(t:X)) (poly_product r (S DELETE t) f)` [x_poly_use] THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (f s:(1->num)->R)` [IN_DELETE] THEN qed [RING_DIVIDES_RMUL;poly_product_poly;x_poly_use] );; let poly_divides_sum = prove(` !(r:R ring) p (f:X->(1->num)->R) S. FINITE S ==> (!s:X. s IN S ==> ring_divides(x_poly r) p (f s)) ==> ring_polynomial r p ==> ring_divides(x_poly r) p (poly_sum r S f) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty;x_poly_use] THEN qed[RING_DIVIDES_0;x_poly_use] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN x INSERT S ==> f s IN ring_carrier(x_poly(r:R ring))` [ring_divides] THEN have `ring_powerseries r (f(x:X):(1->num)->R)` [x_poly_use;ring_polynomial] THEN simp[poly_sum_insert] THEN have `!s:X. s IN S ==> ring_divides(x_poly(r:R ring)) p (f s)` [] THEN have `ring_divides(x_poly(r:R ring)) p (f(x:X))` [] THEN have `ring_divides(x_poly(r:R ring)) p (poly_sum r S (f:X->(1->num)->R))` [] THEN qed[RING_DIVIDES_ADD;x_poly_use] ] );; let poly_divides_delta = prove(` !(r:R ring) p t (f:X->(1->num)->R) S. FINITE S ==> t IN S ==> ring_polynomial r (f t) ==> (!s:X. s IN S ==> ~(s = t) ==> ring_divides(x_poly r) p (f s)) ==> ring_divides(x_poly r) p (poly_sum r S f) ==> ring_divides(x_poly r) p (f t) `, intro THEN have `f(t:X) IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `ring_sum(x_poly(r:R ring)) (S DELETE (t:X)) f = ring_sub(x_poly r) (ring_sum(x_poly r) S f) (f t)` [RING_SUM_DELETE] THEN have `!s:X. s IN S ==> ring_polynomial r (f(s:X):(1->num)->R)` [ring_divides;x_poly_use] THEN set_fact `!s:X. s IN S DELETE t ==> s IN S` THEN have `!s:X. s IN S DELETE t ==> ring_polynomial r (f(s:X):(1->num)->R)` [ring_divides;x_poly_use] THEN have `FINITE (S DELETE(t:X))` [FINITE_DELETE] THEN have `poly_sum(r:R ring) (S DELETE (t:X)) f = ring_sub(x_poly r) (ring_sum(x_poly r) S f) (f t)` [poly_sum_ring_sum_x_poly;x_poly_use] THEN have `poly_sum(r:R ring) (S DELETE (t:X)) f = ring_sub(x_poly r) (poly_sum r S f) (f t)` [poly_sum_ring_sum_x_poly] THEN have `poly_sum(r:R ring) (S DELETE(t:X)) f IN ring_carrier(x_poly(r:R ring))` [poly_sum_poly;x_poly_use] THEN have `poly_sum(r:R ring) (S:X->bool) f IN ring_carrier(x_poly(r:R ring))` [poly_sum_poly;x_poly_use] THEN specialize[`x_poly(r:R ring)` ;`poly_sum(r:R ring) (S DELETE(t:X)) f` ;`poly_sum(r:R ring) (S:X->bool) f` ;`f(t:X):(1->num)->R` ](GENL[`r:R ring`;`A:R`;`B:R`;`C:R` ](RING_RULE `A = ring_sub(r:R ring) B C ==> C = ring_sub r B A`)) THEN set_fact `!s:X. s IN S DELETE t ==> ~(s = t)` THEN have `!s:X. s IN S DELETE t ==> ring_divides(x_poly r) p (f(s:X):(1->num)->R)` [] THEN have `ring_polynomial r (p:(1->num)->R)` [ring_divides;x_poly_use] THEN specialize[`r:R ring`;`p:(1->num)->R`;`f:X->(1->num)->R`;`S DELETE (t:X)`]poly_divides_sum THEN qed[RING_DIVIDES_SUB] );; let coprime_product = prove(` !(r:R ring) p (f:X->R) S. FINITE S ==> (UFD r \/ (integral_domain r /\ bezout_ring r)) ==> p IN ring_carrier r ==> (!s:X. s IN S ==> ring_coprime r (p,f s)) ==> ring_coprime r (p,ring_product r S f) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN STRIP_TAC THENL [ REPEAT DISCH_TAC THEN rw[RING_PRODUCT_CLAUSES] THEN simp[ring_coprime_1] ; REPEAT DISCH_TAC THEN REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN set_fact `(x:X) IN x INSERT S` THEN have `ring_coprime(r:R ring) (p,f(x:X))` [] THEN have `f(x:X) IN ring_carrier(r:R ring)` [ring_coprime] THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN S ==> ring_coprime(r:R ring) (p,f s)` [] THEN have `!s:X. s IN S ==> f s IN ring_carrier(r:R ring)` [ring_coprime] THEN simp[RING_PRODUCT_CLAUSES] THEN qed[RING_COPRIME_RMUL;RING_PRODUCT] ] );; let divides_factor_and_derivative_product = prove(` !(r:R ring) S t p (f:X->(1->num)->R). field r ==> FINITE S ==> t IN S ==> ring_divides(x_poly r) p (f t) ==> (!s:X. s IN S DELETE t ==> ring_coprime(x_poly r) (p,f s)) ==> ring_divides(x_poly r) p (x_derivative r (poly_product r S f)) ==> ring_divides(x_poly r) p (x_derivative r (f t)) `, intro THEN subgoal `!s:X. s IN S ==> f s IN ring_carrier(x_poly(r:R ring))` THENL [ intro THEN case `s = t:X` THENL [ qed[ring_divides] ; qed[ring_coprime;IN_DELETE] ] ; pass ] THEN have `!s:X. s IN S ==> ring_polynomial r (f s:(1->num)->R)` [x_poly_use] THEN have `!s:X. s IN S ==> ring_powerseries r (f s:(1->num)->R)` [ring_polynomial] THEN specialize[`r:R ring`;`f:X->(1->num)->R`;`S:X->bool`]x_derivative_product THEN have `ring_divides(x_poly(r:R ring)) p (poly_sum r S (\s:X. poly_mul r (x_derivative r (f s)) (poly_product r (S DELETE s) f)))` [x_poly_use] THEN subgoal `!s:X. s IN S ==> ~(s = t) ==> ring_divides(x_poly(r:R ring)) p (poly_mul r (x_derivative r (f s)) (poly_product r (S DELETE s) f))` THENL [ intro THEN rw[x_poly_use] THEN sufficesby RING_DIVIDES_LMUL THEN set_fact `t IN S /\ s IN S /\ ~(s = t:X) ==> t IN S DELETE s` THEN have `FINITE (S DELETE (s:X))` [FINITE_DELETE] THEN have `!u:X. u IN S DELETE s ==> ring_polynomial(r:R ring) (f u:(1->num)->R)` [IN_DELETE] THEN specialize[`r:R ring`;`p:(1->num)->R`;`f:X->(1->num)->R`;`S DELETE (s:X)`;`t:X`]poly_divides_product THEN qed[x_derivative_polynomial;x_poly_use] ; pass ] THEN have `FINITE (S DELETE (t:X))` [FINITE_DELETE] THEN have `!s:X. s IN S DELETE t ==> ring_polynomial r (f s:(1->num)->R)` [IN_DELETE] THEN have `ring_polynomial(r:R ring) (poly_mul r (x_derivative r (f(t:X))) (poly_product r (S DELETE t) f))` [poly_product_poly;x_derivative_polynomial;RING_POLYNOMIAL_MUL] THEN specialize[`r:R ring`;`p:(1->num)->R`;`t:X`;`\s:X. poly_mul(r:R ring) (x_derivative r (f s)) (poly_product r (S DELETE s) f)`;`S:X->bool`]poly_divides_delta THEN subgoal `ring_coprime(x_poly(r:R ring)) (p,poly_product r (S DELETE (t:X)) f)` THENL [ have `UFD(x_poly(r:R ring)) \/ (integral_domain(x_poly r) /\ bezout_ring(x_poly r))` [PID_x_poly_field;PID_IMP_UFD] THEN have `p IN ring_carrier(x_poly(r:R ring))` [ring_divides] THEN specialize[`x_poly(r:R ring)`;`p:(1->num)->R`;`f:X->(1->num)->R`;`S DELETE (t:X)`]coprime_product THEN qed[poly_product_ring_product_x_poly] ; have `UFD (x_poly(r:R ring)) \/ integral_domain (x_poly r) /\ bezout_ring (x_poly r)` [PID_x_poly_field;PID_IMP_UFD] THEN have `x_derivative r (f(t:X)) IN ring_carrier(x_poly(r:R ring))` [x_derivative_polynomial;x_poly_use] THEN have `poly_product r (S DELETE (t:X)) f IN ring_carrier (x_poly(r:R ring))` [poly_product_poly;x_poly_use] THEN have `ring_divides (x_poly(r:R ring)) p (ring_mul (x_poly r) (x_derivative r (f(t:X))) (poly_product r (S DELETE t) f))` [x_poly_use] THEN have `ring_coprime (x_poly(r:R ring)) (poly_product r (S DELETE(t:X)) f,p)` [RING_COPRIME_SYM] THEN specialize[`x_poly(r:R ring)`;`x_derivative(r:R ring) (f(t:X))`;`poly_product(r:R ring) (S DELETE(t:X)) f`;`p:(1->num)->R`]RING_COPRIME_DIVPROD_RIGHT THEN qed[] ] );; let coprime_derivative_if_squarefree = prove(` !(r:R ring) f. field r ==> ring_char r = 0 ==> ring_polynomial r f ==> ~(f = poly_0 r) ==> ring_squarefree(x_poly r) f ==> ring_coprime(x_poly r) (f,x_derivative r f) `, intro THEN have `PID (x_poly(r:R ring))` [PID_x_poly_field] THEN have `integral_domain(x_poly(r:R ring))` [PID_IMP_INTEGRAL_DOMAIN] THEN have `f IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `~(f = ring_0(x_poly(r:R ring)))` [x_poly_use] THEN proven_if `ring_unit(x_poly(r:R ring)) f` [ring_coprime_if_unit;x_derivative_polynomial;x_poly_use] THEN have `UFD (x_poly(r:R ring))` [PID_IMP_UFD] THEN ASSUME_TAC(UNDISCH_ALL (ISPECL[`f:(1->num)->R`] (CONJUNCT2 (UNDISCH_ALL (fst (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL [`x_poly(r:R ring)`]UFD_EQ_PRIMEFACT_NONUNIT))))))))) THEN choose2 `n:num` `q:num->(1->num)->R` `1 <= n /\ (!i. 1 <= i /\ i <= n ==> ring_prime(x_poly(r:R ring)) (q i)) /\ ring_product(x_poly r) (1..n) q = f` [] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!i:num. i IN (1..n) ==> q i IN ring_carrier(x_poly(r:R ring))` [ring_prime;IN_NUMSEG] THEN subgoal `!i j. 1 <= i /\ i <= n /\ 1 <= j /\ j <= n /\ ring_divides(x_poly(r:R ring)) (q i) (q j) /\ ~(i = j) ==> F` THENL [ intro THEN have `i IN (1..n)` [IN_NUMSEG] THEN have `j IN (1..n)` [IN_NUMSEG] THEN specialize[`x_poly(r:R ring)`;`1..n`;`q:num->(1->num)->R`;`i:num`;`j:num`]square_divides_product_if_factor_divides_factor THEN have `ring_divides(x_poly(r:R ring)) (poly_mul r (q(i:num)) (q(i))) f` [x_poly_use] THEN have `q(i:num):(1->num)->R IN ring_carrier(x_poly r)` [] THEN have `~(ring_unit(x_poly r) (q(i:num):(1->num)->R))` [ring_prime] THEN have `ring_divides(x_poly r) (ring_mul(x_poly r) (q(i:num)) (q(i))) (f:(1->num)->R)` [x_poly_use] THEN specialize[`x_poly(r:R ring)`;`f:(1->num)->R`;`q(i:num):(1->num)->R`]not_squarefree_if_divisible_by_square THEN qed[] ; pass ] THEN simp[UFD_COPRIME] THEN intro THENL [ qed[x_poly_use;x_derivative_polynomial] ; pass ] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN specialize[`x_poly(r:R ring)`;`p:(1->num)->R`;`1..n`;`q:num->(1->num)->R`]RING_PRIME_DIVIDES_PRODUCT THEN choose `i:num` `i IN 1..n /\ ring_divides(x_poly(r:R ring)) p (q i)` [] THEN have `i IN 1..n` [] THEN have `ring_divides(x_poly(r:R ring)) p (q(i:num))` [] THEN have `!i:num. i IN 1..n ==> ring_polynomial(r:R ring) (q i:(1->num)->R)` [x_poly_use] THEN specialize[`r:R ring`;`q:num->(1->num)->R`;`1..n`]poly_product_ring_product_x_poly THEN have `ring_divides(x_poly(r:R ring)) p (x_derivative r (poly_product r (1..n) q))` [] THEN subgoal `!j:num. j IN (1..n) DELETE i ==> ring_coprime(x_poly(r:R ring)) (p,q j)` THENL [ intro THEN have `j IN (1..n)` [IN_DELETE] THEN have `q(j:num) IN ring_carrier(x_poly(r:R ring))` [] THEN specialize[`x_poly(r:R ring)`;`p:(1->num)->R`;`q(j:num):(1->num)->R`]INTEGRAL_DOMAIN_PRIME_DIVIDES_OR_COPRIME THEN case `ring_divides(x_poly(r:R ring)) p (q(j:num))` THENL [ have `1 <= i /\ i <= n` [IN_NUMSEG] THEN have `1 <= j /\ j <= n` [IN_NUMSEG] THEN have `~(i = j:num)` [IN_DELETE] THEN have `ring_divides(x_poly(r:R ring)) (q(i:num)) (q j)` [prime_divides_prime_and;PID_IMP_UFD] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN specialize[`r:R ring`;`1..n`;`i:num`;`p:(1->num)->R`;`q:num->(1->num)->R`]divides_factor_and_derivative_product THEN have `1 <= i /\ i <= n` [IN_NUMSEG] THEN have `ring_prime(x_poly(r:R ring)) (q(i:num))` [] THEN specialize[`r:R ring`;`q(i:num):(1->num)->R`]coprime_prime_derivative THEN qed[ring_coprime;ring_prime] );; let gcd_poly_linear_combination = prove(` !(r:R ring) a b. field r ==> ring_polynomial r a ==> ring_polynomial r b ==> ?x y. ring_polynomial r x /\ ring_polynomial r y /\ poly_add r (poly_mul r a x) (poly_mul r b y) = ring_gcd(x_poly r) (a,b) `, intro THEN have `PID(x_poly(r:R ring))` [PID_x_poly_field] THEN have `bezout_ring(x_poly(r:R ring))` [PID_IMP_BEZOUT_RING] THEN have `a IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `b IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN specialize[`x_poly(r:R ring)`;`a:(1->num)->R`;`b:(1->num)->R`]BEZOUT_RING_IMP_GCD THEN choose2 `x:(1->num)->R` `y:(1->num)->R` `x IN ring_carrier(x_poly r) /\ y IN ring_carrier(x_poly r) /\ ring_add (x_poly r) (ring_mul (x_poly r) a x) (ring_mul (x_poly r) b y) = ring_gcd (x_poly r) (a,b):(1->num)->R` [] THEN witness `x:(1->num)->R` THEN witness `y:(1->num)->R` THEN have `ring_mul(x_poly(r:R ring)) a x = poly_mul r a x` [x_poly_use] THEN have `ring_mul(x_poly(r:R ring)) b y = poly_mul r b y` [x_poly_use] THEN qed[x_poly_use] );; let linear_combination_if_coprime_poly = prove(` !(r:R ring) a b. field r ==> ring_coprime(x_poly r) (a,b) ==> ?x y. ring_polynomial r x /\ ring_polynomial r y /\ poly_add r (poly_mul r a x) (poly_mul r b y) = poly_1 r `, intro THEN have `PID(x_poly(r:R ring))` [PID_x_poly_field] THEN have `bezout_ring(x_poly(r:R ring))` [PID_IMP_BEZOUT_RING] THEN specialize_forward[`x_poly(r:R ring)`;`a:(1->num)->R`;`b:(1->num)->R`]BEZOUT_RING_COPRIME THEN choose2 `x:(1->num)->R` `y:(1->num)->R` `x IN ring_carrier(x_poly r) /\ y IN ring_carrier(x_poly r) /\ ring_add (x_poly r) (ring_mul (x_poly r) a x) (ring_mul (x_poly r) b y) = ring_1 (x_poly r):(1->num)->R` [] THEN witness `x:(1->num)->R` THEN witness `y:(1->num)->R` THEN have `ring_mul(x_poly(r:R ring)) a x = poly_mul r a x` [x_poly_use] THEN have `ring_mul(x_poly(r:R ring)) b y = poly_mul r b y` [x_poly_use] THEN qed[x_poly_use] );; let coprime_poly_if_linear_combination = prove(` !(r:R ring) a b x y. ring_polynomial r a ==> ring_polynomial r b ==> ring_polynomial r x ==> ring_polynomial r y ==> poly_add r (poly_mul r a x) (poly_mul r b y) = poly_1 r ==> ring_coprime(x_poly r) (a,b) `, rw[ring_coprime] THEN intro THENL [ qed[x_poly_use] ; qed[x_poly_use] ; have `ring_mul(x_poly(r:R ring)) a x = poly_mul r a x` [x_poly_use] THEN have `ring_mul(x_poly(r:R ring)) b y = poly_mul r b y` [x_poly_use] THEN have `ring_add(x_poly(r:R ring)) (ring_mul(x_poly r) a x) (ring_mul(x_poly r) b y) = ring_1(x_poly r)` [x_poly_use] THEN choose `u:(1->num)->R` `u IN ring_carrier(x_poly(r:R ring)) /\ a = ring_mul(x_poly r) d u` [ring_divides] THEN choose `v:(1->num)->R` `v IN ring_carrier(x_poly(r:R ring)) /\ b = ring_mul(x_poly r) d v` [ring_divides] THEN have `a IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `b IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `d IN ring_carrier(x_poly(r:R ring))` [ring_divides;x_poly_use] THEN have `u IN ring_carrier(x_poly(r:R ring))` [] THEN have `v IN ring_carrier(x_poly(r:R ring))` [] THEN have `x IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `y IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN have `a = ring_mul(x_poly(r:R ring)) d u` [x_poly_use] THEN have `b = ring_mul(x_poly(r:R ring)) d v` [x_poly_use] THEN specialize[ `x_poly(r:R ring)`; `a:(1->num)->R`; `b:(1->num)->R`; `d:(1->num)->R`; `u:(1->num)->R`; `v:(1->num)->R`; `x:(1->num)->R`; `y:(1->num)->R` ]( GENL[ `r:R ring`;`a:R`;`b:R`;`d:R`;`u:R`;`v:R`;`x:R`;`y:R` ]( RING_RULE `a = ring_mul(r:R ring) d u /\ b = ring_mul r d v /\ ring_add r (ring_mul r a x) (ring_mul r b y) = ring_1 r ==> ring_mul r d (ring_add r (ring_mul r u x) (ring_mul r v y)) = ring_1 r` ) ) THEN rw[ring_unit] THEN intro THENL [ qed[ring_divides] ; witness `ring_add (x_poly(r:R ring)) (ring_mul (x_poly r) u x) (ring_mul (x_poly r) v y)` THEN qed[x_poly_use;ring_unit;RING_ADD;RING_MUL] ] ] );; let nonzero_poly_mul = prove(` !(r:R ring) p (q:(1->num)->R). integral_domain r ==> ring_powerseries r p ==> ring_powerseries r q ==> ~(p = poly_0 r) ==> ~(q = poly_0 r) ==> ~(poly_mul r p q = poly_0 r) `, rw[x_series_use] THEN intro THEN have `integral_domain(x_series(r:R ring))` [INTEGRAL_DOMAIN_POWSER_RING;x_series] THEN qed[integral_domain] );; let nonzero_poly_pow = prove(` !(r:R ring) (p:(1->num)->R) n. integral_domain r ==> ring_powerseries r p ==> ~(p = poly_0 r) ==> ~(poly_pow r p n = poly_0 r) `, rw[x_series_use;x_series_use_pow] THEN intro THEN have `integral_domain(x_series(r:R ring))` [INTEGRAL_DOMAIN_POWSER_RING;x_series] THEN qed[INTEGRAL_DOMAIN_POW_EQ_0] );; let nonzero_poly_product = prove(` !(r:R ring) S p. integral_domain r ==> FINITE S ==> (!s:X. s IN S ==> ~(p s = poly_0 r)) ==> ~(poly_product r S p = poly_0 r) `, rw[poly_product] THEN intro THEN have `integral_domain(x_series(r:R ring))` [INTEGRAL_DOMAIN_POWSER_RING;x_series] THEN qed[INTEGRAL_DOMAIN_PRODUCT_EQ_0;x_series_use] );; let irred_x_minus_const = prove(` !(r:R ring) c. field r ==> c IN ring_carrier r ==> ( ring_polynomial r (x_minus_const r c) /\ ring_irreducible (x_poly r) (x_minus_const r c) /\ monic r (x_minus_const r c) ) `, intro_gendisch THEN have `ring_polynomial(r:R ring) (x_minus_const r c)` [x_minus_const_poly] THEN have `poly_deg(r:R ring) (x_minus_const r c) = 1` [deg_x_minus_const;field] THEN have `ring_prime(x_poly(r:R ring)) (x_minus_const r c)` [prime_if_deg_1] THEN have `ring_irreducible(x_poly(r:R ring)) (x_minus_const r c)` [prime_iff_irreducible_over_field] THEN have `monic(r:R ring) (x_minus_const r c)` [monic_x_minus_const] THEN qed[] );; let monic_factorization = prove(` !(r:R ring) f. field r ==> ring_polynomial r f ==> monic r f ==> ?n p. ( (!i. i IN (1..n) ==> ring_polynomial r (p i)) /\ (!i. i IN (1..n) ==> monic r (p i)) /\ (!i. i IN (1..n) ==> ring_irreducible(x_poly r) (p i)) /\ poly_product r (1..n) p = f ) `, intro THEN have `UFD(x_poly(r:R ring))` [UFD_x_poly_field] THEN case `poly_deg r (f:(1->num)->R) = 0` THENL [ witness `0:num` THEN witness `\i:num. poly_0 r:(1->num)->R` THEN have `!i. ~(i IN 1..0)` [IN_NUMSEG;ARITH_RULE `~(1 <= i /\ i <= 0)`] THEN set_fact `(!i. ~(i IN 1..0)) ==> 1..0 = {}` THEN simp[poly_product_empty] THEN qed[poly_1_if_monic_deg_0] ; pass ] THEN have `~(f = poly_0(r:R ring):(1->num)->R)` [monic_poly_0;field] THEN have `~(f = ring_0(x_poly(r:R ring)))` [x_poly_use] THEN have `~(ring_unit(x_poly(r:R ring)) f)` [deg_0_if_unit] THEN have `f IN ring_carrier(x_poly(r:R ring))` [x_poly_use] THEN ASSUME_TAC(UNDISCH_ALL (ISPECL[`f:(1->num)->R`] (CONJUNCT2 (UNDISCH_ALL (fst (EQ_IMP_RULE (UNDISCH_ALL (REWRITE_RULE [IMP_CONJ] (ISPECL [`x_poly(r:R ring)`]UFD_EQ_PRIMEFACT_NONUNIT))))))))) THEN choose2 `n:num` `p:num->((1->num)->R)` `1 <= n /\ (!i. 1 <= i ==> i <= n ==> ring_prime (x_poly r) (p i)) /\ ring_product (x_poly r) (1..n) p = f:(1->num)->R` [] THEN witness `n:num` THEN def `q:num->(1->num)->R` `\i:num. @qi. ring_polynomial(r:R ring) qi /\ monic r qi /\ ring_associates(x_poly r) (p i) qi` THEN witness `q:num->(1->num)->R` THEN subgoal `!i. i IN 1..n ==> ring_polynomial (r:R ring) (q i) /\ monic r (q i) /\ ring_associates(x_poly r) (p i) (q i)` THENL [ rw[IN_NUMSEG] THEN intro THEN have `p(i:num) IN ring_carrier(x_poly(r:R ring))` [ring_prime] THEN have `ring_polynomial r (p(i:num):(1->num)->R)` [x_poly_use] THEN have `~(p(i:num):((1->num)->R) = poly_0 r)` [ring_prime;x_poly_use] THEN specialize[ `r:R ring`; `p(i:num):(1->num)->R` ]x_poly_field_monic_associate THEN qed[] ; pass ] THEN have `!i. i IN 1..n ==> ring_prime(x_poly(r:R ring)) (q i)` [RING_ASSOCIATES_PRIME;IN_NUMSEG] THEN specialize_assuming[ `x_poly(r:R ring)`; `1..n`; `p:num->(1->num)->R`; `q:num->(1->num)->R` ]RING_ASSOCIATES_PRODUCT THEN have `FINITE(1..n)` [FINITE_NUMSEG] THEN have `ring_associates(x_poly(r:R ring)) (ring_product(x_poly r) (1..n) p) (ring_product(x_poly r) (1..n) q)` [] THEN have `ring_associates(x_poly(r:R ring)) f (ring_product(x_poly r) (1..n) q)` [FINITE_NUMSEG] THEN have `ring_product(x_poly r) (1..n) q = poly_product(r:R ring) (1..n) q` [poly_product_ring_product_x_poly] THEN specialize_assuming[ `r:R ring`; `q:num->(1->num)->R`; `1..n` ]monic_poly_product THEN have `monic(r:R ring) (poly_product r (1..n) q)` [] THEN have `ring_associates(x_poly(r:R ring)) f (poly_product r (1..n) q)` [] THEN specialize_assuming[ `r:R ring`; `f:(1->num)->R`; `(poly_product r (1..n) q):(1->num)->R` ]monic_associates THEN qed[prime_iff_irreducible_over_field] );; let monic_factorization_exponents = prove(` !(r:R ring) f. field r ==> ring_polynomial r f ==> monic r f ==> ?P e. ( FINITE P /\ (!p. p IN P ==> ring_polynomial r p) /\ (!p. p IN P ==> monic r p) /\ (!p. p IN P ==> ring_irreducible(x_poly r) p) /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product r P (\p. poly_pow r p (e p)) = f ) `, intro THEN choose2 `n:num` `q:num->(1->num)->R` `(!i. i IN 1..n ==> ring_polynomial(r:R ring) (q i)) /\ (!i. i IN 1..n ==> monic r (q i)) /\ (!i. i IN 1..n ==> ring_irreducible (x_poly r) (q i)) /\ poly_product r (1..n) q = f` [monic_factorization] THEN def `P:((1->num)->R)->bool` `IMAGE q (1..n):((1->num)->R)->bool` THEN def `e:((1->num)->R)->num` `\p:(1->num)->R. CARD({i | i IN 1..n /\ q i = p})` THEN witness `P:((1->num)->R)->bool` THEN witness `e:((1->num)->R)->num` THEN have `FINITE(1..n)` [FINITE_NUMSEG] THEN intro THENL [ qed[FINITE_IMAGE] ; choose `i:num` `p:(1->num)->R = q i /\ i IN 1..n` [IN_IMAGE] THEN qed[] ; choose `i:num` `p:(1->num)->R = q i /\ i IN 1..n` [IN_IMAGE] THEN qed[] ; choose `i:num` `p:(1->num)->R = q i /\ i IN 1..n` [IN_IMAGE] THEN qed[] ; specialize[ `1..n`; `\i:num. q i = p:(1->num)->R` ]FINITE_RESTRICT THEN have `{i | i IN 1..n /\ q i = p:(1->num)->R} = {}` [CARD_EQ_0] THEN choose `i:num` `p:(1->num)->R = q i /\ i IN 1..n` [IN_IMAGE] THEN havetac `i IN {i | i IN 1..n /\ q i = p:(1->num)->R}` (simp[IN_ELIM_THM]) THEN have `i:num IN {}` [] THEN set_fact `~(i:num IN {})` THEN qed[] ; have `ring_product(x_poly(r:R ring)) (1..n) q = f` [poly_product_ring_product_x_poly] THEN specialize[ `x_poly(r:R ring)`; `q:num->(1->num)->R`; `q:num->(1->num)->R`; `1..n` ]RING_PRODUCT_IMAGE_GEN THEN have `f = ring_product (x_poly(r:R ring)) (IMAGE q (1..n)) (\y. ring_product (x_poly r) {x | x IN 1..n /\ q x = y} q)` [] THEN subgoal `ring_product (x_poly(r:R ring)) (IMAGE q (1..n)) (\y. ring_product (x_poly r) {x | x IN 1..n /\ q x = y} q) = ring_product (x_poly r) (IMAGE q (1..n)) (\p. ring_pow(x_poly r) p (e p))` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[IN_IMAGE;BETA_THM] THEN intro THEN subgoal `ring_product (x_poly(r:R ring)) {x | x IN 1..n /\ q x = a} q = ring_product (x_poly r) {x | x IN 1..n /\ q x = a} (\x. a)` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[] THEN set_fact `{x' | x' IN 1..n /\ q x' = q x:(1->num)->R} SUBSET 1..n` THEN have `FINITE {x' | x' IN 1..n /\ q x' = q x:(1->num)->R}` [FINITE_SUBSET] THEN have `q(x:num):(1->num)->R IN ring_carrier(x_poly r)` [x_poly_use] THEN simp[ring_product_const] ; pass ] THEN have `f = ring_product (x_poly(r:R ring)) (IMAGE q (1..n)) (\p. ring_pow (x_poly r) p (e p))` [] THEN rw[x_poly_use_pow] THEN subgoal `!p:(1->num)->R. p IN P ==> ring_polynomial r (ring_pow (x_poly r) p (e p))` THENL [ intro THEN rw[GSYM x_poly_use_pow] THEN sufficesby poly_pow_poly THEN choose `i:num` `p:(1->num)->R = q i /\ i IN 1..n` [IN_IMAGE] THEN qed[] ; pass ] THEN have `FINITE(P:((1->num)->R)->bool)` [FINITE_IMAGE] THEN specialize[ `r:R ring`; `\p:(1->num)->R. ring_pow (x_poly r) p (e p)`; `P:((1->num)->R)->bool` ]poly_product_ring_product_x_poly THEN qed[] ] );; (* ===== polynomial reversal *) let x_truncreverse = new_definition ` x_truncreverse (r:R ring) n (p:(1->num)->R) = \m. if m one <= n then p (\v. n - m one) else ring_0 r `;; let coeff_x_truncreverse = prove(` !(r:R ring) n p d. coeff d (x_truncreverse r n p) = if d <= n then coeff (n - d) p else ring_0 r `, rw[x_truncreverse;coeff;x_monomial] );; let x_truncreverse_series = prove(` !(r:R ring) n p. ring_powerseries r p ==> ring_powerseries r (x_truncreverse r n p) `, rw[coeff_series_in_ring] THEN rw[coeff_x_truncreverse] THEN qed[RING_0] );; let x_truncreverse_poly = prove(` !(r:R ring) n p. ring_powerseries r p ==> ring_polynomial r (x_truncreverse r n p) `, intro THEN have `ring_powerseries(r:R ring) (x_truncreverse r n p)` [x_truncreverse_series] THEN have `!d. ~(coeff d (x_truncreverse r n p) = ring_0(r:R ring)) ==> d <= n` [coeff_x_truncreverse] THEN qed[polynomial_if_coeff] );; let deg_x_truncreverse_le = prove(` !(r:R ring) n p. ring_powerseries r p ==> poly_deg r (x_truncreverse r n p) <= n `, intro THEN have `ring_powerseries(r:R ring) (x_truncreverse r n p)` [x_truncreverse_series] THEN have `!d. ~(coeff d (x_truncreverse r n p) = ring_0(r:R ring)) ==> d <= n` [coeff_x_truncreverse] THEN qed[deg_le_coeff] );; let x_truncreverse_poly_add = prove(` !(r:R ring) n p q. x_truncreverse r n (poly_add r p q) = poly_add r (x_truncreverse r n p) (x_truncreverse r n q) `, intro THEN sufficesby eq_coeff THEN rw[coeff_x_truncreverse;coeff_poly_add] THEN qed[RING_ADD_LZERO;RING_0] );; let x_truncreverse_poly_sub = prove(` !(r:R ring) n p q. x_truncreverse r n (poly_sub r p q) = poly_sub r (x_truncreverse r n p) (x_truncreverse r n q) `, intro THEN sufficesby eq_coeff THEN rw[coeff_x_truncreverse;coeff_poly_sub] THEN qed[RING_SUB_RZERO;RING_0] );; let x_truncreverse_poly_const = prove(` !(r:R ring) n c. x_truncreverse r n (poly_const r c) = const_x_pow r c n `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_x_truncreverse] THEN rw[coeff_const_x_pow] THEN rw[coeff_poly_const] THEN case `d = n:num` THENL [ num_linear_fact `n <= n:num` THEN num_linear_fact `n - n = 0:num` THEN qed[] ; num_linear_fact `~(d = n) ==> ~(d <= n) \/ ~(n - d = 0)` THEN qed[] ] );; let x_truncreverse_0_poly_const = prove(` !(r:R ring) c. x_truncreverse r 0 (poly_const r c) = poly_const r c `, rw[x_truncreverse_poly_const;const_x_pow_0] );; let x_truncreverse_poly_1 = prove(` !(r:R ring) n. x_truncreverse r n (poly_1 r) = x_pow r n `, rw[poly_1;x_truncreverse_poly_const;x_pow] );; let x_truncreverse_poly_0 = prove(` !(r:R ring) n. x_truncreverse r n (poly_0 r) = poly_0 r `, rw[poly_0;x_truncreverse_poly_const;const_0_x_pow] );; let x_truncreverse_poly_sum = prove(` !(r:R ring) n p:X->(1->num)->R S:X->bool. FINITE S ==> (!s. s IN S ==> ring_powerseries r (p s)) ==> x_truncreverse r n (poly_sum r S p) = poly_sum r S (\s. x_truncreverse r n (p s)) `, rw[poly_sum] THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;GSYM x_series_use;x_truncreverse_poly_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [] THEN have `p(x:X) IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN have `ring_powerseries(r:R ring) (x_truncreverse r n (p(x:X)))` [x_truncreverse_series] THEN have `x_truncreverse r n (p(x:X)) IN ring_carrier(x_series(r:R ring))` [x_series_use] THEN simp[RING_SUM_CLAUSES] THEN simp[GSYM x_series_use] THEN simp[x_truncreverse_poly_add] ] );; let x_truncreverse_const_x_pow = prove(` !(r:R ring) n c. x_truncreverse r n (const_x_pow r c n) = poly_const r c `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_x_truncreverse] THEN rw[coeff_const_x_pow] THEN rw[coeff_poly_const] THEN case `d = 0:num` THENL [ num_linear_fact `0 <= n:num` THEN num_linear_fact `n-0 = n:num` THEN qed[] ; num_linear_fact `~(d = 0) ==> ~(d <= n) \/ ~(n - d = n)` THEN qed[] ] );; let x_truncreverse_x_pow = prove(` !(r:R ring) n. x_truncreverse r n (x_pow r n) = poly_1 r `, rw[x_pow;x_truncreverse_const_x_pow;poly_1] );; let x_truncreverse_x_minus_const = prove(` !(r:R ring) c. x_truncreverse r 1 (x_minus_const r c) = one_minus_constx r c `, intro THEN rw[x_minus_const;one_minus_constx] THEN rw[x_truncreverse_poly_sub] THEN rw[x_truncreverse_x_pow] THEN rw[x_truncreverse_poly_const] THEN rw[x_pow_0] );; let x_truncreverse_one_minus_constx = prove(` !(r:R ring) c. x_truncreverse r 1 (one_minus_constx r c) = x_minus_const r c `, intro THEN rw[x_minus_const;one_minus_constx] THEN rw[x_truncreverse_poly_sub] THEN rw[x_pow_0] THEN rw[x_truncreverse_poly_1] THEN rw[x_truncreverse_const_x_pow] );; let x_truncreverse_mul = prove(` !(r:R ring) m p n q. ring_polynomial r p ==> ring_polynomial r q ==> poly_deg r p <= m ==> poly_deg r q <= n ==> x_truncreverse r (m+n) (poly_mul r p q) = poly_mul r (x_truncreverse r m p) (x_truncreverse r n q) `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_poly_mul_oneindex;coeff_x_truncreverse] THEN have `!d. ~(coeff d p = ring_0(r:R ring)) ==> d <= m` [coeff_deg_le] THEN have `!d. ~(coeff d q = ring_0(r:R ring)) ==> d <= n` [coeff_deg_le] THEN case `m+n < d:num` THENL [ simp[ARITH_RULE `m+n < d:num ==> ~(d <= m+n)`] THEN sufficesby(GSYM RING_SUM_EQ_0) THEN intro THEN case `a <= m:num` THENL [ num_linear_fact `a <= m:num /\ m+n < d ==> ~(d-a <= n)` THEN qed[coeff_poly_in_ring;RING_MUL_RZERO] ; qed[coeff_poly_in_ring;RING_0;RING_MUL_LZERO] ] ; pass ] THEN simp[ARITH_RULE `~(m+n < d:num) ==> d <= m+n`] THEN specialize[`r:R ring`;`(m+n)-d:num`;`\a. ring_mul(r:R ring) (coeff a p) (coeff ((m + n) - d - a) q)`]ring_sum_numseg_le_reflect THEN specialize[`r:R ring`;`(m+n)-d:num`;`d:num`;`\a. ring_mul(r:R ring) (coeff ((m + n) - d - a) p) (coeff ((m + n) - d - ((m + n) - d - a)) q)`]ring_sum_numseg_le_offset THEN simp[ARITH_RULE `~(m+n < d:num) ==> (m+n)-d+d = m+n`] THEN num_linear_fact `m+n <= m+n+d:num` THEN specialize[`r:R ring`;`m+n:num`;`m+n+d:num`;`\b. if d <= b then ring_mul r (coeff ((m + n) - d - (b - d)) p) (coeff ((m + n) - d - ((m + n) - d - (b - d))) q) else ring_0(r:R ring)`]ring_sum_numseg_le_expand THEN specialize[`r:R ring`;`d:num`;`n:num`;`\a. ring_mul(r:R ring) (if a <= m then coeff (m - a) p else ring_0 r) (if d - a <= n then coeff (n - (d - a)) q else ring_0 r)`]ring_sum_numseg_le_offset THEN num_linear_fact `d+n <= m+n+d:num` THEN specialize[`r:R ring`;`d+n:num`;`m+n+d:num`;`\b. if n <= b then ring_mul r (if b - n <= m then coeff (m - (b - n)) p else ring_0 r) (if d - (b - n) <= n then coeff (n - (d - (b - n))) q else ring_0 r) else ring_0(r:R ring)`]ring_sum_numseg_le_expand THEN simp[] THEN sufficesby RING_SUM_EQ THEN intro THEN case `a < n:num` THENL [ simp[ARITH_RULE `a < n:num ==> ~(n <= a)`] THEN case `d <= a:num` THENL [ simp[ARITH_RULE `a < n:num ==> a <= m+n`] THEN num_linear_fact `a < n:num ==> d <= a ==> ~((m+n)-d-(a-d) <= m)` THEN have `coeff ((m+n)-d-(a-d)) p = ring_0(r:R ring)` [] THEN simp[] THEN qed[RING_MUL_LZERO;coeff_poly_in_ring] ; pass ] THEN simp[] THEN qed[COND_ID] ; pass ] THEN simp[ARITH_RULE `~(a < n:num) ==> n <= a`] THEN case `a <= m+n:num` THENL [ simp[ARITH_RULE `a <= m+n:num ==> a-n <= m`] THEN case `d <= a:num` THENL [ simp[ARITH_RULE `~(a < n) ==> d <= a:num ==> d-(a-n) <= n`] THEN simp[ARITH_RULE `~(a < n) ==> d <= a:num ==> (m+n)-d-(a-d) = m-(a-n)`] THEN simp[ARITH_RULE `a <= m+n ==> ~(a < n) ==> d <= a:num ==> (m+n)-d-(m-(a-n)) = a-d`] THEN case `a <= d+n:num` THENL [ simp[ARITH_RULE `d <= a ==> ~(a < n) ==> a <= d+n ==> a-d = n-(d-(a-n)):num`] ; num_linear_fact `d <= a:num ==> ~(a <= d+n) ==> ~(a-d <= n)` THEN have `coeff (a-d) q = ring_0(r:R ring)` [] THEN simp[] THEN qed[RING_MUL_RZERO;coeff_poly_in_ring] ] ; case `a <= d+n:num` THENL [ simp[ARITH_RULE `~(d <= a) ==> ~(a < n) ==> ~(d-(a-n) <= n:num)`] THEN qed[RING_MUL_RZERO;coeff_poly_in_ring] ; simp[] ] ] ; simp[ARITH_RULE `~(a <= m+n) ==> ~(a-n <= m:num)`] THEN qed[RING_MUL_LZERO;coeff_poly_in_ring] ] );; let x_truncreverse_product = prove(` !(r:R ring) n p S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> poly_deg r (p s) <= (n s)) ==> x_truncreverse r (nsum S n) (poly_product r S p) = poly_product r S (\s. x_truncreverse r (n s) (p s)) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;NSUM_CLAUSES] THEN qed[x_truncreverse_0_poly_const;poly_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_polynomial(r:R ring) (p(x:X):(1->num)->R)` [] THEN have `ring_powerseries(r:R ring) (p(x:X):(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries(r:R ring) (x_truncreverse r (n x) (p(x:X):(1->num)->R))` [x_truncreverse_series] THEN simp[poly_product_insert;NSUM_CLAUSES] THEN have `ring_polynomial(r:R ring) (poly_product r S (p:X->(1->num)->R))` [poly_product_poly] THEN have `poly_deg(r:R ring) (p(x:X):(1->num)->R) <= n x` [] THEN have `poly_deg(r:R ring) (poly_product r (S:X->bool) p) <= nsum S n` [poly_deg_product_le] THEN specialize[`r:R ring`;`n(x:X):num`;`p(x:X):(1->num)->R`;`nsum (S:X->bool) n`;`poly_product(r:R ring) (S:X->bool) p`]x_truncreverse_mul THEN qed[] ] );; let x_truncreverse_pow = prove(` !(r:R ring) p d n. ring_polynomial r p ==> poly_deg r p <= d ==> x_truncreverse r (n*d) (poly_pow r p n) = poly_pow r (x_truncreverse r d p) n `, intro THEN have `ring_powerseries r (p:(1->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (x_truncreverse r d (p:(1->num)->R))` [x_truncreverse_series] THEN simp[poly_pow_is_product] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `CARD (1..n) = n` [CARD_NUMSEG;ARITH_RULE `(n+1)-1 = n`] THEN have `nsum (1..n) (\i. d) = CARD(1..n) * d` [NSUM_CONST] THEN have `nsum (1..n) (\i. d) = n * d` [] THEN have `!i. i IN 1..n ==> ring_polynomial r (p:(1->num)->R)` [] THEN have `!i. i IN 1..n ==> poly_deg r (p:(1->num)->R) <= d` [] THEN specialize[`r:R ring`;`\i:num. d:num`;`\i:num. p:(1->num)->R`;`1..n`]x_truncreverse_product THEN qed[] );; let deg_monic_vanishing_at = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_deg r (monic_vanishing_at r S c) = if ring_1 r = ring_0 r then 0 else CARD S `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[monic_vanishing_at_empty;POLY_DEG_1;card_empty] THEN ARITH_TAC ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN S ==> c s IN ring_carrier(r:R ring)` [] THEN simp[monic_vanishing_at_insert;card_insert] THEN have `c(x:X) IN ring_carrier(r:R ring)` [] THEN have `ring_polynomial r (x_minus_const(r:R ring) (c(x:X)))` [x_minus_const_poly] THEN have `ring_polynomial r (monic_vanishing_at(r:R ring) S (c:X->R))` [monic_vanishing_at_poly] THEN have `monic r (x_minus_const(r:R ring) (c(x:X)))` [monic_x_minus_const] THEN have `monic r (monic_vanishing_at(r:R ring) S (c:X->R))` [monic_vanishing_at_monic] THEN have `poly_deg(r:R ring) (poly_mul r (x_minus_const r (c(x:X))) (monic_vanishing_at r S c)) = poly_deg r (x_minus_const r (c x)) + poly_deg r (monic_vanishing_at r S c)` [deg_monic_poly_mul] THEN proven_if `ring_1 r = ring_0(r:R ring)` [deg_zero_ring;ring_polynomial;RING_POLYNOMIAL_MUL] THEN have `poly_deg(r:R ring) (x_minus_const r (c(x:X))) = 1` [deg_x_minus_const] THEN have `poly_deg(r:R ring) (monic_vanishing_at r S (c:X->R)) = CARD S` [] THEN qed[ARITH_RULE `C + 1 = 1 + C`] ] );; let deg_monic_vanishing_at_le = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_deg r (monic_vanishing_at r S c) <= CARD S `, simp[deg_monic_vanishing_at] THEN ARITH_TAC );; let topcoeff_monic_vanishing_at = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> coeff (CARD S) (monic_vanishing_at r S c) = ring_1 r `, intro THEN case `ring_1 r = ring_0(r:R ring)` THENL [ have `trivial_ring(r:R ring)` [TRIVIAL_RING_10] THEN have `ring_carrier(r:R ring) = {ring_0 r}` [trivial_ring] THEN have `ring_polynomial r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_poly] THEN have `coeff (CARD S) (monic_vanishing_at r S (c:X->R)) IN ring_carrier r` [coeff_poly_in_ring] THEN qed[IN_SING] ; pass ] THEN have `monic r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_monic] THEN have `CARD S = poly_deg(r:R ring) (monic_vanishing_at r S (c:X->R))` [deg_monic_vanishing_at] THEN qed[monic] );; (* could use HAS_SIZE, RESTRICTED_POWERSET, etc. *) let coeff_monic_vanishing_at_lemma = prove(` !(r:R ring) c:X->R S t n. FINITE S ==> ~(t IN S) ==> (!s:X. s IN t INSERT S ==> c s IN ring_carrier r) ==> ~(n = 0) ==> n <= CARD S ==> ring_add r ( ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U c) ) ( ring_mul r (c t) ( ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_product r U c) ) ) = ring_sum r {U | U SUBSET t INSERT S /\ CARD U = n} (\U. ring_product r U c) `, intro THEN have `c(t:X):R IN ring_carrier r` [IN_INSERT] THEN specialize[`S:X->bool`;`n-1`]finite_subsets_card THEN specialize[`S:X->bool`;`n:num`]finite_subsets_card THEN subgoal `FINITE {U:X->bool | U SUBSET t INSERT S /\ t IN U /\ CARD U = n}` THENL [ set_fact `{U:X->bool | U SUBSET t INSERT S /\ t IN U /\ CARD U = n} SUBSET {U:X->bool | U SUBSET t INSERT S}` THEN have `FINITE ((t:X) INSERT S)` [FINITE_INSERT] THEN specialize[`(t:X) INSERT S`]FINITE_POWERSET THEN qed[FINITE_SUBSET] ; pass ] THEN subgoal `DISJOINT {U:X->bool | U SUBSET S /\ CARD U = n} {U | U SUBSET t INSERT S /\ t IN U /\ CARD U = n}` THENL [ rw[DISJOINT;INTER;EXTENSION;IN_ELIM_THM;EMPTY] THEN qed[SUBSET] ; pass ] THEN specialize[ `r:R ring`; `\U. ring_product r U (c:X->R)`; `{U:X->bool | U SUBSET S /\ CARD U = n}`; `{U:X->bool | U SUBSET t INSERT S /\ t IN U /\ CARD U = n}` ]RING_SUM_UNION THEN subgoal `{U:X->bool | U SUBSET t INSERT S /\ CARD U = n} = {U | U SUBSET S /\ CARD U = n} UNION {U | U SUBSET t INSERT S /\ t IN U /\ CARD U = n}` THENL [ rw[UNION;EXTENSION;IN_ELIM_THM] THEN SET_TAC[] ; pass ] THEN simp[] THEN subgoal `ring_mul r (c(t:X):R) (ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_product r U c)) = ring_sum r {U | U SUBSET t INSERT S /\ t IN U /\ CARD U = n} (\U. ring_product r U c)` THENL [ subgoal `ring_sum r {U | U SUBSET t INSERT S /\ t IN U /\ CARD U = n} (\U. ring_product r U (c:X->R)) = ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_mul r (c t) (ring_product r U c))` THENL [ subgoal `!A B. A IN {U:X->bool | U SUBSET S /\ CARD U = n - 1} ==> B IN {U | U SUBSET S /\ CARD U = n - 1} ==> t INSERT A = t INSERT B ==> A = B` THENL [ rw[IN_ELIM_THM;EXTENSION;SUBSET;IN_INSERT] THEN qed[] ; pass ] THEN subgoal `{U | U SUBSET t INSERT S /\ t IN U /\ CARD U = n} = IMAGE (\A. t INSERT A) {U:X->bool | U SUBSET S /\ CARD U = n - 1}` THENL [ rw[IMAGE;EXTENSION;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN witness `x DELETE (t:X)` THEN intro THENL [ qed[SUBSET_INSERT_DELETE] ; qed[CARD_DELETE;FINITE_INSERT;FINITE_SUBSET] ; qed[IN_INSERT;IN_DELETE] ] ; intro THENL [ qed[SUBSET;IN_INSERT] ; qed[IN_INSERT] ; have `FINITE(x':X->bool)` [FINITE_SUBSET] THEN have `~((t:X) IN x')` [SUBSET] THEN have `CARD((t:X) INSERT x') = (n-1)+1` [card_insert] THEN num_linear_fact `~(n = 0) ==> n = (n-1)+1` THEN have `CARD((t:X) INSERT x') = n` [] THEN qed[EXTENSION] ] ] ; pass ] THEN specialize[ `r:R ring`; `\A:X->bool. t INSERT A`; `\U. ring_product r U (c:X->R)`; `{U:X->bool | U SUBSET S /\ CARD U = n - 1}` ]RING_SUM_IMAGE THEN subgoal `ring_sum r {U | U SUBSET S /\ CARD U = n - 1} ((\U. ring_product r U (c:X->R)) o (\A. t INSERT A)) = ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_mul r (c t) (ring_product r U c))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[o_THM] THEN set_fact `a IN {U | U SUBSET S /\ CARD U = n - 1} ==> (a:X->bool) SUBSET S` THEN have `FINITE(a:X->bool)` [FINITE_SUBSET] THEN have `~((t:X) IN a)` [SUBSET] THEN specialize[ `r:R ring`; `t:X`; `c:X->R`; `a:X->bool` ](CONJUNCT2 RING_PRODUCT_CLAUSES) THEN qed[] ; pass ] THEN qed[] ; pass ] THEN specialize_assuming[ `r:R ring`; `\U. ring_product r U (c:X->R)`; `c(t:X):R`; `{U:X->bool | U SUBSET S /\ CARD U = n - 1}` ]RING_SUM_LMUL THEN have `(!a. a IN {U | U SUBSET S /\ CARD U = n - 1} ==> ring_product r a (c:X->R) IN ring_carrier r)` [RING_PRODUCT] THEN qed[] ; pass ] THEN qed[] );; let coeff_monic_vanishing_at_lemma2 = prove(` !(r:R ring) c:X->R S t n. FINITE S ==> ~(t IN S) ==> (!s:X. s IN t INSERT S ==> c s IN ring_carrier r) ==> ~(n = 0) ==> n <= CARD S ==> ring_sub r ( ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) n ) ( ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U c) ) ) ( ring_mul r (c t) ( ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) (n - 1) ) ( ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_product r U c) ) ) ) = ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) n ) ( ring_sum r {U | U SUBSET t INSERT S /\ CARD U = n} (\U. ring_product r U c) ) `, intro THEN simp[GSYM coeff_monic_vanishing_at_lemma] THEN num_linear_fact `~(n = 0) ==> (n-1)+1 = n` THEN have `ring_pow(r:R ring) (ring_neg r (ring_1 r)) n = ring_neg r (ring_pow r (ring_neg r (ring_1 r)) (n-1))` [ring_pow_neg_1_plus1] THEN simp[] THEN specialize_assuming[ `r:R ring`; `c(t:X):R`; `ring_pow(r:R ring) (ring_neg r (ring_1 r)) (n - 1)`; `ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))`; `ring_sum r {U | U SUBSET S /\ CARD U = n - 1} (\U. ring_product r U (c:X->R))` ](GENL[ `r:R ring`;`C:R`;`M:R`;`N:R`;`N1:R` ](RING_RULE ` ring_sub(r:R ring) (ring_mul r (ring_neg r (M)) (N)) (ring_mul r (C) (ring_mul r (M) (N1))) = ring_mul r (ring_neg r (M)) (ring_add r (N) (ring_mul r (C) (N1))) `)) THEN qed[RING_1;RING_NEG;RING_POW;RING_SUM;IN_INSERT] );; let coeff_monic_vanishing_at = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> !n. n <= CARD S ==> coeff (CARD S - n) (monic_vanishing_at r S c) = ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) n ) ( ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U c) ) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ have `CARD({}:X->bool) = 0` [card_empty] THEN have `n = 0` [ARITH_RULE `n <= 0 ==> n = 0`] THEN simp[ARITH_RULE `0 - 0 = 0`;monic_vanishing_at;poly_product_empty] THEN have `{U | U SUBSET {} /\ CARD U = 0} = {{}:X->bool}` [subsets_full_card_empty] THEN simp[RING_SUM_SING;RING_PRODUCT_CLAUSES;RING_1;RING_POW_0;coeff_poly_1;RING_MUL_LID;RING_1] ; pass ] THEN case `n = 0` THENL [ simp[ARITH_RULE `x - 0 = x`] THEN simp[topcoeff_monic_vanishing_at;subsets_card_0;FINITE_INSERT] THEN rw[RING_SUM_SING;RING_PRODUCT;RING_PRODUCT_CLAUSES;RING_1] THEN RING_TAC ; pass ] THEN set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[card_insert;monic_vanishing_at_insert] THEN have `c(x:X) IN ring_carrier(r:R ring)` [] THEN have `ring_powerseries r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_series] THEN simp[coeff_x_minus_const_times] THEN case `n = CARD(S:X->bool) + 1` THENL [ simp[ARITH_RULE `x - x = 0`] THEN have `{U | U SUBSET S /\ CARD U = CARD S} = {S:X->bool}` [subsets_full_card] THEN have `coeff (CARD S - CARD S) (monic_vanishing_at r S (c:X->R)) = ring_mul r (ring_pow r (ring_neg r (ring_1 r)) (CARD S)) (ring_sum r {U | U SUBSET S /\ CARD U = CARD S} (\U. ring_product r U c))` [LE_REFL] THEN have `coeff 0 (monic_vanishing_at r S (c:X->R)) = ring_mul r (ring_pow r (ring_neg r (ring_1 r)) (CARD S)) (ring_sum r {S} (\U. ring_product r U c))` [ARITH_RULE `x - x = 0`] THEN have `ring_sum r {S} (\U. ring_product r U c) = ring_product r S (c:X->R)` [RING_PRODUCT;RING_SUM_SING] THEN have `coeff 0 (monic_vanishing_at r S (c:X->R)) = ring_mul r (ring_pow r (ring_neg r (ring_1 r)) (CARD S)) (ring_product r S c)` [] THEN have `{U | U SUBSET x INSERT S /\ CARD U = CARD(x INSERT S)} = {x INSERT S:X->bool}` [subsets_full_card;FINITE_INSERT] THEN have `{U | U SUBSET x INSERT S /\ CARD U = CARD S + 1} = {x INSERT S:X->bool}` [card_insert] THEN simp[RING_SUM_SING;RING_PRODUCT] THEN simp[RING_PRODUCT_CLAUSES] THEN rw[ring_pow_neg_1_plus1] THEN RING_TAC THEN qed[coeff_poly_in_ring;monic_vanishing_at_poly;RING_PRODUCT;RING_SUM] ; pass ] THEN have `n <= CARD(S:X->bool) + 1` [card_insert] THEN num_linear_fact `n <= CARD(S:X->bool) + 1 ==> ~(n = CARD S + 1) ==> n <= CARD S` THEN num_linear_fact `n <= CARD(S:X->bool) + 1 ==> ~(n = CARD S + 1) ==> ~((CARD S + 1) - n = 0)` THEN num_linear_fact `n <= CARD(S:X->bool) + 1 ==> ~(n = CARD S + 1) ==> (CARD S + 1) - n - 1 = CARD S - n` THEN num_linear_fact `~(n = 0) ==> n <= CARD(S:X->bool) + 1 ==> ~(n = CARD S + 1) ==> (CARD S + 1) - n = CARD S - (n - 1)` THEN num_linear_fact `n <= CARD(S:X->bool) ==> n-1 <= CARD S` THEN simp[] THEN qed[coeff_monic_vanishing_at_lemma2] );; let x_truncreverse_monic_vanishing_at = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> x_truncreverse r (CARD S) (monic_vanishing_at r S c) = poly_product r S (\s. one_minus_constx r (c s)) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[monic_vanishing_at_empty;poly_product_empty;card_empty;poly_1;x_truncreverse_0_poly_const] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN S ==> c s IN ring_carrier(r:R ring)` [] THEN simp[monic_vanishing_at_insert;poly_product_insert;card_insert] THEN have `c(x:X) IN ring_carrier(r:R ring)` [] THEN have `ring_polynomial r (x_minus_const(r:R ring) (c(x:X)))` [x_minus_const_poly] THEN have `ring_powerseries r (one_minus_constx(r:R ring) (c(x:X)))` [one_minus_constx_series] THEN have `ring_polynomial r (monic_vanishing_at(r:R ring) S (c:X->R))` [monic_vanishing_at_poly] THEN have `poly_deg(r:R ring) (x_minus_const r (c(x:X))) <= 1` [deg_x_minus_const_le] THEN have `poly_deg(r:R ring) (monic_vanishing_at r S (c:X->R)) <= CARD S` [deg_monic_vanishing_at_le] THEN num_linear_fact `CARD (S:X->bool) + 1 = 1 + CARD S` THEN simp[x_truncreverse_mul] THEN rw[x_truncreverse_x_minus_const] ] );; let x_truncreverse_derivative_monic_vanishing_at = prove(` !(r:R ring) S c:X->R. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> x_truncreverse r (CARD S - 1) (x_derivative r (monic_vanishing_at r S c)) = poly_sum r S (\s. poly_product r (S DELETE s) (\s. one_minus_constx r (c s))) `, intro THEN simp[x_derivative_monic_vanishing_at] THEN have `!s:X. s IN S ==> FINITE(S DELETE s)` [FINITE_DELETE] THEN set_fact `!s:X. !t:X. t IN (S DELETE s) ==> t IN S` THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (monic_vanishing_at r (S DELETE s) c)` [monic_vanishing_at_series] THEN simp[x_truncreverse_poly_sum] THEN sufficesby poly_sum_eq THEN intro THEN simp[] THEN have `CARD(S DELETE (s:X)) = CARD S - 1` [CARD_DELETE] THEN have `!t:X. t IN (S DELETE s) ==> c t IN ring_carrier(r:R ring)` [IN_DELETE] THEN qed[x_truncreverse_monic_vanishing_at] );; let x_truncreverse_subring = prove(` !(r:R ring) G n (p:(1->num)->R). x_truncreverse (subring_generated r G) n p = x_truncreverse r n p `, rw[x_truncreverse;SUBRING_GENERATED] );; let botcoeff1_if_x_truncreverse_monic = prove(` !(r:R ring) p. monic r p ==> coeff 0 (x_truncreverse r (poly_deg r p) p) = ring_1 r `, rw[monic;coeff_x_truncreverse;ARITH_RULE `0 <= d`;ARITH_RULE `d - 0 = d`] );; let nonzero_if_x_truncreverse_monic = prove(` !(r:R ring) p. ~(ring_1 r = ring_0 r) ==> monic r p ==> ~(x_truncreverse r (poly_deg r p) p = poly_0 r) `, qed[botcoeff1_if_x_truncreverse_monic;coeff_poly_0] );; (* ===== newton identities *) (* (prod_i x_i) sum_i 1/x_i = sum_i prod_(j!=i) x_i *) let product_times_sum_reciprocals = prove(` !(r:R ring) (f:X->R) (g:X->R) S. FINITE S ==> (!s. s IN S ==> f s IN ring_carrier r) ==> (!s. s IN S ==> g s IN ring_carrier r) ==> (!s. s IN S ==> ring_mul r (f s) (g s) = ring_1 r) ==> ring_mul r (ring_product r S (\s. f s)) (ring_sum r S (\s. g s)) = ring_sum r S (\t. ring_product r (S DELETE t) (\s. f s)) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;RING_SUM_CLAUSES] THEN qed[RING_0;RING_MUL_LID] ; pass ] THEN set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `!s:X. s IN S ==> f s IN ring_carrier(r:R ring)` [] THEN have `!s:X. s IN S ==> g s IN ring_carrier(r:R ring)` [] THEN set_fact `~(x IN S) ==> ((x:X) INSERT S) DELETE x = S` THEN have `((x:X) INSERT S) DELETE x = S` [] THEN set_fact `!s:X. s IN S ==> ~(x IN S) ==> (x INSERT S) DELETE s = x INSERT (S DELETE s)` THEN have `f(x:X) IN ring_carrier(r:R ring)` [] THEN have `g(x:X) IN ring_carrier(r:R ring)` [] THEN have `ring_product r S (\s:X. f s) IN ring_carrier(r:R ring)` [RING_PRODUCT] THEN have `ring_sum r S (\s:X. g s) IN ring_carrier(r:R ring)` [RING_SUM] THEN simp[RING_PRODUCT_CLAUSES;RING_SUM_CLAUSES] THEN simp[RING_RULE `ring_mul(r:R ring) F0 G0 = ring_1 r ==> ring_mul r (ring_mul r F0 F1) (ring_add r G0 G1) = ring_add r F1 (ring_mul r F0 (ring_mul r F1 G1))`] THEN have `!a:X. a IN S ==> ring_product(r:R ring) (S DELETE a) (\s. f s) IN ring_carrier r` [RING_PRODUCT] THEN specialize[`r:R ring`;`\t:X. ring_product(r:R ring) (S DELETE t) (\s. f s)`;`f(x:X):R`;`S:X->bool`](GSYM RING_SUM_LMUL) THEN have `ring_sum(r:R ring) S (\x':X. ring_mul r (f x) (ring_product r (S DELETE x') (\s. f s))) IN ring_carrier r` [RING_SUM] THEN have `ring_sum(r:R ring) S (\a:X. ring_product r (x INSERT (S DELETE a)) (\s. f s)) IN ring_carrier r` [RING_SUM] THEN simp[RING_ADD_LCANCEL] THEN sufficesby RING_SUM_EQ THEN intro THEN simp[] THEN have `FINITE (S DELETE a:X)` [FINITE_DELETE] THEN set_fact `!s:X. s IN S ==> ~(x IN S) ==> ~(x IN S DELETE s)` THEN simp[CONJUNCT2 RING_PRODUCT_CLAUSES] );; let newton_identities_natural = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r (poly_product r S (\s. one_minus_constx r (c s))) (poly_sum r S (\s. infinite_geometric_series r (c s))) = poly_sum r S (\t. poly_product r (S DELETE t) (\s. one_minus_constx r (c s))) `, rw[x_series_use] THEN rw[poly_product;poly_sum] THEN intro THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_series] THEN have `!s:X. s IN S ==> one_minus_constx r (c s) IN ring_carrier(x_series(r:R ring))` [x_series_carrier;SET_RULE `ring_powerseries(r:R ring) (p:(1->num)->R) ==> p IN {q | ring_powerseries r q}`] THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (infinite_geometric_series r (c s))` [infinite_geometric_series_powerseries] THEN have `!s:X. s IN S ==> infinite_geometric_series r (c s) IN ring_carrier(x_series(r:R ring))` [x_series_carrier;SET_RULE `ring_powerseries(r:R ring) (p:(1->num)->R) ==> p IN {q | ring_powerseries r q}`] THEN have `!s:X. s IN S ==> ring_mul (x_series(r:R ring)) (one_minus_constx r (c s)) (infinite_geometric_series r (c s)) = ring_1 (x_series r)` [infinite_geometric_series_inverse;x_series_use] THEN specialize[`x_series(r:R ring)`;`\s:X. one_minus_constx(r:R ring) (c s)`;`\s:X. infinite_geometric_series(r:R ring) (c s)`;`S:X->bool`]product_times_sum_reciprocals THEN qed[] );; let newton_identities_monic_vanishing_at = prove(` !(r:R ring) c:X->R S. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r (x_truncreverse r (CARD S) (monic_vanishing_at r S c)) (poly_sum r S (\s. infinite_geometric_series r (c s))) = x_truncreverse r (CARD S - 1) (x_derivative r (monic_vanishing_at r S c)) `, simp[x_truncreverse_monic_vanishing_at] THEN simp[x_truncreverse_derivative_monic_vanishing_at] THEN qed[newton_identities_natural] );; (* maybe the most canonical form *) let newton_identities = prove(` !(r:R ring) c:X->R S d. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_sum r (0..d) (\a. ring_mul r (if a <= CARD S then coeff(CARD S-a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d-a))) ) = if d <= CARD S-1 then coeff(CARD S-1-d) (x_derivative r (monic_vanishing_at r S c)) else ring_0 r `, intro THEN specialize[`r:R ring`;`c:X->R`;`S:X->bool`]newton_identities_monic_vanishing_at THEN specialize[`r:R ring`;`d:num`;`x_truncreverse(r:R ring) (CARD S) (monic_vanishing_at r S (c:X->R))`;`poly_sum(r:R ring) S (\s:X. infinite_geometric_series r (c s))`](GSYM coeff_poly_mul_oneindex) THEN specialize[`r:R ring`;`CARD(S:X->bool)`;`monic_vanishing_at(r:R ring) S (c:X->R)`](GSYM coeff_x_truncreverse) THEN recall(GSYM coeff_infinite_geometric_series) THEN subgoal `!a:num. ring_sum(r:R ring) S (\s:X. coeff (d - a) (infinite_geometric_series r (c s))) = coeff (d - a) (poly_sum r S (\s. infinite_geometric_series r (c s)))` THENL [ intro THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (infinite_geometric_series r (c s))` [infinite_geometric_series_powerseries] THEN specialize[`r:R ring`;`\s:X. infinite_geometric_series (r:R ring) (c s)`;`d-a:num`;`S:X->bool`]coeff_poly_sum THEN qed[] ; pass ] THEN simp[] THEN qed[coeff_x_truncreverse] );; (* pulling out the leading 1 *) let newton_identities_recurrence = prove(` !(r:R ring) c:X->R S d. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_sum r S (\s. ring_pow r (c s) d) = ring_sub r (if d <= CARD S-1 then coeff(CARD S-1-d) (x_derivative r (monic_vanishing_at r S c)) else ring_0 r) (ring_sum r (1..d) (\a. ring_mul r (if a <= CARD S then coeff(CARD S-a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d-a))) ) ) `, intro THEN specialize[`r:R ring`;`c:X->R`;`S:X->bool`;`d:num`](GSYM newton_identities) THEN simp[] THEN num_linear_fact `0 <= d` THEN have `FINITE (0..d)` [FINITE_NUMSEG] THEN have `0 IN 0..d` [IN_NUMSEG_0] THEN have `ring_mul(r:R ring) (if 0 <= CARD S then coeff (CARD S - 0) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s:X. ring_pow r (c s) (d - 0))) IN ring_carrier r` [RING_MUL;RING_SUM;coeff_poly_in_ring;monic_vanishing_at_poly;RING_0] THEN specialize[`r:R ring`;`\a. ring_mul(r:R ring) (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s:X. ring_pow r (c s) (d - a)))`;`0..d`;`0:num`]RING_SUM_DELETE THEN subgoal `1..d = (0..d) DELETE 0` THENL [ rw[EXTENSION;IN_DELETE;IN_NUMSEG] THEN ARITH_TAC ; pass ] THEN simp[] THEN have `ring_sum r (0..d) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s:X. ring_pow r (c s) (d - a)))) IN ring_carrier(r:R ring)` [RING_SUM] THEN have `ring_mul r (coeff (CARD S - 0) (monic_vanishing_at r S c)) (ring_sum r S (\s:X. ring_pow r (c s) (d - 0))) IN ring_carrier(r:R ring)` [RING_SUM;RING_MUL;monic_vanishing_at_poly;coeff_poly_in_ring] THEN simp[RING_RULE `ring_sub(r:R ring) x (ring_sub r x y) = y`] THEN simp[ARITH_RULE `x - 0 = x`] THEN simp[topcoeff_monic_vanishing_at] THEN qed[RING_MUL_LID;RING_SUM] );; let powersums_subring_if_poly_subring = prove(` !(r:R ring) G S c:X->R d. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!n. coeff n (monic_vanishing_at r S c) IN ring_carrier(subring_generated r G)) ==> ring_sum r S (\s. ring_pow r (c s) d) IN ring_carrier(subring_generated r G) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby num_WF THEN intro THEN specialize[`r:R ring`;`c:X->R`;`S:X->bool`;`d:num`]newton_identities_recurrence THEN have `ring_0(r:R ring) IN ring_carrier(subring_generated r G)` [RING_0;SUBRING_GENERATED] THEN specialize[`r:R ring`;`monic_vanishing_at r S (c:X->R)`;`CARD(S:X->bool) - 1 - d`]coeff_x_derivative THEN have `ring_of_num(r:R ring) (CARD(S:X->bool)-1-d+1) IN ring_carrier(subring_generated r G)` [RING_OF_NUM;RING_OF_NUM_SUBRING_GENERATED] THEN have `coeff (CARD S-1-d+1) (monic_vanishing_at r S (c:X->R)) IN ring_carrier(subring_generated r G)` [] THEN specialize[`subring_generated(r:R ring) G`;`ring_of_num(r:R ring) (CARD(S:X->bool) - 1 - d + 1)`;`coeff (CARD S - 1 - d + 1) (monic_vanishing_at r S (c:X->R))`]RING_MUL THEN have `ring_mul (subring_generated r G) (ring_of_num r (CARD S - 1 - d + 1)) (coeff (CARD S - 1 - d + 1) (monic_vanishing_at r S c)) = ring_mul r (ring_of_num r (CARD S - 1 - d + 1)) (coeff (CARD S - 1 - d + 1) (monic_vanishing_at r S (c:X->R)))` [SUBRING_GENERATED] THEN have `ring_mul r (ring_of_num r (CARD S - 1 - d + 1)) (coeff (CARD S - 1 - d + 1) (monic_vanishing_at r S (c:X->R))) IN ring_carrier(subring_generated r G)` [] THEN have `coeff(CARD S - 1 - d) (x_derivative r (monic_vanishing_at r S (c:X->R))) IN ring_carrier(subring_generated r G)` [RING_MUL] THEN subgoal `!a. a IN (1..d) ==> ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s:X. ring_pow r (c s) (d - a))) IN ring_carrier(subring_generated(r:R ring) G)` THENL [ intro THEN have `!m. m < d ==> ring_sum r S (\s:X. ring_pow r (c s:R) m) IN ring_carrier (subring_generated r G)` [] THEN have `1 <= a` [IN_NUMSEG] THEN have `a <= d:num` [IN_NUMSEG] THEN num_linear_fact `1 <= a /\ a <= d ==> d-a < d` THEN have `ring_sum r S (\s:X. ring_pow r (c s:R) (d-a)) IN ring_carrier (subring_generated r G)` [] THEN have `(if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) IN ring_carrier(subring_generated r G)` [] THEN specialize[`subring_generated(r:R ring) G`;`(if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r)`;`ring_sum r S (\s:X. ring_pow r (c s:R) (d - a))`]RING_MUL THEN qed[SUBRING_GENERATED] ; pass ] THEN specialize[`r:R ring`;`G:R->bool`;`1..d`;`\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d - a)))`](GSYM ring_sum_subring_generated_v2) THEN specialize[`subring_generated(r:R ring) G`;`1..d`;`\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d - a)))`]RING_SUM THEN have `(if d <= CARD S - 1 then coeff (CARD S - 1 - d) (x_derivative r (monic_vanishing_at r S (c:X->R))) else ring_0 r) IN ring_carrier(subring_generated r G)` [] THEN specialize[`subring_generated(r:R ring) G`;`if d <= CARD S - 1 then coeff (CARD S - 1 - d) (x_derivative r (monic_vanishing_at r S (c:X->R))) else ring_0 r`;`ring_sum (subring_generated r G) (1..d) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d - a))))`]RING_SUB THEN have `ring_sub r (if d <= CARD S - 1 then coeff (CARD S - 1 - d) (x_derivative r (monic_vanishing_at r S (c:X->R))) else ring_0 r) (ring_sum (subring_generated r G) (1..d) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d - a))))) IN ring_carrier (subring_generated r G)` [RING_SUB_SUBRING_GENERATED] THEN have `ring_sub r (if d <= CARD S - 1 then coeff (CARD S - 1 - d) (x_derivative r (monic_vanishing_at r S (c:X->R))) else ring_0 r) (ring_sum r (1..d) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (d - a))))) IN ring_carrier (subring_generated r G)` [] THEN qed[] );; (* the D=1 case boils down to powersums_subring_if_poly_subring_denominators *) (* but giving shorter proof of that case feels better *) let powersums_subring_if_poly_subring_denominators_waterfall = prove(` !(r:R ring) G S c:X->R D. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> D IN ring_carrier r ==> (!i. ring_mul r (ring_pow r D i) (coeff (CARD S - i) (monic_vanishing_at r S c)) IN ring_carrier(subring_generated r G)) ==> !j. ring_mul r (ring_pow r D j) (ring_sum r S (\s. ring_pow r (c s) j)) IN ring_carrier(subring_generated r G) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN sufficesby num_WF THEN intro THEN case `j = 0` THENL [ simp[RING_POW_0;ring_sum_1] THEN simp[RING_MUL_LID;RING_OF_NUM] THEN qed[RING_OF_NUM_SUBRING_GENERATED;RING_OF_NUM] ; pass ] THEN labelspecialize "newton" [`r:R ring`;`c:X->R`;`S:X->bool`;`j:num`] newton_identities_recurrence THEN (ASM once_rw)[] THEN removelabeled "newton" THEN rw[coeff_x_derivative] THEN have `coeff (CARD S-1-j+1) (monic_vanishing_at r S (c:X->R)) IN ring_carrier r` [coeff_poly_in_ring;monic_vanishing_at_poly] THEN have `ring_pow r D j IN ring_carrier(r:R ring)` [RING_POW] THEN have `(if j <= CARD S - 1 then ring_mul r (ring_of_num r (CARD S - 1 - j + 1)) (coeff (CARD S - 1 - j + 1) (monic_vanishing_at r S (c:X->R))) else ring_0 r) IN ring_carrier r` [RING_MUL;RING_OF_NUM;RING_0] THEN have `ring_sum r (1..j) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - a)))) IN ring_carrier r` [RING_SUM] THEN labelspecialize "distrib" [`r:R ring`;`ring_pow(r:R ring) D j`;`if j <= CARD S - 1 then ring_mul r (ring_of_num r (CARD S - 1 - j + 1)) (coeff (CARD S - 1 - j + 1) (monic_vanishing_at r S (c:X->R))) else ring_0 r`;`ring_sum r (1..j) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - a))))`]RING_SUB_LDISTRIB THEN simp[] THEN removelabeled "distrib" THEN have `FINITE (1..j)` [FINITE_NUMSEG] THEN have `!a. a IN (1..j) ==> ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - a))) IN ring_carrier r` [RING_MUL;coeff_poly_in_ring;monic_vanishing_at_poly;RING_SUM;RING_0] THEN labelspecialize "distrib2" [`r:R ring`;`\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - a)))`;`ring_pow(r:R ring) D j`;`1..j`](GSYM RING_SUM_LMUL) THEN simp[] THEN removelabeled "distrib2" THEN subgoal `ring_mul r (ring_pow r D j) (if j <= CARD S - 1 then ring_mul r (ring_of_num r (CARD S - 1 - j + 1)) (coeff (CARD S - 1 - j + 1) (monic_vanishing_at r S (c:X->R))) else ring_0 r) IN ring_carrier(subring_generated r G)` THENL [ case `j <= CARD(S:X->bool) - 1` THENL [ have `ring_of_num(r:R ring) (CARD(S:X->bool)-1-j+1) IN ring_carrier r` [RING_OF_NUM] THEN simp[RING_MUL_ASSOC] THEN simp[RING_MUL_SYM] THEN simp[GSYM RING_MUL_ASSOC] THEN have `ring_mul r (ring_pow r D j) (coeff (CARD S - j) (monic_vanishing_at r S (c:X->R))) IN ring_carrier (subring_generated r G)` [] THEN num_linear_fact `~(j = 0) ==> j <= CARD(S:X->bool) - 1 ==> CARD S - 1 - j + 1 = CARD S - j` THEN have `ring_mul r (ring_pow r D j) (coeff (CARD S - j) (monic_vanishing_at r S (c:X->R))) IN ring_carrier (subring_generated r G)` [] THEN have `ring_mul r (ring_pow r D j) (coeff (CARD S - 1 - j + 1) (monic_vanishing_at r S (c:X->R))) IN ring_carrier (subring_generated r G)` [] THEN have `ring_of_num(r:R ring) (CARD(S:X->bool) - 1 - j + 1) IN ring_carrier(subring_generated r G)` [RING_OF_NUM_SUBRING_GENERATED;RING_OF_NUM] THEN qed[SUBRING_GENERATED;RING_MUL] ; pass ] THEN simp[] THEN have `ring_mul(r:R ring) (ring_pow r D j) (ring_0 r) = ring_0 r` [RING_MUL_RZERO;RING_POW] THEN qed[RING_0;SUBRING_GENERATED] ; pass ] THEN subgoal `ring_sum r (1..j) (\x. ring_mul r (ring_pow r D j) (ring_mul r (if x <= CARD S then coeff (CARD S - x) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - x))))) IN ring_carrier(subring_generated r G)` THENL [ subgoal `!x. x IN 1..j ==> ring_mul r (ring_pow r D j) (ring_mul r (if x <= CARD S then coeff (CARD S - x) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - x)))) IN ring_carrier(subring_generated r G)` THENL [ intro THEN case `x <= CARD(S:X->bool)` THENL [ have `x <= j:num` [IN_NUMSEG] THEN num_linear_fact `x <= j:num ==> x + (j-x) = j` THEN have `ring_pow(r:R ring) D j = ring_mul r (ring_pow r D x) (ring_pow r D (j-x))` [RING_POW_ADD] THEN simp[] THEN have `ring_pow(r:R ring) D x IN ring_carrier r` [RING_POW] THEN have `ring_pow(r:R ring) D (j-x) IN ring_carrier r` [RING_POW] THEN have `ring_sum r S (\s:X. ring_pow r (c s:R) (j-x)) IN ring_carrier r` [RING_SUM] THEN have `coeff (CARD S-x) (monic_vanishing_at r S (c:X->R)) IN ring_carrier r` [coeff_poly_in_ring;monic_vanishing_at_poly] THEN specialize[`r:R ring` ;`coeff(CARD S-x) (monic_vanishing_at r S (c:X->R))` ;`ring_sum r S (\s:X. ring_pow r (c s:R) (j - x))` ;`ring_pow(r:R ring) D x` ;`ring_pow(r:R ring) D (j-x)`; ](GENL[`r:R ring`;`a:R`;`b:R`;`x:R`;`y:R`]( RING_RULE `ring_mul(r:R ring) (ring_mul r x y) (ring_mul r a b) = ring_mul r (ring_mul r x a) (ring_mul r y b)` )) THEN simp[] THEN have `!s:X. s IN S ==> ring_pow r (c s:R) (j-x) IN ring_carrier r` [RING_POW] THEN have `1 <= x` [IN_NUMSEG] THEN num_linear_fact `1 <= x ==> x <= j ==> j-x < j` THEN have `ring_mul r (ring_pow r D (j-x)) (ring_sum r S (\s:X. ring_pow r (c s:R) (j-x))) IN ring_carrier(subring_generated r G)` [] THEN have `ring_mul(subring_generated r G) (ring_mul r (ring_pow r D x) (coeff (CARD S - x) (monic_vanishing_at r S (c:X->R)))) (ring_mul r (ring_pow r D (j - x)) (ring_sum r S (\s. ring_pow r (c s) (j - x)))) IN ring_carrier (subring_generated r G)` [RING_MUL] THEN qed[SUBRING_GENERATED] ; pass ] THEN simp[RING_MUL_LZERO;RING_SUM;RING_MUL_RZERO;RING_POW] THEN qed[RING_0;SUBRING_GENERATED] ; pass ] THEN specialize[`r:R ring`;`G:R->bool`;`1..j`;`\x. ring_mul r (ring_pow r D j) (ring_mul r (if x <= CARD S then coeff (CARD S - x) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - x))))`](GSYM ring_sum_subring_generated_v2) THEN simp[] THEN qed[RING_SUM] ; pass ] THEN have `ring_sub(subring_generated r G) (ring_mul r (ring_pow r D j) (if j <= CARD S - 1 then ring_mul r (ring_of_num r (CARD S - 1 - j + 1)) (coeff (CARD S - 1 - j + 1) (monic_vanishing_at r S (c:X->R))) else ring_0 r)) (ring_sum r (1..j) (\x. ring_mul r (ring_pow r D j) (ring_mul r (if x <= CARD S then coeff (CARD S - x) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (j - x)))))) IN ring_carrier (subring_generated r G)` [RING_SUB] THEN qed[RING_SUB_SUBRING_GENERATED] );; let powersums_subring_if_poly_subring_denominators = prove(` !(r:R ring) G S c:X->R D j. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> D IN ring_carrier r ==> (!i. ring_mul r (ring_pow r D i) (coeff (CARD S - i) (monic_vanishing_at r S c)) IN ring_carrier(subring_generated r G)) ==> ring_mul r (ring_pow r D j) (ring_sum r S (\s. ring_pow r (c s) j)) IN ring_carrier(subring_generated r G) `, qed[powersums_subring_if_poly_subring_denominators_waterfall] );; (* ===== reverse Newton recurrence: coeffs from power sums *) let newton_identities_reverse = prove(` !(r:R ring) c:X->R S n. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> n <= CARD S ==> ring_mul r ( ring_neg r (ring_of_num r n) ) ( coeff(CARD S - n) (monic_vanishing_at r S c) ) = ring_sum r (1..n) (\a. ring_mul r ( coeff(CARD S - (n - a)) (monic_vanishing_at r S c) ) ( ring_sum r S (\s. ring_pow r (c s) a) ) ) `, intro THEN have `ring_polynomial r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_poly] THEN case `n = 0` THENL [ have `1..n = {}` [EMPTY_NUMSEG;ARITH_RULE `0 < 1`] THEN simp[RING_OF_NUM_0;RING_NEG_0;RING_SUM_CLAUSES;RING_MUL_LZERO;coeff_poly_in_ring] ; pass ] THEN specialize[`r:R ring`;`c:X->R`;`S:X->bool`;`n:num`]newton_identities THEN subgoal `ring_sum r (0..n) (\a. ring_mul r (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (n - a)))) = ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S c))` THENL [ subgoal `(if n <= CARD S - 1 then coeff (CARD S - 1 - n) (x_derivative r (monic_vanishing_at r S (c:X->R))) else ring_0 r) = ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S c))` THENL [ rw[coeff_x_derivative] THEN case `n = CARD(S:X->bool):num` THENL [ num_linear_fact `~(n = 0) ==> n = CARD(S:X->bool) ==> ~(n <= CARD S - 1)` THEN have `ring_polynomial r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_poly] THEN simp[ARITH_RULE `CARD(S:X->bool) - CARD S = 0`;RING_OF_NUM_0;RING_NEG_0;RING_SUM_CLAUSES;RING_MUL_LZERO;coeff_poly_in_ring] ; pass ] THEN num_linear_fact `~(n = 0) ==> ~(n = CARD(S:X->bool)) ==> n <= CARD S ==> n <= CARD S - 1` THEN num_linear_fact `~(n = 0) ==> ~(n = CARD(S:X->bool)) ==> n <= CARD S ==> CARD S - 1 - n + 1 = CARD S - n` THEN simp[] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_sum r (0..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) b))) = ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S c))` THENL [ specialize[ `r:R ring`; `n:num`; `\a. ring_mul(r:R ring) (if a <= CARD S then coeff (CARD S - a) (monic_vanishing_at r S c) else ring_0 r) (ring_sum r S (\s:X. ring_pow r (c s) (n - a)))` ]ring_sum_numseg_le_reflect THEN subgoal `ring_sum r (0..n) (\b. ring_mul r (if n - b <= CARD S then coeff (CARD S - (n - b)) (monic_vanishing_at r S (c:X->R)) else ring_0 r) (ring_sum r S (\s. ring_pow r (c s) (n - (n - b))))) = ring_sum r (0..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) (b))))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN subgoal `ring_sum r S (\s. ring_pow r (c(s:X):R) (n - (n - a))) = ring_sum r S (\s. ring_pow r (c s) a)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN have `a <= n:num` [IN_NUMSEG] THEN num_linear_fact `a <= n: num ==> n - (n - a) = a` THEN qed[] ; pass ] THEN num_linear_fact `n <= CARD S ==> n - a <= CARD(S:X->bool)` THEN qed[] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_add r (ring_mul r (coeff (CARD S - (n - 0)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) 0))) (ring_sum r (0 + 1..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S c)) (ring_sum r S (\s. ring_pow r (c s) b)))) = ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S c))` THENL [ num_linear_fact `0 <= n` THEN specialize[ `r:R ring`; `\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) b))`; `0`; `n:num`; ]RING_SUM_CLAUSES_LEFT THEN have `ring_mul r (coeff (CARD S - (n - 0)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) 0)) IN ring_carrier r` [RING_MUL;coeff_poly_in_ring;monic_vanishing_at_poly;RING_SUM] THEN qed[] ; pass ] THEN subgoal `ring_add r (ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))) (ring_of_num r (CARD S))) (ring_sum r (1..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S c)) (ring_sum r S (\s. ring_pow r (c s) b)))) = ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S c))` THENL [ subgoal `ring_add r (ring_mul r (coeff (CARD S - (n - 0)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) 0))) (ring_sum r (0 + 1..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S c)) (ring_sum r S (\s. ring_pow r (c s) b)))) = ring_add r (ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S c)) (ring_of_num r (CARD S))) (ring_sum r (1..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S c)) (ring_sum r S (\s. ring_pow r (c s) b))))` THENL [ rw[ARITH_RULE `0 + 1 = 1`] THEN rw[ARITH_RULE `n - 0 = n`] THEN rw[RING_POW_0] THEN simp[ring_sum_1] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_add r (ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))) (ring_of_num r (CARD S))) (ring_sum r (1..n) (\b. ring_mul r (coeff (CARD S - (n - b)) (monic_vanishing_at r S c)) (ring_sum r S (\s. ring_pow r (c s) b)))) = ring_add r (ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S c)) (ring_of_num r (CARD S))) (ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S c)))` THENL [ subgoal `ring_mul r (ring_of_num r (CARD S - n)) (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))) = ring_add r (ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S c)) (ring_of_num r (CARD S))) (ring_mul r (ring_neg r (ring_of_num r (n))) (coeff (CARD S - n) (monic_vanishing_at r S c)))` THENL [ num_linear_fact `n <= CARD(S:X->bool) ==> CARD S = (CARD S - n) + n` THEN have `ring_of_num(r:R ring) (CARD(S:X->bool)) = ring_add r (ring_of_num r (CARD S - n)) (ring_of_num r n)` [RING_OF_NUM_ADD] THEN simp[] THEN specialize_assuming[ `r:R ring`; `coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))`; `ring_of_num(r:R ring) (CARD(S:X->bool) - n)`; `ring_of_num(r:R ring) n`; ](GENL[ `r:R ring`;`C:R`;`Sn:R`;`n:R` ](RING_RULE ` ring_mul(r:R ring) Sn C = ring_add r (ring_mul r C (ring_add r Sn n)) (ring_mul r (ring_neg r n) C) `)) THEN qed[RING_OF_NUM;coeff_poly_in_ring;monic_vanishing_at_poly] ; pass ] THEN qed[] ; pass ] THEN specialize_assuming[ `r:R ring`; `ring_mul r (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))) (ring_of_num r (CARD S))`; `ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)))`; `ring_sum r (1..n) (\a. ring_mul r (coeff (CARD S - (n - a)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) a)))` ]RING_ADD_LCANCEL THEN qed[RING_OF_NUM;coeff_poly_in_ring;monic_vanishing_at_poly;RING_MUL;RING_SUM;RING_NEG] );; let coeff_poly_subring_if_powersums_subring_lemma = prove(` !(r:R ring) G S c:X->R. ring_hasQ(subring_generated r G) ==> FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!d. ring_sum r S (\s. ring_pow r (c s) d) IN ring_carrier(subring_generated r G)) ==> !n. n <= CARD S ==> coeff (CARD S - n) (monic_vanishing_at r S c) IN ring_carrier(subring_generated r G) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN sufficesby num_WF THEN intro THEN subgoal `ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))) IN ring_carrier (subring_generated r G)` THENL [ specialize[`r:R ring`;`c:X->R`;`S:X->bool`;`n:num`]newton_identities_reverse THEN subgoal `ring_sum r (1..n) (\a. ring_mul r (coeff (CARD S - (n - a)) (monic_vanishing_at r S (c:X->R))) (ring_sum r S (\s. ring_pow r (c s) a))) IN ring_carrier(subring_generated r G)` THENL [ sufficesby ring_sum_in_subring THEN intro THEN rw[BETA_THM] THEN sufficesby RING_MUL_IN_SUBRING_GENERATED THEN have `1 <= s` [IN_NUMSEG] THEN have `s <= n:num` [IN_NUMSEG] THEN num_linear_fact `1 <= s ==> s <= n ==> n - s < n` THEN num_linear_fact `n - s < n ==> n <= CARD(S:X->bool) ==> n - s <= CARD S` THEN qed[] ; pass ] THEN qed[] ; pass ] THEN case `n = 0` THENL [ have `monic r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_monic] THEN have `ring_char(r:R ring) = 0` [RING_CHAR_SUBRING_GENERATED;ring_hasQ] THEN have `CARD S = poly_deg r (monic_vanishing_at r S (c:X->R))` [deg_monic_vanishing_at;TRIVIAL_RING_10;RING_CHAR_EQ_1;ARITH_RULE `~(1 = 0)`] THEN num_linear_fact `n = 0 ==> CARD(S:X->bool) - n = CARD S` THEN have `coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)) = ring_1 r` [monic] THEN qed[SUBRING_GENERATED;RING_1] ; pass ] THEN have `ring_neg r (ring_of_num r n) = ring_neg(subring_generated(r:R ring) G) (ring_of_num(subring_generated r G) n)` [RING_OF_NUM_SUBRING_GENERATED;SUBRING_GENERATED] THEN have `ring_neg r (ring_of_num r n) IN ring_carrier(subring_generated(r:R ring) G)` [RING_OF_NUM;RING_NEG] THEN have `ring_unit(subring_generated(r:R ring) G) (ring_neg r (ring_of_num r n))` [ring_hasQ_neg] THEN choose `v:R` `v IN ring_carrier(subring_generated(r:R ring) G) /\ ring_mul(subring_generated r G) (ring_neg r (ring_of_num r n)) v = ring_1(subring_generated r G)` [ring_unit] THEN specialize_assuming[ `r:R ring`; `coeff (CARD S - n) (monic_vanishing_at r S (c:X->R))`; `ring_neg(r:R ring) (ring_of_num r n)`; `v:R` ](GENL[ `r:R ring`;`c:R`;`u:R`;`v:R` ](RING_RULE ` ring_mul(r:R ring) u v = ring_1 r ==> c = ring_mul r v (ring_mul r u c) `)) THEN have `coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)) = ring_mul r v (ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S c)))` [coeff_poly_in_ring;monic_vanishing_at_poly;RING_CARRIER_SUBRING_GENERATED_SUBSET;SUBSET;SUBRING_GENERATED] THEN have `ring_mul(subring_generated r G) v (ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)))) IN ring_carrier(subring_generated r G)` [RING_MUL] THEN have `ring_mul r v (ring_mul r (ring_neg r (ring_of_num r n)) (coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)))) IN ring_carrier(subring_generated r G)` [SUBRING_GENERATED] THEN qed[] );; let coeff_poly_subring_if_powersums_subring = prove(` !(r:R ring) G S c:X->R. ring_hasQ(subring_generated r G) ==> FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!d. ring_sum r S (\s. ring_pow r (c s) d) IN ring_carrier(subring_generated r G)) ==> !n. coeff n (monic_vanishing_at r S c) IN ring_carrier(subring_generated r G) `, intro THEN case `n <= CARD(S:X->bool)` THENL [ num_linear_fact `n <= CARD(S:X->bool) ==> CARD S - (CARD S - n) = n` THEN num_linear_fact `CARD S - n <= CARD(S:X->bool)` THEN specialize[ `r:R ring`; `G:R->bool`; `S:X->bool`; `c:X->R` ]coeff_poly_subring_if_powersums_subring_lemma THEN qed[] ; pass ] THEN proven_if `coeff n (monic_vanishing_at r S (c:X->R)) = ring_0 r` [RING_0;SUBRING_GENERATED] THEN have `ring_polynomial r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_poly] THEN have `poly_deg r (monic_vanishing_at r S (c:X->R)) <= CARD S` [deg_monic_vanishing_at_le] THEN qed[coeff_deg_le] );; let poly_subring_if_powersums_subring = prove(` !(r:R ring) G S c:X->R. ring_hasQ(subring_generated r G) ==> FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!d. ring_sum r S (\s. ring_pow r (c s) d) IN ring_carrier(subring_generated r G)) ==> ring_polynomial(subring_generated r G) (monic_vanishing_at r S c) `, intro THEN have `!n. coeff n (monic_vanishing_at r S (c:X->R)) IN ring_carrier(subring_generated r G)` [coeff_poly_subring_if_powersums_subring] THEN have `ring_polynomial r (monic_vanishing_at r S (c:X->R))` [monic_vanishing_at_poly] THEN qed[ring_polynomial_subring_if_coeffs] );; let symfun_subring_if_powersums_subring = prove(` !(r:R ring) G S c:X->R n. ring_hasQ(subring_generated r G) ==> FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!d. ring_sum r S (\s. ring_pow r (c s) d) IN ring_carrier(subring_generated r G)) ==> ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U c) IN ring_carrier(subring_generated r G) `, intro THEN case `n <= CARD(S:X->bool)` THENL [ have `coeff (CARD S - n) (monic_vanishing_at r S (c:X->R)) IN ring_carrier(subring_generated r G)` [coeff_poly_subring_if_powersums_subring] THEN have `ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))) IN ring_carrier(subring_generated r G)` [coeff_monic_vanishing_at] THEN specialize[`r:R ring`;`n:num`]ring_pow_neg_1_mul_refl THEN subgoal `ring_mul r (ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_pow r (ring_neg r (ring_1 r)) n)) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))) IN ring_carrier (subring_generated r G)` THENL [ have `ring_pow(subring_generated r G) (ring_neg r (ring_1 r)) n:R IN ring_carrier(subring_generated r G)` [RING_POW;SUBRING_GENERATED;RING_NEG;RING_1] THEN have `ring_pow r (ring_neg r (ring_1 r)) n:R IN ring_carrier(subring_generated r G)` [RING_POW_SUBRING_GENERATED] THEN have `ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R)))) IN ring_carrier (subring_generated r G)` [RING_MUL;SUBRING_GENERATED] THEN have `ring_mul r (ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_pow r (ring_neg r (ring_1 r)) n)) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))) = ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_mul r (ring_pow r (ring_neg r (ring_1 r)) n) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))))` [RING_MUL_ASSOC;RING_POW;RING_NEG;RING_1;RING_SUM] THEN qed[] ; pass ] THEN have `ring_mul r (ring_1 r) (ring_sum r {U | U SUBSET S /\ CARD U = n} (\U. ring_product r U (c:X->R))) IN ring_carrier (subring_generated r G)` [ring_pow_neg_1_mul_refl] THEN qed[RING_MUL_LID;RING_SUM] ; pass ] THEN simp[subsets_card_toobig] THEN rw[RING_SUM_CLAUSES] THEN qed[SUBRING_GENERATED;RING_0] );; (* ===== generalized Newton identities for 1/(1-cx)^(e+1) *) (* note: the e=0 case is given somewhat shorter proofs above *) let pow_newton_identities_natural = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r (poly_product r S (\s. poly_pow r (one_minus_constx r (c s)) (e+1) )) (poly_sum r S (\s. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r (c s) n) ) )) = poly_sum r S (\t. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1) ) ) `, rw[x_series_use] THEN rw[poly_product;poly_sum] THEN intro THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_series] THEN have `!s:X. s IN S ==> one_minus_constx r (c s) IN ring_carrier(x_series(r:R ring))` [x_series_carrier;SET_RULE `ring_powerseries(r:R ring) (p:(1->num)->R) ==> p IN {q | ring_powerseries r q}`] THEN have `!s:X. s IN S ==> poly_pow r (one_minus_constx r (c s)) (e+1) IN ring_carrier(x_series(r:R ring))` [x_series_use_pow;RING_POW] THEN def `q:X->(1->num)->R` `\s:X. series_from_coeffs (\n. ring_mul(r:R ring) (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n))` THEN have `!s:X n:num. s IN S ==> ring_mul(r:R ring) (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n) IN ring_carrier r` [RING_OF_NUM;RING_POW;RING_MUL] THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (q s:(1->num)->R)` [series_series_from_coeffs] THEN have `!s:X. s IN S ==> q s IN ring_carrier(x_series(r:R ring))` [x_series_carrier;SET_RULE `ring_powerseries(r:R ring) (p:(1->num)->R) ==> p IN {q | ring_powerseries r q}`] THEN subgoal `!s:X. s IN S ==> poly_mul(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e+1)) (q s) = poly_1 r` THENL [ intro THEN have `c(s:X) IN ring_carrier(r:R ring)` [] THEN specialize[`r:R ring`;`c(s:X):R`;`e:num`]pow_infinite_geometric_series_inverse THEN qed[] ; pass ] THEN have `!s:X. s IN S ==> ring_mul (x_series(r:R ring)) (poly_pow r (one_minus_constx r (c s)) (e+1)) (q s) = ring_1(x_series r)` [x_series_use] THEN specialize[`x_series(r:R ring)`;`\s:X. poly_pow r (one_minus_constx(r:R ring) (c s)) (e+1)`;`q:X->(1->num)->R`;`S:X->bool`](product_times_sum_reciprocals) THEN qed[RING_SUM_EQ] );; let pow_newton_identities_monic_vanishing_at = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r (poly_pow r (x_truncreverse r (CARD S) (monic_vanishing_at r S c)) (e+1) ) (poly_sum r S (\s. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r (c s) n) ) )) = poly_sum r S (\t. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1) ) ) `, intro THEN simp[x_truncreverse_monic_vanishing_at] THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_series] THEN simp[poly_product_pow] THEN qed[pow_newton_identities_natural] );; let scaled_pow_newton_identities_monic_vanishing_at_lemma = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r ( const_x_pow r (ring_of_num r (FACT(e))) e ) ( poly_sum r S (\s. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom(n+e,e))) (ring_pow r (c s) n) ) ) ) = poly_sum r S (\s. series_from_coeffs (\n. ring_mul r (ring_of_num r (FACT(e) * binom(n,e))) (ring_pow r (c s) (n-e)) ) ) `, intro THEN sufficesby eq_coeff THEN intro THEN have `ring_of_num r (FACT e) IN ring_carrier(r:R ring)` [RING_OF_NUM] THEN have `!s:X n. s IN S ==> ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n) IN ring_carrier(r:R ring)` [RING_POW;RING_OF_NUM;RING_MUL] THEN have `!s:X. s IN S ==> ring_powerseries(r:R ring) (series_from_coeffs (\n. ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n)))` [series_series_from_coeffs] THEN have `ring_powerseries(r:R ring) (poly_sum r S (\s:X. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n))))` [poly_sum_series] THEN simp[coeff_const_x_pow_times] THEN have `!s:X n. s IN S ==> ring_mul r (ring_of_num r (FACT e * binom (n,e))) (ring_pow r (c s) (n - e)) IN ring_carrier(r:R ring)` [RING_POW;RING_OF_NUM;RING_MUL] THEN subgoal `!s:X. s IN S ==> ring_powerseries(r:R ring) (series_from_coeffs (\n. ring_mul r (ring_of_num r (FACT e * binom (n,e))) (ring_pow r (c s) (n - e))))` THENL [ simp[series_series_from_coeffs] ; pass ] THEN simp[coeff_poly_sum] THEN rw[coeff_series_from_coeffs] THEN case `d < e:num` THENL [ simp[BINOM_LT;MULT_0;RING_OF_NUM_0;RING_MUL_LZERO;RING_POW] THEN qed[RING_SUM_0] ; pass ] THEN num_linear_fact `~(d < e:num) ==> d-e+e = d` THEN simp[] THEN have `ring_of_num(r:R ring) (FACT e) IN ring_carrier r` [RING_OF_NUM] THEN have `!s:X. s IN S ==> ring_mul(r:R ring) (ring_of_num r (binom(d,e))) (ring_pow r (c s) (d-e)) IN ring_carrier(r:R ring)` [RING_OF_NUM;RING_POW;RING_MUL] THEN specialize[`r:R ring`;`\a:X. ring_mul(r:R ring) (ring_of_num r (binom (d,e))) (ring_pow r (c a) (d - e))`;`ring_of_num(r:R ring) (FACT e)`;`S:X->bool`](GSYM RING_SUM_LMUL) THEN simp[] THEN sufficesby RING_SUM_EQ THEN intro THEN simp[] THEN simp[RING_MUL_ASSOC;RING_OF_NUM;RING_POW] THEN rw[RING_OF_NUM_MUL] );; let scaled_pow_newton_rightside = new_definition ` scaled_pow_newton_rightside (r:R ring) (c:X->R) (S:X->bool) (e:num) = poly_mul r ( const_x_pow r (ring_of_num r (FACT(e))) e ) ( poly_sum r S (\t. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1) ) ) ) `;; let scaled_pow_newton_identities_monic_vanishing_at = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_mul r (poly_pow r (x_truncreverse r (CARD S) (monic_vanishing_at r S c)) (e+1) ) (poly_sum r S (\s. series_from_coeffs (\n. ring_mul r (ring_of_num r (FACT(e) * binom(n,e))) (ring_pow r (c s) (n-e)) ) )) = scaled_pow_newton_rightside r c S e `, rw[scaled_pow_newton_rightside] THEN intro THEN simp[GSYM pow_newton_identities_monic_vanishing_at] THEN simp[GSYM scaled_pow_newton_identities_monic_vanishing_at_lemma] THEN have `ring_powerseries(r:R ring) (const_x_pow r (ring_of_num r (FACT e)) e)` [const_x_pow_series;RING_OF_NUM] THEN have `ring_powerseries(r:R ring) (poly_pow r (x_truncreverse r (CARD S) (monic_vanishing_at r (S:X->bool) c)) (e + 1))` [monic_vanishing_at_series;x_truncreverse_series;poly_pow_series] THEN have `!s:X n. s IN S ==> ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n) IN ring_carrier(r:R ring)` [RING_MUL;RING_OF_NUM;RING_POW] THEN have `ring_powerseries(r:R ring) (poly_sum r S (\s:X. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n))))` [poly_sum_series;series_series_from_coeffs] THEN specialize[`r:R ring` ;`const_x_pow(r:R ring) (ring_of_num r (FACT e)) e` ;`poly_pow(r:R ring) (x_truncreverse r (CARD S) (monic_vanishing_at r (S:X->bool) c)) (e + 1)` ;`poly_sum(r:R ring) S (\s:X. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n)))` ]POLY_MUL_ASSOC THEN specialize[`r:R ring` ;`const_x_pow(r:R ring) (ring_of_num r (FACT e)) e` ;`poly_pow(r:R ring) (x_truncreverse r (CARD S) (monic_vanishing_at r (S:X->bool) c)) (e + 1)` ]POLY_MUL_SYM THEN specialize[`r:R ring` ;`poly_pow(r:R ring) (x_truncreverse r (CARD S) (monic_vanishing_at r (S:X->bool) c)) (e + 1)` ;`const_x_pow(r:R ring) (ring_of_num r (FACT e)) e` ;`poly_sum(r:R ring) S (\s:X. series_from_coeffs (\n. ring_mul r (ring_of_num r (binom (n + e,e))) (ring_pow r (c s) n)))` ]POLY_MUL_ASSOC THEN qed[] );; let poly_pow_newton_identities_monic_vanishing_at_lemma = prove(` !(r:R ring) c:X->R S e t. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> t IN S ==> ring_polynomial(r:R ring) (poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e + 1))) `, intro THEN have `FINITE(S DELETE (t:X))` [FINITE_DELETE] THEN set_fact `!s:X. s IN S DELETE t ==> s IN S` THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_poly] THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e + 1))` [poly_pow_poly] THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e+1))` [] THEN specialize[`r:R ring`;`\s:X. poly_pow(r:R ring) (one_minus_constx r (c s)) (e+1)`;`S DELETE (t:X)`]poly_product_poly THEN qed[] );; let poly_pow_newton_identities_monic_vanishing_at = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_polynomial r ( poly_sum r S (\t. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1) ) ) ) `, intro THEN set_fact `!s t:X. s IN S DELETE t ==> s IN S` THEN have `!t:X. t IN S ==> !s. s IN S DELETE t ==> ring_polynomial(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_poly] THEN have `!t:X. t IN S ==> !s. s IN S DELETE t ==> ring_polynomial(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e + 1))` [poly_pow_poly] THEN have `!t:X. t IN S ==> ring_polynomial(r:R ring) (poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e + 1)))` [poly_pow_newton_identities_monic_vanishing_at_lemma] THEN qed[poly_sum_poly] );; let poly_scaled_pow_newton_rightside = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> ring_polynomial r (scaled_pow_newton_rightside r c S e) `, rw[scaled_pow_newton_rightside] THEN qed[poly_pow_newton_identities_monic_vanishing_at;RING_OF_NUM;const_x_pow_poly;RING_POLYNOMIAL_MUL] );; let deg_pow_newton_identities_monic_vanishing_at_lemma = prove(` !(r:R ring) c:X->R S e t. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> t IN S ==> poly_deg(r:R ring) (poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e + 1))) <= (CARD S - 1) * (e+1) `, intro THEN set_fact `!s t:X. s IN S DELETE t ==> s IN S` THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_poly] THEN subgoal `!s:X. s IN S DELETE t ==> poly_deg(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e+1)) <= e+1` THENL [ intro THEN have `poly_deg(r:R ring) (one_minus_constx r (c(s:X))) <= 1` [deg_one_minus_constx_le] THEN have `(e+1) * poly_deg(r:R ring) (one_minus_constx r (c(s:X))) <= (e+1)*1` [LE_MULT_LCANCEL] THEN have `poly_deg(r:R ring) (poly_pow r (one_minus_constx r (c(s:X))) (e+1)) <= (e+1) * poly_deg r (one_minus_constx r (c(s:X)))` [poly_deg_pow_le] THEN qed[LE_TRANS;ARITH_RULE `(e+1)*1 = e+1`] ; pass ] THEN have `FINITE(S DELETE (t:X))` [FINITE_DELETE] THEN have `!s:X. s IN S DELETE t ==> ring_polynomial(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e+1))` [poly_pow_poly;one_minus_constx_poly] THEN have `!s:X. s IN S DELETE t ==> poly_deg(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e+1)) <= e+1` [poly_pow_poly;one_minus_constx_poly] THEN specialize[`r:R ring`;`\s:X. poly_pow r (one_minus_constx r (c s:R)) (e+1)`;`e+1`;`S DELETE (t:X)`]poly_deg_product_each_le THEN qed[CARD_DELETE] );; let deg_pow_newton_identities_monic_vanishing_at = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_deg r ( poly_sum r S (\t. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1) ) ) ) <= (CARD S - 1) * (e+1) `, intro THEN set_fact `!s t:X. s IN S DELETE t ==> s IN S` THEN have `!t:X. t IN S ==> !s. s IN S DELETE t ==> ring_polynomial(r:R ring) (one_minus_constx r (c s))` [one_minus_constx_poly] THEN have `!t:X. t IN S ==> !s. s IN S DELETE t ==> ring_polynomial(r:R ring) (poly_pow r (one_minus_constx r (c s)) (e + 1))` [poly_pow_poly] THEN have `!t:X. t IN S ==> ring_polynomial(r:R ring) (poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e + 1)))` [poly_pow_newton_identities_monic_vanishing_at_lemma] THEN have `!t:X. t IN S ==> poly_deg(r:R ring) (poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e + 1))) <= (CARD S - 1) * (e+1)` [deg_pow_newton_identities_monic_vanishing_at_lemma] THEN specialize[`r:R ring`;`(\t:X. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s:R)) (e + 1)))`;`(CARD(S:X->bool)-1)*(e+1)`;`S:X->bool`]poly_deg_sum_le THEN qed[] );; let deg_scaled_pow_newton_rightside = prove(` !(r:R ring) c:X->R S e. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> poly_deg r (scaled_pow_newton_rightside r c S e) <= e + (CARD S - 1) * (e+1) `, intro THEN rw[scaled_pow_newton_rightside] THEN have `ring_polynomial(r:R ring) (const_x_pow r (ring_of_num r (FACT(e))) e)` [const_x_pow_poly;RING_OF_NUM] THEN have `poly_deg(r:R ring) (const_x_pow r (ring_of_num r (FACT(e))) e) <= e` [deg_const_x_pow_le;RING_OF_NUM] THEN have `ring_polynomial(r:R ring) (poly_sum r S (\t:X. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1))))` [poly_pow_newton_identities_monic_vanishing_at] THEN have `poly_deg(r:R ring) (poly_sum r S (\t:X. poly_product r (S DELETE t) (\s. poly_pow r (one_minus_constx r (c s)) (e+1)))) <= (CARD S - 1) * (e+1)` [deg_pow_newton_identities_monic_vanishing_at] THEN qed[POLY_DEG_MUL_LE;ARITH_RULE `d <= d1 + d2 /\ d1 <= e:num /\ d2 <= C ==> d <= e + C`] );; (* ===== the ring of complex numbers *) let complex_ring = new_definition ` complex_ring = ring((:complex),Cx(&0),Cx(&1),( -- ),( + ),( * )) `;; let complex_of_num = new_definition ` complex_of_num(n:num) = Cx(real_of_num n) `;; let complex_of_int = new_definition ` complex_of_int(n:int) = Cx(real_of_int n) `;; let complex_ring_clauses = prove (`ring_carrier complex_ring = (:complex) /\ ring_0 complex_ring = Cx(&0) /\ ring_1 complex_ring = Cx(&1) /\ ring_neg complex_ring = ( -- ) /\ ring_add complex_ring = ( + ) /\ ring_mul complex_ring = ( * )`, PURE_REWRITE_TAC [GSYM PAIR_EQ; ring_carrier; ring_0; ring_1; ring_neg; ring_add; ring_mul; BETA_THM; PAIR] THEN PURE_REWRITE_TAC[complex_ring; GSYM(CONJUNCT2 ring_tybij)] THEN REWRITE_TAC[IN_UNIV] THEN CONV_TAC COMPLEX_RING);; let in_complex_ring = prove(` !z. z IN ring_carrier complex_ring `, qed[complex_ring_clauses;IN_UNIV] );; let field_complex = prove(` field complex_ring `, rw[field;complex_ring_clauses;IN_UNIV] THEN qed[CX_INJ;COMPLEX_MUL_RINV;REAL_ARITH `~(&1 = &0:real)`] );; let integral_domain_complex = prove(` integral_domain complex_ring `, qed[field_complex;FIELD_IMP_INTEGRAL_DOMAIN] );; let ring_of_num_complex = prove(` ring_of_num complex_ring = complex_of_num `, rw[FUN_EQ_THM;complex_of_num] THEN INDUCT_TAC THEN simp[ring_of_num;complex_ring_clauses] THEN rw[GSYM CX_ADD;REAL_OF_NUM_SUC] );; let ring_of_int_complex = prove(` ring_of_int complex_ring = complex_of_int `, rw[FUN_EQ_THM] THEN rw[FORALL_INT_CASES;RING_OF_INT_CASES] THEN rw[ring_of_num_complex;complex_of_num] THEN rw[complex_of_int;int_of_num_th;int_neg_th;CX_NEG] THEN rw[complex_ring_clauses] );; let ring_char_complex = prove(` ring_char complex_ring = 0 `, rw[RING_CHAR_EQ_0;ring_of_num_complex;complex_of_num;complex_ring_clauses;CX_INJ;REAL_OF_NUM_EQ] );; let ring_sub_complex = prove(` ring_sub complex_ring = (-) `, rw[FUN_EQ_THM;ring_sub;complex_ring_clauses;complex_sub] );; (* from Complex/complexnumbers.ml *) let COMPLEX_MUL_RINV_UNIQ = prove (`!w z. w * z = Cx(&1) ==> inv w = z`, CONV_TAC COMPLEX_FIELD);; let ring_inv_complex = prove(` ring_inv complex_ring = inv `, rw[FUN_EQ_THM;ring_inv;complex_ring_clauses] THEN simp[FIELD_UNIT;field_complex;complex_ring_clauses;IN_UNIV;COND_SWAP] THEN intro THEN case `x = Cx(&0)` THENL [ qed[COMPLEX_INV_0] ; qed[COMPLEX_MUL_RINV;COMPLEX_MUL_RINV_UNIQ] ] );; let ring_div_complex = prove(` ring_div complex_ring = (/) `, rw[FUN_EQ_THM;ring_div;complex_ring_clauses;ring_inv_complex;complex_div] );; let ring_pow_complex = prove(` ring_pow complex_ring = (pow) `, rw[FUN_EQ_THM] THEN GEN_TAC THEN INDUCT_TAC THEN simp[ring_pow;complex_ring_clauses] THEN qed[complex_pow] );; let complex_field_clauses = prove (`ring_carrier complex_ring = (:complex) /\ ring_0 complex_ring = Cx(&0) /\ ring_1 complex_ring = Cx(&1) /\ ring_neg complex_ring = ( -- ) /\ ring_add complex_ring = ( + ) /\ ring_mul complex_ring = ( * ) /\ ring_of_num complex_ring = complex_of_num /\ ring_sub complex_ring = (-) /\ ring_inv complex_ring = inv /\ ring_div complex_ring = (/) /\ ring_pow complex_ring = (pow)`, rw[complex_ring_clauses;ring_of_num_complex;ring_sub_complex;ring_inv_complex;ring_div_complex;ring_pow_complex]);; let vsum_ring_sum_complex = prove(` !(f:X->complex) S. FINITE S ==> vsum S f = ring_sum complex_ring S f `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[VSUM_CLAUSES;RING_SUM_CLAUSES] THEN rw[complex_ring_clauses;COMPLEX_VEC_0] ; simp[VSUM_CLAUSES;RING_SUM_CLAUSES;in_complex_ring] THEN rw[complex_ring_clauses] ] );; let cproduct_ring_product_complex = prove(` !(f:X->complex) S. FINITE S ==> cproduct S f = ring_product complex_ring S f `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[CPRODUCT_CLAUSES;RING_PRODUCT_CLAUSES] THEN rw[complex_ring_clauses] ; simp[CPRODUCT_CLAUSES;RING_PRODUCT_CLAUSES;in_complex_ring] THEN rw[complex_ring_clauses] ] );; let vsum_delta_complex = prove(` !(S:X->bool) t a. FINITE S ==> vsum S (\s. if s = t then a else Cx(&0)) = if t IN S then a else Cx(&0) `, rw[GSYM COMPLEX_VEC_0] THEN qed[VSUM_DELTA] );; let ring_1_0_complex = prove(` ~(ring_1 complex_ring = ring_0 complex_ring) `, qed[integral_domain_complex;integral_domain] );; (* ===== \C[x] *) let series_complex = prove(` !(p:(1->num)->complex). ring_powerseries complex_ring p `, rw[ring_powerseries;in_complex_ring] THEN qed[FINITE_MONOMIAL_VARS_1;INFINITE] );; let complex_coeff_x_pow_times = prove(` !d p e. coeff e (poly_mul complex_ring (x_pow complex_ring d) p) = if e < d then Cx(&0) else coeff (e - d) p `, qed[coeff_x_pow_times;series_complex;complex_ring_clauses] );; let poly_1_0_complex = prove(` ~(poly_1 complex_ring = poly_0 complex_ring:(V->num)->complex) `, rw[poly_1;poly_0;POLY_CONST_EQ] THEN qed[ring_1_0_complex] );; (* ===== copy of \Z inside \C *) let ZinC = new_definition ` ZinC = {z:complex | ?i:int. z = complex_of_int i} `;; let ZinC_ring = new_definition ` ZinC_ring = ring(ZinC,Cx(&0),Cx(&1),( -- ),( + ),( * )) `;; let ii_not_in_ZinC = prove(` ~(ii IN ZinC) `, rw[ZinC;IN_ELIM_THM;complex_of_int] THEN rw[ii;COMPLEX_EQ;CX_DEF;IM] THEN ARITH_TAC );; let half_not_in_ZinC = prove(` ~(Cx(&1 / &2) IN ZinC) `, rw[ZinC;IN_ELIM_THM;complex_of_int] THEN rw[CX_INJ] THEN intro THEN int_linear_fact `&1 / &2 = real_of_int i ==> &1 = &2 * (real_of_int i)` THEN have `&1 = real_of_int(&1)` [REAL_OF_INT_CLAUSES] THEN have `&2 = real_of_int(&2)` [REAL_OF_INT_CLAUSES] THEN have `real_of_int(&1) = real_of_int(&2) * real_of_int i` [] THEN have `real_of_int(&1) = real_of_int(&2 * i)` [REAL_OF_INT_CLAUSES] THEN have `&1 = &2 * i:int` [REAL_OF_INT_CLAUSES] THEN DISJ_CASES_TAC(SPEC `i:int` INT_IMAGE) THENL [ choose `n:num` `i:int = &n` [] THEN have `&1 = &2 * &n:int` [] THEN num_linear_fact `&1 = &2 * &n:int ==> 1 = 2 * n` THEN have `EVEN 1` [EVEN_EXISTS] THEN num_linear_fact `1 = SUC(2 * 0)` THEN have `ODD 1` [ODD_EXISTS] THEN qed[NOT_EVEN] ; ASM_ARITH_TAC ] );; let num_in_ZinC = prove(` !d. Cx(&d) IN ZinC `, intro THEN rw[ZinC;IN_ELIM_THM;complex_of_int] THEN witness `&d:int` THEN rw[int_of_num_th] );; let int_in_ZinC = prove(` !i. complex_of_int i IN ZinC `, SET_TAC[ZinC] );; let ZinC_0 = prove(` Cx(&0) IN ZinC `, qed[num_in_ZinC] );; let ZinC_1 = prove(` Cx(&1) IN ZinC `, qed[num_in_ZinC] );; let neg_in_ZinC = prove(` !x:complex. x IN ZinC ==> --x IN ZinC `, rw[ZinC;IN_ELIM_THM;complex_of_int] THEN intro THEN witness `int_neg i` THEN rw[GSYM REAL_OF_INT_CLAUSES;CX_NEG] THEN qed[] );; let add_in_ZinC = prove(` !x:complex y:complex. x IN ZinC /\ y IN ZinC ==> x+y IN ZinC `, rw[ZinC;IN_ELIM_THM;complex_of_int] THEN intro THEN witness `i+i':int` THEN rw[GSYM REAL_OF_INT_CLAUSES;CX_ADD] THEN qed[] );; let mul_in_ZinC = prove(` !x:complex y:complex. x IN ZinC /\ y IN ZinC ==> x*y IN ZinC `, rw[ZinC;IN_ELIM_THM;complex_of_int] THEN intro THEN witness `i*i':int` THEN rw[GSYM REAL_OF_INT_CLAUSES;CX_MUL] THEN qed[] );; let ZinC_ring_clauses = prove(` ring_carrier ZinC_ring = ZinC /\ ring_0 ZinC_ring = Cx(&0) /\ ring_1 ZinC_ring = Cx(&1) /\ ring_neg ZinC_ring = ( -- ) /\ ring_add ZinC_ring = ( + ) /\ ring_mul ZinC_ring = ( * ) `, rw[ring_carrier;ring_0;ring_1;ring_neg;ring_add;ring_mul;GSYM PAIR_EQ] THEN rw[ZinC_ring] THEN rw[GSYM(CONJUNCT2 ring_tybij)] THEN rw[ZinC_0;ZinC_1;neg_in_ZinC;add_in_ZinC;mul_in_ZinC] THEN qed[COMPLEX_ADD_SYM;COMPLEX_ADD_ASSOC;COMPLEX_ADD_LID;COMPLEX_ADD_LINV;COMPLEX_MUL_ASSOC;COMPLEX_MUL_SYM;COMPLEX_MUL_LID;COMPLEX_ADD_LDISTRIB] );; let ZinC_subring_complex = prove(` ZinC subring_of complex_ring `, rw[subring_of] THEN rw[complex_ring_clauses] THEN rw[ZinC_0;ZinC_1;neg_in_ZinC;add_in_ZinC;mul_in_ZinC;SUBSET_UNIV] );; let ZinC_subring_generated_carrier = prove(` ZinC = INTERS {S | S subring_of complex_ring} `, rw[INTERS] THEN rw[ZinC;IN_ELIM_THM;EXTENSION] THEN intro THEN splitiff THENL [ rw[GSYM ring_of_int_complex] THEN qed[int_in_subring] ; intro THEN have `ZinC subring_of complex_ring` [ZinC_subring_complex] THEN have `x IN ZinC` [] THEN have `x IN {z | ?i. z = complex_of_int i}` [ZinC] THEN set_fact `x IN {z | ?i. z = complex_of_int i} ==> ?i. x = complex_of_int i` THEN qed[] ] );; let subring_complex_empty_lemma = prove(` INTERS {S | S subring_of complex_ring /\ (:complex) INTER {} SUBSET S} = ZinC `, rw[ZinC_subring_generated_carrier] THEN SET_TAC [] );; let subring_complex_empty = prove(` subring_generated complex_ring {} = ZinC_ring `, rw[subring_generated] THEN rw[complex_ring_clauses] THEN rw[ZinC_ring;subring_complex_empty_lemma] );; let ring_char_ZinC = prove(` ring_char ZinC_ring = 0 `, recall ring_char_complex THEN specialize[`complex_ring`;`{}:complex->bool`]RING_CHAR_SUBRING_GENERATED THEN qed[subring_complex_empty] );; let subring_complex_ZinC_lemma = prove(` INTERS {S | S subring_of complex_ring /\ (:complex) INTER ZinC SUBSET S} = ZinC `, rw[ZinC_subring_generated_carrier] THEN SET_TAC [] );; let subring_complex_ZinC = prove(` subring_generated complex_ring ZinC = ZinC_ring `, rw[subring_generated] THEN rw[complex_ring_clauses] THEN rw[ZinC_ring;subring_complex_ZinC_lemma] );; let ring_of_num_ZinC = prove(` ring_of_num ZinC_ring = complex_of_num `, rw[FUN_EQ_THM;complex_of_num] THEN INDUCT_TAC THEN simp[ring_of_num;ZinC_ring_clauses] THEN rw[GSYM CX_ADD;REAL_OF_NUM_SUC] );; let ring_pow_ZinC = prove(` ring_pow ZinC_ring = (pow) `, qed[RING_POW_SUBRING_GENERATED;ring_pow_complex;subring_complex_empty] );; let integral_domain_ZinC = prove(` integral_domain ZinC_ring `, qed[integral_domain_complex;INTEGRAL_DOMAIN_SUBRING_GENERATED;subring_complex_empty] );; (* ===== \Z[x] *) let poly_complex_if_poly_ZinC = prove(` !p:(V->num)->complex. ring_polynomial ZinC_ring p ==> ring_polynomial complex_ring p `, rw[ring_polynomial] THEN rw[ring_powerseries] THEN rw[complex_ring_clauses;ZinC_ring_clauses] THEN qed[IN_UNIV] );; let poly_0_ZinC_eq_poly_0_complex = prove(` poly_0 ZinC_ring = (poly_0 complex_ring):(V->num)->complex `, rw[poly_0;poly_const] THEN rw[ZinC_ring_clauses;complex_ring_clauses] );; let zero_if_ZinC_norm_lt1 = prove(` !z. z IN ZinC ==> norm z < &1 ==> z = Cx(&0) `, intro THEN have `z IN {z | ?i. z = complex_of_int i}` [ZinC] THEN set_fact `z IN {z | ?i. z = complex_of_int i} ==> ?i. z = complex_of_int i` THEN choose `i:int` `z = complex_of_int i` [ZinC] THEN have `z = Cx(real_of_int i)` [complex_of_int] THEN have `norm(z:complex) = abs(real_of_int i)` [COMPLEX_NORM_CX] THEN choose `n:num` `abs(real_of_int i) = &n` [INTEGER_REAL_OF_INT;integer] THEN proven_if `n = 0` [REAL_ABS_ZERO] THEN num_linear_fact `~(n = 0) ==> 1 <= n` THEN have `&1 <= &n:real` [REAL_OF_NUM_LE] THEN have `norm(z:complex) < norm z` [REAL_LTE_TRANS] THEN qed[REAL_LT_REFL] );; let zero_if_ZinC_scale_bound = prove(` !(c:complex) (r:real) z. ~(c = Cx(&0)) ==> c * z IN ZinC ==> norm z <= r ==> norm c * r < &1 ==> z = Cx(&0) `, intro THEN subgoal `norm (c*z) < &1` THENL [ rw[COMPLEX_NORM_MUL] THEN qed[REAL_LE_LMUL;NORM_POS_LE;REAL_LET_TRANS] ; have `c*z = Cx(&0)` [zero_if_ZinC_norm_lt1] THEN complex_field_fact `~(c = Cx(&0)) ==> c*z = Cx(&0) ==> z = Cx(&0)` THEN qed[] ] );; (* ===== copy of \Q inside \C *) let QinC = new_definition ` QinC = {z:complex | ?i:int d:num. ~(d = 0) /\ Cx(&d) * z = complex_of_int i} `;; let QinC_ring = new_definition ` QinC_ring = ring(QinC,Cx(&0),Cx(&1),( -- ),( + ),( * )) `;; let ZinC_in_QinC = prove(` !z. z IN ZinC ==> z IN QinC `, rw[ZinC;QinC;IN_ELIM_THM] THEN intro THEN witness `i:int` THEN witness `1` THEN simp[COMPLEX_MUL_LID] THEN qed[ARITH_RULE `~(1 = 0)`] );; let ZinC_subset_QinC = prove(` ZinC SUBSET QinC `, SET_TAC[ZinC_in_QinC] );; let QinC_0 = prove(` Cx(&0) IN QinC `, SET_TAC[ZinC_in_QinC;ZinC_0] );; let QinC_1 = prove(` Cx(&1) IN QinC `, SET_TAC[ZinC_in_QinC;ZinC_1] );; let neg_in_QinC = prove(` !z. z IN QinC ==> --z IN QinC `, rw[QinC;IN_ELIM_THM] THEN rw[complex_of_int] THEN intro THEN witness `--i:int` THEN witness `d:num` THEN rw[GSYM REAL_OF_INT_CLAUSES] THEN rw[CX_NEG] THEN qed[COMPLEX_MUL_RNEG] );; let add_in_QinC = prove(` !y z. y IN QinC /\ z IN QinC ==> y+z IN QinC `, rw[QinC;IN_ELIM_THM] THEN rw[complex_of_int] THEN intro THEN witness `i * &d'+i' * &d:int` THEN witness `d*d':num` THEN have `~(d * d' = 0)` [MULT_EQ_0] THEN rw[int_add_th] THEN rw[int_mul_th] THEN rw[GSYM REAL_OF_NUM_MUL] THEN rw[CX_MUL;CX_ADD] THEN rw[GSYM REAL_OF_INT_CLAUSES] THEN complex_field_fact `(Cx(&d) * Cx(&d'))*(y+z) = (Cx(&d)*y)*Cx(&d') + (Cx(&d')*z)*Cx(&d)` THEN qed[] );; let mul_in_QinC = prove(` !y z. y IN QinC /\ z IN QinC ==> y*z IN QinC `, rw[QinC;IN_ELIM_THM] THEN rw[complex_of_int] THEN intro THEN witness `i*i':int` THEN witness `d*d':num` THEN have `~(d*d' = 0)` [MULT_EQ_0] THEN rw[int_mul_th] THEN rw[GSYM REAL_OF_NUM_MUL] THEN rw[CX_MUL] THEN complex_field_fact `(Cx(&d)*Cx(&d'))*(y*z) = (Cx(&d)*y)*(Cx(&d')*z)` THEN qed[] );; let QinC_ring_clauses = prove(` ring_carrier QinC_ring = QinC /\ ring_0 QinC_ring = Cx(&0) /\ ring_1 QinC_ring = Cx(&1) /\ ring_neg QinC_ring = ( -- ) /\ ring_add QinC_ring = ( + ) /\ ring_mul QinC_ring = ( * ) `, rw[ring_carrier;ring_0;ring_1;ring_neg;ring_add;ring_mul;GSYM PAIR_EQ] THEN rw[QinC_ring] THEN rw[GSYM(CONJUNCT2 ring_tybij)] THEN rw[QinC_0;QinC_1;neg_in_QinC;add_in_QinC;mul_in_QinC] THEN qed[COMPLEX_ADD_SYM;COMPLEX_ADD_ASSOC;COMPLEX_ADD_LID;COMPLEX_ADD_LINV;COMPLEX_MUL_ASSOC;COMPLEX_MUL_SYM;COMPLEX_MUL_LID;COMPLEX_ADD_LDISTRIB] );; let ZinC_subring_QinC = prove(` ZinC subring_of QinC_ring `, rw[subring_of] THEN rw[QinC_ring_clauses] THEN rw[ZinC_0;ZinC_1;neg_in_ZinC;add_in_ZinC;mul_in_ZinC;ZinC_subset_QinC] );; let QinC_subring_complex = prove(` QinC subring_of complex_ring `, rw[subring_of] THEN rw[complex_ring_clauses] THEN rw[QinC_0;QinC_1;neg_in_QinC;add_in_QinC;mul_in_QinC;SUBSET_UNIV] );; let series_QinC_if_series_ZinC = prove(` !p:(V->num)->complex. ring_powerseries ZinC_ring p ==> ring_powerseries QinC_ring p `, rw[ring_powerseries] THEN rw[ZinC_ring_clauses;QinC_ring_clauses] THEN qed[ZinC_in_QinC] );; let poly_QinC_if_poly_ZinC = prove(` !p:(V->num)->complex. ring_polynomial ZinC_ring p ==> ring_polynomial QinC_ring p `, rw[ring_polynomial] THEN rw[ring_powerseries] THEN rw[ZinC_ring_clauses;QinC_ring_clauses] THEN qed[ZinC_in_QinC] );; let poly_complex_if_poly_QinC = prove(` !p:(V->num)->complex. ring_polynomial QinC_ring p ==> ring_polynomial complex_ring p `, rw[ring_polynomial] THEN rw[ring_powerseries] THEN rw[complex_ring_clauses;QinC_ring_clauses] THEN qed[IN_UNIV] );; let series_complex_if_series_QinC = prove(` !p:(V->num)->complex. ring_powerseries QinC_ring p ==> ring_powerseries complex_ring p `, rw[ring_powerseries] THEN rw[complex_ring_clauses;QinC_ring_clauses] THEN qed[IN_UNIV] );; let poly_0_QinC_eq_poly_0_complex = prove(` poly_0 QinC_ring = (poly_0 complex_ring):(V->num)->complex `, rw[poly_0;poly_const] THEN rw[QinC_ring_clauses;complex_ring_clauses] );; let poly_1_QinC_eq_poly_1_complex = prove(` poly_1 QinC_ring = (poly_1 complex_ring):(V->num)->complex `, rw[poly_1;poly_const] THEN rw[QinC_ring_clauses;complex_ring_clauses] );; let poly_neg_QinC_eq_poly_neg_complex = prove(` !p:(V->num)->complex. poly_neg QinC_ring p = poly_neg complex_ring p `, rw[poly_neg] THEN rw[QinC_ring_clauses;complex_ring_clauses] );; let poly_add_QinC_eq_poly_add_complex = prove(` !p:(V->num)->complex q. poly_add QinC_ring p q = poly_add complex_ring p q `, rw[poly_add] THEN rw[QinC_ring_clauses;complex_ring_clauses] );; let x_pow_QinC_eq_x_pow_complex = prove(` !n. x_pow QinC_ring n = x_pow complex_ring n `, intro THEN sufficesby eq_coeff THEN rw[coeff_x_pow;QinC_ring_clauses;complex_ring_clauses] );; let poly_0_ZinC_eq_poly_0_QinC = prove(` poly_0 ZinC_ring = (poly_0 QinC_ring):(V->num)->complex `, qed[poly_0_ZinC_eq_poly_0_complex;poly_0_QinC_eq_poly_0_complex] );; let ring_of_num_QinC = prove(` ring_of_num QinC_ring = complex_of_num `, rw[FUN_EQ_THM;complex_of_num] THEN INDUCT_TAC THEN simp[ring_of_num;QinC_ring_clauses] THEN rw[GSYM CX_ADD;REAL_OF_NUM_SUC] );; let num_in_QinC = prove(` !n. Cx(&n) IN QinC `, rw[GSYM complex_of_num;GSYM ring_of_num_QinC;GSYM QinC_ring_clauses;RING_OF_NUM] );; let num_over_num_in_QinC = prove(` !n d. Cx(&n) / Cx(&d) IN QinC `, rw[QinC;IN_ELIM_THM] THEN intro THEN case `d = 0` THENL [ witness `&0:int` THEN witness `1:num` THEN rw[complex_of_int;GSYM REAL_OF_INT_CLAUSES] THEN simp[complex_div] THEN rw[COMPLEX_INV_0] THEN rw[COMPLEX_MUL_RZERO] THEN qed[ARITH_RULE `~(1 = 0)`] ; witness `&n:int` THEN witness `d:num` THEN rw[complex_of_int;GSYM REAL_OF_INT_CLAUSES] THEN have `~(d = 0) ==> ~(&d = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&d) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx (&d) = Cx (&0)) ==> Cx (&d) * Cx (&n) / Cx (&d) = Cx (&n)` THEN qed[] ] );; let neg_num_over_num_in_QinC = prove(` !n d. -- (Cx(&n) / Cx(&d)) IN QinC `, qed[num_over_num_in_QinC;neg_in_QinC] );; let field_QinC = prove(` field QinC_ring `, rw[field] THEN rw[QinC_ring_clauses] THEN intro THENL [ qed[CX_INJ;REAL_ARITH `~(&1 = &0:real)`] ; choose2 `i:int` `d:num` `~(d = 0) /\ Cx(&d) * x = complex_of_int i` [QinC;IN_ELIM_THM] THEN choose `n:num` `real_of_int i = &n \/ real_of_int i = -- &n` [dest_int_rep] THEN case `n = 0` THENL [ num_linear_fact `-- &0 = &0:real` THEN have `real_of_int i = &0` [] THEN have `complex_of_int i = Cx(&0)` [complex_of_int] THEN have `Cx(&d) * x = Cx(&0)` [] THEN complex_field_fact `~(x = Cx(&0)) /\ Cx(&d) * x = Cx(&0) ==> Cx(&d) = Cx(&0)` THEN have `&d = &0:real` [CX_INJ] THEN qed[REAL_OF_NUM_EQ] ; pass ] THEN have `~(n = 0) ==> ~(&n = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&n) = Cx(&0))` [CX_INJ] THEN case `real_of_int i = &n` THENL [ have `complex_of_int i = Cx(&n)` [complex_of_int] THEN witness `Cx(&d) / Cx(&n)` THEN have `Cx(&d) * x = Cx(&n)` [] THEN complex_field_fact `Cx(&d)*x = Cx(&n) ==> ~(Cx(&n) = Cx(&0)) ==> x * Cx(&d) / Cx(&n) = Cx(&1)` THEN qed[num_over_num_in_QinC] ; have `real_of_int i = -- &n` [] THEN have `complex_of_int i = Cx(-- &n)` [complex_of_int] THEN witness `-- (Cx(&d) / Cx(&n))` THEN have `Cx(&d) * x = -- Cx(&n)` [CX_NEG] THEN complex_field_fact `Cx(&d)*x = -- Cx(&n) ==> ~(Cx(&n) = Cx(&0)) ==> x * -- (Cx(&d) / Cx(&n)) = Cx(&1)` THEN qed[neg_num_over_num_in_QinC] ] ] );; let integral_domain_QinC = prove(` integral_domain QinC_ring `, qed[field_QinC;FIELD_IMP_INTEGRAL_DOMAIN] );; let QinC_over_num = prove(` !z e. z IN QinC ==> ~(e = 0) ==> z / Cx(&e) IN QinC `, rw[QinC;IN_ELIM_THM] THEN intro THEN witness `i:int` THEN witness `d*e:num` THEN rw[GSYM REAL_OF_NUM_MUL] THEN rw[CX_MUL] THEN have `~(e = 0) ==> ~(&e = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&e) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx(&e) = Cx(&0)) ==> (Cx(&d) * Cx(&e)) * z / Cx(&e) = Cx(&d) * z` THEN have `~(d * e = 0)` [MULT_EQ_0] THEN simp[] );; let div_in_QinC = prove(` !y z. y IN QinC /\ z IN QinC /\ ~(z = Cx(&0)) ==> y/z IN QinC `, intro THEN choose `zinv:complex` `zinv IN ring_carrier QinC_ring /\ ring_mul QinC_ring z zinv = ring_1 QinC_ring` [field;QinC_ring_clauses;field_QinC] THEN have `z * zinv = Cx(&1)` [QinC_ring_clauses] THEN complex_field_fact `~(z = Cx(&0)) ==> z * zinv = Cx(&1) ==> y / z = y * zinv` THEN qed[RING_MUL;QinC_ring_clauses] );; let QinC_to_ZinC = prove(` !z. z IN QinC ==> ?d. (~(d = 0) /\ Cx(&d) * z IN ZinC) `, rw[QinC;ZinC;IN_ELIM_THM] THEN qed[] );; let multi_QinC_to_ZinC = prove(` !S f:X->complex. FINITE S ==> (!s. s IN S ==> f s IN QinC) ==> ?d. (~(d = 0) /\ !s. s IN S ==> Cx(&d) * f s IN ZinC) `, intro THEN witness `nproduct S (\s:X. @d. ~(d = 0) /\ Cx(&d) * f s IN ZinC)` THEN conjunction THENL [ simp[NPRODUCT_EQ_0] THEN rw[NOT_EXISTS_THM] THEN qed[QinC_to_ZinC] ; intro THEN specialize[`\t:X. @d. ~(d = 0) /\ Cx (&d) * f t IN ZinC`;`S:X->bool`;`s:X`](GSYM NPRODUCT_DELETE) THEN simp[] THEN once_rw[MULT_SYM] THEN simp[prove(`&(m*n):real = &m * &n`,qed[REAL_OF_NUM_CLAUSES])] THEN rw[CX_MUL;GSYM COMPLEX_MUL_ASSOC] THEN have `(Cx (&(@u:num. ~(u = 0) /\ Cx (&u) * f(s:X) IN ZinC))) * f s IN ZinC` [QinC_ring_clauses;QinC_to_ZinC] THEN qed[mul_in_ZinC;ZinC_ring_clauses;num_in_ZinC] ] );; let subring_complex_QinC_lemma = prove(` INTERS {S | S subring_of complex_ring /\ (:complex) INTER QinC SUBSET S} = QinC `, havetac `QinC IN {S | S subring_of complex_ring /\ (:complex) INTER QinC SUBSET S}` ( rw[IN_ELIM_THM] THEN qed[QinC_subring_complex;INTER_UNIV;SUBSET_REFL] ) THEN set_fact `!x. x IN {S | S subring_of complex_ring /\ (:complex) INTER QinC SUBSET S} ==> QinC SUBSET x` THEN qed[is_inters] );; let subring_complex_QinC = prove(` subring_generated complex_ring QinC = QinC_ring `, rw[subring_generated] THEN rw[complex_ring_clauses] THEN rw[QinC_ring;subring_complex_QinC_lemma] );; let ring_char_QinC = prove(` ring_char QinC_ring = 0 `, qed[ring_char_complex;RING_CHAR_SUBRING_GENERATED;subring_complex_QinC] );; let ring_hasQ_QinC = prove(` ring_hasQ QinC_ring `, rw[ring_hasQ;ring_char_QinC] THEN rw[ring_of_num_QinC;complex_of_num;ring_unit] THEN GEN_TAC THEN DISCH_TAC THEN have `Cx(&n) IN ring_carrier QinC_ring` [num_in_QinC;QinC_ring_clauses] THEN simp[] THEN have `~(Cx(&n) = ring_0 QinC_ring)` [CX_INJ;REAL_OF_NUM_EQ;QinC_ring_clauses] THEN qed[field;field_QinC] );; let ring_sub_QinC = prove(` ring_sub QinC_ring = (-) `, rw[FUN_EQ_THM;ring_sub;QinC_ring_clauses;complex_sub] );; let x_minus_const_QinC_eq_x_minus_const_complex = prove(` !c. x_minus_const QinC_ring c = x_minus_const complex_ring c `, intro THEN rw[x_minus_const] THEN rw[poly_sub;poly_const] THEN rw[ring_sub_QinC;ring_sub_complex] THEN rw[x_pow_QinC_eq_x_pow_complex] THEN rw[QinC_ring_clauses;complex_ring_clauses] );; let subring_QinC_empty_lemma = prove(` INTERS {S | S subring_of QinC_ring /\ QinC INTER {} SUBSET S} = ZinC `, subgoal `ZinC IN {S | S subring_of QinC_ring /\ QinC INTER {} SUBSET S}` THENL [ rw[IN_ELIM_THM] THEN rw[ZinC_subring_QinC;INTER_EMPTY;EMPTY_SUBSET] ; pass ] THEN subgoal `!x. x IN {S | S subring_of QinC_ring /\ QinC INTER {} SUBSET S} ==> ZinC SUBSET x` THENL [ rw[IN_ELIM_THM] THEN intro THEN have `x subring_of complex_ring` [subring_complex_QinC;SUBRING_OF_SUBRING_GENERATED_EQ] THEN set_fact `x subring_of complex_ring ==> INTERS {S | S subring_of complex_ring} SUBSET x` THEN qed[ZinC_subring_generated_carrier] ; pass ] THEN qed[is_inters] );; let subring_QinC_empty = prove(` subring_generated QinC_ring {} = ZinC_ring `, rw[subring_generated] THEN rw[QinC_ring_clauses] THEN rw[ZinC_ring;subring_QinC_empty_lemma] );; let ring_1_0_QinC = prove(` ~(ring_1 QinC_ring = ring_0 QinC_ring) `, qed[integral_domain_QinC;integral_domain] );; let ring_pow_QinC = prove(` ring_pow QinC_ring = (pow) `, qed[RING_POW_SUBRING_GENERATED;ring_pow_complex;subring_complex_QinC] );; let sum_QinC_eq_sum_complex = prove(` !S f. (!s:X. s IN S ==> f s IN QinC) ==> ring_sum QinC_ring S f = ring_sum complex_ring S f `, intro THEN have `!s:X. s IN S ==> f s IN ring_carrier(subring_generated complex_ring QinC)` [subring_complex_QinC;QinC_ring_clauses] THEN have `ring_sum(subring_generated complex_ring QinC) S (f:X->complex) = ring_sum complex_ring S f` [ring_sum_subring_generated_v2] THEN qed[subring_complex_QinC] );; let poly_mul_QinC_eq_poly_mul_complex = prove(` !p:(V->num)->complex q. ring_powerseries QinC_ring p ==> ring_powerseries QinC_ring q ==> poly_mul QinC_ring p q = poly_mul complex_ring p q `, intro THEN rw[poly_mul;FUN_EQ_THM] THEN intro THEN subgoal `ring_sum QinC_ring {m1,m2 | !y. monomial_mul m1 m2 y = x y} (\(m1,m2). ring_mul QinC_ring (p(m1:V->num)) (q m2)) = ring_sum QinC_ring {m1,m2 | !y. monomial_mul m1 m2 y = x y} (\(m1,m2). ring_mul complex_ring (p m1) (q m2))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[LAMBDA_PAIR;BETA_THM] THEN qed[QinC_ring_clauses;complex_ring_clauses] ; pass ] THEN simp[] THEN sufficesby sum_QinC_eq_sum_complex THEN rw[IN_ELIM_THM;LAMBDA_PAIR] THEN intro THEN simp[FST;SND] THEN have `p(m1:V->num) IN ring_carrier(QinC_ring)` [ring_powerseries] THEN have `q(m2:V->num) IN ring_carrier(QinC_ring)` [ring_powerseries] THEN qed[RING_MUL;QinC_ring_clauses;complex_ring_clauses] );; let ring_divides_poly_complex_if_ring_divides_poly_QinC = prove(` !p:(1->num)->complex q. ring_divides(x_poly QinC_ring) p q ==> ring_divides(x_poly complex_ring) p q `, rw[ring_divides] THEN rw[GSYM x_poly_use] THEN intro THENL [ qed[poly_complex_if_poly_QinC] ; qed[poly_complex_if_poly_QinC] ; witness `x:(1->num)->complex` THEN qed[ring_polynomial;poly_complex_if_poly_QinC;poly_mul_QinC_eq_poly_mul_complex] ] );; (* ===== \Q[x] *) let PID_x_poly_QinC = prove(` PID (x_poly QinC_ring) `, qed[field_QinC;PID_x_poly_field] );; let integral_domain_x_poly_QinC = prove(` integral_domain (x_poly QinC_ring) `, qed[PID_x_poly_QinC;PID_IMP_INTEGRAL_DOMAIN;x_poly] );; let UFD_x_poly_QinC = prove(` UFD (x_poly QinC_ring) `, qed[PID_x_poly_QinC;PID_IMP_UFD;x_poly] );; let x_poly_mul_in_QinC_eq_0 = prove(` !(p:(1->num)->complex) (q:(1->num)->complex). ring_polynomial QinC_ring p ==> ring_polynomial QinC_ring q ==> (poly_mul QinC_ring p q = poly_0 QinC_ring <=> p = poly_0 QinC_ring \/ q = poly_0 QinC_ring) `, rw[x_poly_use] THEN qed[integral_domain_x_poly_QinC;INTEGRAL_DOMAIN_MUL_EQ_0] );; let x_poly_ZinC_denominator_is_QinC = prove(` !(p:(1->num)->complex) e. ring_polynomial ZinC_ring p ==> ~(e = 0) ==> ring_polynomial QinC_ring (series_from_coeffs (\n. (coeff n p) / Cx(&e))) `, rw[poly_series_from_coeffs] THEN intro THENL [ have `coeff d p IN ring_carrier ZinC_ring` [coeff_poly_in_ring] THEN have `coeff d p IN ZinC` [ZinC_ring_clauses] THEN have `coeff d p IN QinC` [ZinC_in_QinC] THEN rw[QinC_ring_clauses] THEN qed[QinC_over_num] ; specialize[`ZinC_ring`;`p:(1->num)->complex`]finite_coeff THEN subgoal `!d. ~(coeff d p / Cx(&e) = ring_0 QinC_ring) ==> ~(coeff d p = ring_0 ZinC_ring)` THENL [ simp[QinC_ring_clauses;ZinC_ring_clauses] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN set_fact `(!d. ~(coeff d p / Cx(&e) = ring_0 QinC_ring) ==> ~(coeff d p = ring_0 ZinC_ring)) ==> {d | ~(coeff d p / Cx(&e) = ring_0 QinC_ring)} SUBSET {d | ~(coeff d p = ring_0 ZinC_ring)}` THEN qed[FINITE_SUBSET] ] );; let x_poly_ZinC_denominator_is_QinC_v2_lemma = prove(` !(p:(1->num)->complex) e. series_from_coeffs (\n. (coeff n p) / Cx(&e)) = poly_mul complex_ring (poly_const complex_ring (Cx(&1) / Cx(&e))) p `, intro THEN have `ring_powerseries complex_ring (p:(1->num)->complex)` [series_complex] THEN sufficesby eq_coeff THEN intro THEN rw[coeff_series_from_coeffs] THEN have `Cx(&1) / Cx(&e) IN ring_carrier complex_ring` [in_complex_ring] THEN simp[coeff_poly_const_times] THEN rw[complex_ring_clauses] THEN CONV_TAC COMPLEX_FIELD );; let poly_const_QinC = prove(` !c. c IN QinC <=> ring_polynomial QinC_ring (poly_const QinC_ring c:(V->num)->complex) `, rw[RING_POLYNOMIAL_CONST] THEN qed[QinC_ring_clauses] );; let poly_const_QinC_poly_const_complex = prove(` !c. poly_const QinC_ring c = poly_const complex_ring c:(V->num)->complex `, qed[poly_const_subring;subring_complex_QinC] );; let poly_mul_QinC_poly_mul_complex = prove(` !(p:(1->num)->complex) (q:(1->num)->complex). ring_powerseries QinC_ring p ==> ring_powerseries QinC_ring q ==> poly_mul QinC_ring p q = poly_mul complex_ring p q `, qed[poly_mul_subring;subring_complex_QinC] );; let x_poly_ZinC_denominator_is_QinC_v2_lemma2 = prove(` !(p:(1->num)->complex) e. ring_powerseries QinC_ring p ==> series_from_coeffs (\n. (coeff n p) / Cx(&e)) = poly_mul QinC_ring (poly_const QinC_ring (Cx(&1) / Cx(&e))) p `, intro THEN have `Cx(&1) / Cx(&e) IN QinC` [num_over_num_in_QinC] THEN have `ring_powerseries QinC_ring (poly_const QinC_ring (Cx(&1) / Cx(&e)):(1->num)->complex)` [RING_POWERSERIES_CONST;QinC_ring_clauses] THEN simp[poly_mul_QinC_poly_mul_complex] THEN simp[poly_const_QinC_poly_const_complex] THEN qed[x_poly_ZinC_denominator_is_QinC_v2_lemma] );; let x_poly_ZinC_denominator_is_QinC_v2 = prove(` !(p:(1->num)->complex) e. ring_polynomial ZinC_ring p ==> ~(e = 0) ==> ring_polynomial QinC_ring (poly_mul QinC_ring (poly_const QinC_ring (Cx(&1) / Cx(&e))) p) `, intro THEN have `ring_powerseries ZinC_ring (p:(1->num)->complex)` [ring_polynomial] THEN have `ring_powerseries QinC_ring (p:(1->num)->complex)` [series_QinC_if_series_ZinC] THEN qed[x_poly_ZinC_denominator_is_QinC;x_poly_ZinC_denominator_is_QinC_v2_lemma2] );; (* XXX: should factor proof through multi_QinC_to_ZinC *) let x_poly_QinC_is_ZinC_denominator = prove(` !p:(1->num)->complex. ring_polynomial QinC_ring p ==> ?e. (~(e = 0) /\ ring_polynomial ZinC_ring (series_from_coeffs (\n. Cx(&e) * coeff n p)) ) `, intro THEN witness `nproduct (0..poly_deg QinC_ring p) (\n. @d. ~(d = 0) /\ Cx(&d) * coeff n p IN ZinC)` THEN conjunction THENL [ rw[NPRODUCT_EQ_0_NUMSEG] THEN qed[coeff_poly_in_ring;QinC_ring_clauses;QinC_to_ZinC] ; rw[poly_series_from_coeffs] THEN intro THENL [ case `d <= poly_deg QinC_ring(p:(1->num)->complex)` THENL [ have `d IN (0..poly_deg QinC_ring(p:(1->num)->complex))` [IN_NUMSEG_0] THEN have `FINITE (0..poly_deg QinC_ring(p:(1->num)->complex))` [FINITE_NUMSEG] THEN specialize[`\n. @d. ~(d = 0) /\ Cx (&d) * coeff n p IN ZinC`;`0..poly_deg QinC_ring(p:(1->num)->complex)`;`d:num`](GSYM NPRODUCT_DELETE) THEN simp[] THEN once_rw[MULT_SYM] THEN simp[prove(`&(m*n):real = &m * &n`,qed[REAL_OF_NUM_CLAUSES])] THEN rw[CX_MUL;GSYM COMPLEX_MUL_ASSOC] THEN have `(Cx (&(@u:num. ~(u = 0) /\ Cx (&u) * coeff d p IN ZinC))) * coeff d p IN ZinC` [coeff_poly_in_ring;QinC_ring_clauses;QinC_to_ZinC] THEN qed[mul_in_ZinC;ZinC_ring_clauses;num_in_ZinC] ; have `coeff d p = ring_0 QinC_ring` [coeff_le_deg] THEN have `ring_0 QinC_ring = Cx(&0)` [QinC_ring_clauses] THEN simp[COMPLEX_MUL_RZERO;ZinC_ring_clauses] THEN qed[ZinC_0] ] ; specialize[`QinC_ring`;`p:(1->num)->complex`]finite_coeff THEN subgoal `{d | ~(Cx (&(nproduct (0..poly_deg QinC_ring p) (\n. @d. ~(d = 0) /\ Cx (&d) * coeff n p IN ZinC))) * coeff d p = ring_0 ZinC_ring)} SUBSET {d | ~(coeff d p = ring_0 QinC_ring)}` THENL [ rw[SUBSET;IN_ELIM_THM] THEN rw[ZinC_ring_clauses;QinC_ring_clauses] THEN qed[COMPLEX_MUL_RZERO] ; qed[FINITE_SUBSET] ] ] ] );; let x_poly_QinC_is_ZinC_denominator_v2_lemma = prove(` !(p:(1->num)->complex) e. series_from_coeffs (\n. Cx(&e) * coeff n p) = poly_mul complex_ring (poly_const complex_ring (Cx(&e))) p `, intro THEN have `ring_powerseries complex_ring (p:(1->num)->complex)` [series_complex] THEN sufficesby eq_coeff THEN intro THEN rw[coeff_series_from_coeffs] THEN have `Cx(&e) IN ring_carrier complex_ring` [in_complex_ring] THEN simp[coeff_poly_const_times] THEN rw[complex_ring_clauses] );; let x_poly_QinC_is_ZinC_denominator_v2_lemma2 = prove(` !(p:(1->num)->complex) e. ring_powerseries QinC_ring p ==> series_from_coeffs (\n. Cx(&e) * coeff n p) = poly_mul QinC_ring (poly_const QinC_ring (Cx(&e))) p `, intro THEN have `Cx(&e) IN QinC` [num_in_QinC] THEN have `ring_powerseries QinC_ring (poly_const QinC_ring (Cx(&e)):(1->num)->complex)` [RING_POWERSERIES_CONST;QinC_ring_clauses] THEN simp[poly_mul_QinC_poly_mul_complex] THEN simp[poly_const_QinC_poly_const_complex] THEN qed[x_poly_QinC_is_ZinC_denominator_v2_lemma] );; let x_poly_QinC_is_ZinC_denominator_v2 = prove(` !p:(1->num)->complex. ring_polynomial QinC_ring p ==> ?e. (~(e = 0) /\ ring_polynomial ZinC_ring (poly_mul QinC_ring (poly_const QinC_ring (Cx(&e))) p) ) `, intro THEN have `ring_powerseries QinC_ring (p:(1->num)->complex)` [ring_polynomial] THEN simp[GSYM x_poly_QinC_is_ZinC_denominator_v2_lemma2] THEN qed[x_poly_QinC_is_ZinC_denominator] );; let ring_hasQ_subring_series_complex = prove(` ring_hasQ ( subring_generated ( x_series complex_ring ) ( ring_carrier (x_series QinC_ring) ) ) `, rw[ring_hasQ] THEN intro THENL [ rw[RING_CHAR_SUBRING_GENERATED] THEN rw[ring_char_x_series] THEN rw[ring_char_complex] ; rw[ring_unit] THEN intro THENL [ qed[RING_OF_NUM] ; witness `poly_const QinC_ring (Cx(&1) / Cx(&n)):(1->num)->complex` THEN intro THENL [ have `Cx(&1) / Cx(&n) IN QinC` [num_over_num_in_QinC] THEN have `Cx(&1) / Cx(&n) IN ring_carrier(QinC_ring)` [QinC_ring_clauses] THEN have `poly_const QinC_ring (Cx(&1) / Cx(&n)) IN ring_carrier(x_series QinC_ring)` [x_series_use;RING_POWERSERIES_CONST] THEN have `ring_carrier(x_series QinC_ring) SUBSET ring_carrier(x_series complex_ring)` [x_series_use;series_complex_if_series_QinC;SUBSET] THEN have `ring_carrier(x_series QinC_ring) SUBSET ring_carrier (subring_generated (x_series complex_ring) (ring_carrier(x_series QinC_ring)))` [SUBSET_CARRIER_SUBRING_GENERATED;SUBSET_REFL] THEN qed[SUBSET] ; rw[RING_OF_NUM_SUBRING_GENERATED] THEN rw[SUBRING_GENERATED] THEN rw[ring_of_num_x_series] THEN rw[GSYM x_series_use] THEN rw[poly_const_QinC_poly_const_complex] THEN rw[ring_of_num_complex;complex_of_num] THEN simp[GSYM POLY_CONST_MUL;in_complex_ring] THEN rw[complex_ring_clauses;poly_1] THEN have `~(&n = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&n) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx(&n) = Cx(&0)) ==> Cx (&n) * Cx (&1) / Cx (&n) = Cx(&1)` THEN qed[] ] ] ] );; (* ===== algebraic numbers *) let algebraic_number = new_definition ` algebraic_number (z:complex) <=> ?p. (ring_polynomial ZinC_ring p /\ ~(p = poly_0 ZinC_ring) /\ poly_eval complex_ring p z = Cx(&0) ) `;; let algebraic_number_ZinC_explicit = prove(` !z. z IN ZinC ==> ( ring_polynomial ZinC_ring (x_minus_const ZinC_ring z) /\ ~((x_minus_const ZinC_ring z) = poly_0 ZinC_ring) /\ poly_eval complex_ring (x_minus_const ZinC_ring z) z = Cx(&0) ) `, intro_gendisch THEN have `z IN ring_carrier(ZinC_ring)` [ZinC_ring_clauses] THEN intro THENL [ qed[x_minus_const_poly] ; have `coeff 1 (x_minus_const ZinC_ring z) = ring_1 ZinC_ring` [coeff_x_minus_const] THEN have `coeff 1 (poly_0 ZinC_ring) = ring_0 ZinC_ring` [coeff_poly_0] THEN have `ring_1 ZinC_ring = Cx(&1)` [ZinC_ring_clauses] THEN have `ring_0 ZinC_ring = Cx(&0)` [ZinC_ring_clauses] THEN have `Cx(&1) = Cx(&0)` [] THEN qed[CX_INJ;REAL_ARITH `~(&1 = &0:real)`] ; have `ZinC subring_of complex_ring` [ZinC_subring_complex] THEN have `z IN ring_carrier(ZinC_ring)` [ZinC_ring_clauses] THEN have `ring_powerseries(subring_generated complex_ring ZinC) (x_minus_const ZinC_ring z)` [x_minus_const_series;subring_complex_ZinC] THEN specialize[`complex_ring`;`ZinC`;`x_minus_const ZinC_ring z`;`z:complex`]poly_eval_subring THEN simp[] THEN qed[eval_x_minus_const_refl;subring_complex_ZinC;ZinC_0;ZinC_ring_clauses] ] );; let algebraic_number_ZinC = prove(` !z. z IN ZinC ==> algebraic_number z `, rw[algebraic_number] THEN qed[algebraic_number_ZinC_explicit] );; let algebraic_number_ii = prove(` algebraic_number ii `, rw[algebraic_number] THEN intro THEN witness `poly_add ZinC_ring (x_pow ZinC_ring 2) (x_pow ZinC_ring 0)` THEN intro THENL [ qed[x_pow_poly;RING_POLYNOMIAL_ADD] ; have `coeff 0 (x_pow ZinC_ring 2) = ring_0 ZinC_ring` [coeff_x_pow;ARITH_RULE `~(2 = 0)`] THEN have `coeff 0 (x_pow ZinC_ring 0) = ring_1 ZinC_ring` [coeff_x_pow] THEN have `coeff 0 (poly_add ZinC_ring (x_pow ZinC_ring 2) (x_pow ZinC_ring 0)) = ring_add ZinC_ring (ring_0 ZinC_ring) (ring_1 ZinC_ring)` [coeff_poly_add] THEN have `coeff 0 (poly_add ZinC_ring (x_pow ZinC_ring 2) (x_pow ZinC_ring 0)) = ring_1 ZinC_ring` [RING_ADD_LZERO;RING_1] THEN have `coeff 0 (poly_0 ZinC_ring) = ring_0 ZinC_ring` [coeff_poly_0] THEN have `ring_1 ZinC_ring = Cx(&1)` [ZinC_ring_clauses] THEN have `ring_0 ZinC_ring = Cx(&0)` [ZinC_ring_clauses] THEN have `Cx(&1) = Cx(&0)` [] THEN qed[CX_INJ;REAL_ARITH `~(&1 = &0:real)`] ; pass ] THEN have `x_pow ZinC_ring 2 = x_pow complex_ring 2` [subring_x_pow;subring_complex_empty] THEN have `x_pow ZinC_ring 0 = x_pow complex_ring 0` [subring_x_pow;subring_complex_empty] THEN simp[poly_add_subring;GSYM subring_complex_empty] THEN have `ii IN ring_carrier complex_ring` [in_complex_ring] THEN have `poly_deg complex_ring (x_pow complex_ring 2) <= 2` [deg_x_pow_le] THEN have `poly_deg complex_ring (x_pow complex_ring 0) <= 2` [deg_x_pow_le;ARITH_RULE `d <= 0 ==> d <= 2`] THEN have `ring_polynomial complex_ring (x_pow complex_ring 2)` [x_pow_poly] THEN have `ring_polynomial complex_ring (x_pow complex_ring 0)` [x_pow_poly] THEN have `ring_polynomial complex_ring (poly_add complex_ring (x_pow complex_ring 2) (x_pow complex_ring 0))` [RING_POLYNOMIAL_ADD] THEN have `poly_deg complex_ring (poly_add complex_ring (x_pow complex_ring 2) (x_pow complex_ring 0)) <= MAX (poly_deg complex_ring (x_pow complex_ring 2)) (poly_deg complex_ring (x_pow complex_ring 0))` [POLY_DEG_ADD_LE] THEN have `poly_deg complex_ring (poly_add complex_ring (x_pow complex_ring 2) (x_pow complex_ring 0)) <= 2` [ARITH_RULE `d <= 2 /\ e <= 2 /\ c <= MAX d e ==> c <= 2`] THEN specialize[`complex_ring`;`poly_add complex_ring (x_pow complex_ring 2) (x_pow complex_ring 0)`;`ii`;`2`]poly_eval_expand_coeff THEN simp[] THEN rw[coeff_poly_add;coeff_x_pow] THEN simp[RING_SUM_CLAUSES_RIGHT; ARITH_RULE `0 < 2 /\ 0 <= 2`;ARITH_RULE `2 - 1 = 1`;ARITH_RULE `~(2 = 0)`; ARITH_RULE `0 < 1 /\ 0 <= 1`;ARITH_RULE `1 - 1 = 0`;ARITH_RULE `~(1 = 2)`;ARITH_RULE `~(1 = 0)`; RING_CLAUSES] THEN once_rw[RING_SUM_CLAUSES_NUMSEG] THEN simp[ARITH_RULE `~(0 = 2)`;ARITH_RULE `~(0 = 1)`; RING_ADD_LZERO;RING_ADD_RZERO;RING_MUL_LID;RING_MUL_LZERO; RING_CLAUSES] THEN rw[ring_pow_complex;complex_ring_clauses] THEN simp[COMPLEX_POW_II_2;complex_pow] THEN CONV_TAC COMPLEX_FIELD );; let algebraic_number_half = prove(` algebraic_number (Cx(&1 / &2)) `, rw[algebraic_number] THEN intro THEN witness `poly_sub ZinC_ring (const_x_pow ZinC_ring (Cx(&2)) 1) (x_pow ZinC_ring 0)` THEN intro THENL [ have `Cx(&2) IN ring_carrier ZinC_ring` [complex_of_num;RING_OF_NUM;ring_of_num_ZinC] THEN qed[x_pow_poly;const_x_pow_poly;RING_POLYNOMIAL_SUB] ; have `coeff 0 (const_x_pow ZinC_ring (Cx(&2)) 1) = ring_0 ZinC_ring` [coeff_const_x_pow;ARITH_RULE `~(1 = 0)`] THEN have `coeff 0 (x_pow ZinC_ring 0) = ring_1 ZinC_ring` [coeff_x_pow] THEN have `coeff 0 (poly_sub ZinC_ring (const_x_pow ZinC_ring (Cx(&2)) 1) (x_pow ZinC_ring 0)) = ring_sub ZinC_ring (ring_0 ZinC_ring) (ring_1 ZinC_ring)` [coeff_poly_sub] THEN have `coeff 0 (poly_sub ZinC_ring (const_x_pow ZinC_ring (Cx(&2)) 1) (x_pow ZinC_ring 0)) = ring_neg ZinC_ring (ring_1 ZinC_ring)` [RING_SUB_LZERO;RING_1] THEN have `coeff 0 (poly_0 ZinC_ring) = ring_0 ZinC_ring` [coeff_poly_0] THEN have `ring_1 ZinC_ring = Cx(&1)` [ZinC_ring_clauses] THEN have `ring_0 ZinC_ring = Cx(&0)` [ZinC_ring_clauses] THEN have `ring_neg ZinC_ring = ( -- )` [ZinC_ring_clauses] THEN have `-- Cx(&1) = Cx(&0)` [] THEN have `Cx(&0) + Cx(&1) = Cx(&0)` [COMPLEX_LNEG_UNIQ] THEN have `Cx(&0 + &1) = Cx(&0)` [CX_ADD] THEN qed[CX_INJ;REAL_ARITH `~(&0 + &1 = &0:real)`] ; pass ] THEN have `const_x_pow ZinC_ring (Cx(&2)) 1 = const_x_pow complex_ring (Cx(&2)) 1` [subring_const_x_pow;subring_complex_empty] THEN have `x_pow ZinC_ring 0 = x_pow complex_ring 0` [subring_x_pow;subring_complex_empty] THEN simp[poly_sub_subring;GSYM subring_complex_empty] THEN have `Cx(&1 / &2) IN ring_carrier complex_ring` [in_complex_ring] THEN have `Cx(&2) IN ring_carrier complex_ring` [in_complex_ring] THEN have `poly_deg complex_ring (const_x_pow complex_ring (Cx(&2)) 1) <= 1` [deg_const_x_pow_le] THEN have `poly_deg complex_ring (x_pow complex_ring 0) <= 1` [deg_x_pow_le;ARITH_RULE `d <= 0 ==> d <= 1`] THEN have `ring_polynomial complex_ring (const_x_pow complex_ring (Cx(&2)) 1)` [const_x_pow_poly] THEN have `ring_polynomial complex_ring (x_pow complex_ring 0)` [x_pow_poly] THEN have `ring_polynomial complex_ring (poly_sub complex_ring (const_x_pow complex_ring (Cx(&2)) 1) (x_pow complex_ring 0))` [RING_POLYNOMIAL_SUB] THEN have `poly_deg complex_ring (poly_sub complex_ring (const_x_pow complex_ring (Cx(&2)) 1) (x_pow complex_ring 0)) <= 1` [POLY_DEG_SUB_LE;ARITH_RULE `d <= 1 /\ e <= 1 /\ c <= MAX d e ==> c <= 1`] THEN specialize[`complex_ring`;`poly_sub complex_ring (const_x_pow complex_ring (Cx(&2)) 1) (x_pow complex_ring 0)`;`Cx(&1 / &2)`;`1`]poly_eval_expand_coeff THEN simp[] THEN rw[coeff_poly_sub;coeff_const_x_pow;coeff_x_pow] THEN simp[RING_SUM_CLAUSES_RIGHT; ARITH_RULE `0 < 1 /\ 0 <= 1`;ARITH_RULE `1 - 1 = 0`;ARITH_RULE `~(1 = 0)`; RING_CLAUSES] THEN once_rw[RING_SUM_CLAUSES_NUMSEG] THEN simp[ARITH_RULE `~(0 = 1)`; RING_ADD_LZERO;RING_ADD_RZERO;RING_MUL_LID;RING_MUL_LZERO; RING_CLAUSES] THEN rw[ring_pow_complex;complex_ring_clauses;ring_sub_complex] THEN simp[complex_pow;COMPLEX_POW_1] THEN CONV_TAC COMPLEX_FIELD );; let algebraic_number_if_monic_vanishing_at = prove(` !S (c:X->complex) s. FINITE S ==> ring_polynomial ZinC_ring (monic_vanishing_at complex_ring S c) ==> s IN S ==> algebraic_number (c s) `, rw[algebraic_number] THEN intro THEN witness `monic_vanishing_at complex_ring S (c:X->complex)` THEN intro THENL [ qed[] ; have `!s:X. s IN S ==> c s IN ring_carrier complex_ring` [in_complex_ring] THEN have `monic complex_ring (monic_vanishing_at complex_ring (S:X->bool) c)` [monic_vanishing_at_monic] THEN have `monic complex_ring (poly_0 ZinC_ring)` [] THEN have `monic complex_ring (poly_0 complex_ring)` [poly_0_ZinC_eq_poly_0_complex] THEN have `ring_1 complex_ring = ring_0 complex_ring` [monic_poly_0] THEN qed[field_complex;field] ; have `!s:X. s IN S ==> c s IN ring_carrier complex_ring` [in_complex_ring] THEN simp[eval_monic_vanishing_at_refl] THEN rw[complex_ring_clauses] ] );; (* ===== algebraic numbers via \Q[x] *) let algebraic_number_root_QinC_poly = prove(` !z:complex. algebraic_number z <=> ?p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ poly_eval complex_ring p z = Cx(&0) ) `, rw[algebraic_number] THEN intro THEN splitiff THENL [ intro THEN witness `p:(1->num)->complex` THEN qed[poly_QinC_if_poly_ZinC;poly_0_ZinC_eq_poly_0_QinC] ; intro THEN choose `e:num` `~(e = 0) /\ ring_polynomial ZinC_ring (poly_mul QinC_ring (poly_const QinC_ring (Cx (&e))) (p:(1->num)->complex))` [x_poly_QinC_is_ZinC_denominator_v2] THEN witness `poly_mul QinC_ring (poly_const QinC_ring (Cx (&e))) (p:(1->num)->complex)` THEN rw[poly_0_ZinC_eq_poly_0_QinC] THEN intro THENL [ simp[] ; have `Cx(&e) IN QinC` [num_in_QinC] THEN have `ring_polynomial QinC_ring (poly_const QinC_ring (Cx(&e)):(1->num)->complex)` [poly_const_QinC] THEN have `(poly_const QinC_ring (Cx(&e)):(1->num)->complex) = poly_0 QinC_ring` [x_poly_mul_in_QinC_eq_0] THEN have `Cx(&e) = ring_0 QinC_ring` [poly_0;POLY_CONST_EQ] THEN have `Cx(&e) = Cx(&0)` [QinC_ring_clauses] THEN have `&e = &0:real` [CX_INJ] THEN qed[REAL_OF_NUM_EQ] ; have `Cx(&e) IN QinC` [num_in_QinC] THEN have `ring_polynomial QinC_ring (poly_const QinC_ring (Cx(&e)):(1->num)->complex)` [poly_const_QinC] THEN have `ring_powerseries QinC_ring (poly_const QinC_ring (Cx(&e)):(1->num)->complex)` [ring_polynomial] THEN have `ring_powerseries QinC_ring (p:(1->num)->complex)` [ring_polynomial] THEN simp[poly_mul_QinC_poly_mul_complex] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (poly_const QinC_ring (Cx(&e)):(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN simp[POLY_EVAL_MUL;complex_ring_clauses] THEN qed[COMPLEX_MUL_RZERO;POLY_EVAL] ] ] );; let algebraic_number_if_monic_vanishing_at_QinC = prove(` !S (c:X->complex) s. FINITE S ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring S c) ==> s IN S ==> algebraic_number (c s) `, rw[algebraic_number_root_QinC_poly] THEN intro THEN witness `monic_vanishing_at complex_ring S (c:X->complex)` THEN intro THENL [ qed[] ; have `!s:X. s IN S ==> c s IN ring_carrier complex_ring` [in_complex_ring] THEN have `monic complex_ring (monic_vanishing_at complex_ring (S:X->bool) c)` [monic_vanishing_at_monic] THEN have `monic complex_ring (poly_0 QinC_ring)` [] THEN have `monic complex_ring (poly_0 complex_ring)` [poly_0_QinC_eq_poly_0_complex] THEN have `ring_1 complex_ring = ring_0 complex_ring` [monic_poly_0] THEN qed[field_complex;field] ; have `!s:X. s IN S ==> c s IN ring_carrier complex_ring` [in_complex_ring] THEN simp[eval_monic_vanishing_at_refl] THEN rw[complex_ring_clauses] ] );; let algebraic_number_if_root_irreducible_QinC_poly = prove(` !z:complex. (?p. (ring_polynomial QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) ) ) ==> algebraic_number z `, rw[algebraic_number_root_QinC_poly] THEN intro THEN witness `p:(1->num)->complex` THEN have `~(p = ring_0(x_poly QinC_ring))` [ring_irreducible] THEN qed[x_poly_use] );; let algebraic_number_is_root_irreducible_QinC_poly = prove(` !z:complex. algebraic_number z ==> ?p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) ) `, rw[algebraic_number_root_QinC_poly] THEN intro THEN have `UFD(x_poly QinC_ring)` [UFD_x_poly_QinC] THEN have `p IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `~(p = ring_0(x_poly QinC_ring))` [x_poly_use] THEN choose2 `n:num` `q:num->(1->num)->complex` `(!i. 1 <= i /\ i <= n ==> ring_prime(x_poly QinC_ring) (q i)) /\ ring_associates(x_poly QinC_ring) (ring_product(x_poly QinC_ring) (1..n) q) p` [UFD_EQ_PRIMEFACT] THEN def `qprod:(1->num)->complex` `(ring_product(x_poly QinC_ring) (1..n) q)` THEN have `ring_divides(x_poly QinC_ring) p qprod` [ring_associates] THEN choose `c:(1->num)->complex` `c IN ring_carrier(x_poly QinC_ring) /\ qprod = ring_mul(x_poly QinC_ring) p c` [ring_divides] THEN have `(qprod:(1->num)->complex) = poly_mul QinC_ring p c` [x_poly;POLY_RING_CLAUSES] THEN have `ring_powerseries QinC_ring (p:(1->num)->complex)` [ring_polynomial] THEN have `p IN ring_carrier(x_poly QinC_ring)` [ring_divides] THEN have `qprod IN ring_carrier(x_poly QinC_ring)` [RING_MUL] THEN have `ring_polynomial QinC_ring (qprod:(1->num)->complex)` [x_poly_use] THEN have `ring_powerseries QinC_ring (qprod:(1->num)->complex)` [ring_polynomial] THEN have `c IN ring_carrier(x_poly QinC_ring)` [ring_divides] THEN have `ring_polynomial QinC_ring (c:(1->num)->complex)` [x_poly_use] THEN have `ring_powerseries QinC_ring (c:(1->num)->complex)` [ring_polynomial] THEN have `(qprod:(1->num)->complex) = poly_mul complex_ring p c` [poly_mul_QinC_poly_mul_complex] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (c:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `poly_eval complex_ring qprod z = ring_mul complex_ring (poly_eval complex_ring p z) (poly_eval complex_ring c z)` [POLY_EVAL_MUL] THEN have `poly_eval complex_ring qprod z = ring_mul complex_ring (ring_0 complex_ring) (poly_eval complex_ring c z)` [complex_ring_clauses] THEN have `poly_eval complex_ring qprod z = ring_0 complex_ring` [RING_MUL_LZERO;POLY_EVAL] THEN have `poly_eval complex_ring qprod z = Cx(&0)` [complex_ring_clauses] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!i. i IN 1..n ==> 1 <= i /\ i <= n` [IN_NUMSEG] THEN have `!i. i IN 1..n ==> q i IN ring_carrier(x_poly QinC_ring)` [ring_prime] THEN have `!i. i IN 1..n ==> ring_polynomial QinC_ring (q i:(1->num)->complex)` [x_poly_use] THEN have `!i. i IN 1..n ==> ring_polynomial complex_ring (q i:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `!i. i IN 1..n ==> ring_powerseries QinC_ring (q i:(1->num)->complex)` [ring_polynomial] THEN have `!i. i IN 1..n ==> ring_powerseries (subring_generated complex_ring QinC) (q i:(1->num)->complex)` [subring_complex_QinC] THEN specialize[`complex_ring`;`q:num->(1->num)->complex`;`z:complex`;`1..n`]eval_poly_product THEN have `qprod = poly_product QinC_ring (1..n) q` [poly_product_ring_product_x_poly] THEN specialize[`complex_ring`;`QinC`;`q:num->(1->num)->complex`;`1..n`]poly_product_subring THEN have `qprod = poly_product complex_ring (1..n) q` [subring_complex_QinC] THEN have `ring_product complex_ring (1..n) (\i. poly_eval complex_ring (q i) z) = ring_0 complex_ring` [] THEN have `integral_domain complex_ring` [integral_domain_complex] THEN specialize[`complex_ring`;`1..n`;`\i. poly_eval complex_ring (q(i:num)) z`]INTEGRAL_DOMAIN_PRODUCT_EQ_0 THEN choose `i:num` `i IN 1..n /\ poly_eval complex_ring (q i) z = ring_0 complex_ring` [] THEN witness `q(i:num):(1->num)->complex` THEN have `integral_domain(x_poly QinC_ring)` [integral_domain_x_poly_QinC] THEN have `ring_irreducible(x_poly QinC_ring) (q(i:num))` [INTEGRAL_DOMAIN_PRIME_IMP_IRREDUCIBLE] THEN have `(q(i:num)) IN ring_carrier(x_poly QinC_ring)` [ring_irreducible] THEN have `ring_polynomial(QinC_ring) (q(i:num):(1->num)->complex)` [x_poly_use] THEN have `~((q(i:num):(1->num)->complex) = ring_0(x_poly QinC_ring))` [ring_irreducible] THEN have `~((q(i:num):(1->num)->complex) = poly_0 QinC_ring)` [x_poly_use] THEN qed[] );; let algebraic_number_root_irreducible_QinC_poly = prove(` !z:complex. algebraic_number z <=> ?p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) ) `, qed[ algebraic_number_if_root_irreducible_QinC_poly; algebraic_number_is_root_irreducible_QinC_poly ] );; let algebraic_number_is_root_monic_irreducible_QinC_poly = prove(` !z:complex. algebraic_number z ==> ?p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ monic QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) ) `, rw[algebraic_number_root_irreducible_QinC_poly] THEN intro THEN have `field QinC_ring` [field_QinC] THEN choose `q:(1->num)->complex` `ring_polynomial QinC_ring q /\ monic QinC_ring (q:(1->num)->complex) /\ ring_associates(x_poly QinC_ring) p q` [x_poly_field_monic_associate] THEN witness `q:(1->num)->complex` THEN have `~((q:(1->num)->complex) = poly_0 QinC_ring)` [monic_poly_0;field] THEN have `integral_domain(x_poly QinC_ring)` [integral_domain_x_poly_QinC] THEN have `ring_irreducible (x_poly QinC_ring) q` [RING_ASSOCIATES_IRREDUCIBLE] THEN choose `w:(1->num)->complex` `w IN ring_carrier(x_poly QinC_ring) /\ q = ring_mul(x_poly QinC_ring) p w` [ring_associates;ring_divides] THEN have `(q:(1->num)->complex) = poly_mul QinC_ring p w` [x_poly_use] THEN have `ring_polynomial QinC_ring (w:(1->num)->complex)` [x_poly_use] THEN have `(q:(1->num)->complex) = poly_mul complex_ring p w` [poly_mul_subring;subring_complex_QinC;ring_polynomial] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (w:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN specialize[`complex_ring`;`p:(1->num)->complex`;`w:(1->num)->complex`;`z:complex`]POLY_EVAL_MUL THEN qed[COMPLEX_MUL_LZERO;complex_ring_clauses] );; let coprimes_sharing_root = prove(` !(r:R ring) s G p q z. s = subring_generated r G ==> field s ==> ring_polynomial s p ==> ring_polynomial s q ==> ~(p = poly_0 s) ==> ~(q = poly_0 s) ==> ring_irreducible(x_poly s) p ==> ring_irreducible(x_poly s) q ==> z IN ring_carrier r ==> poly_eval r p z = ring_0 r ==> poly_eval r q z = ring_0 r ==> ring_associates(x_poly s) p q `, intro THEN def `g:(1->num)->R` `ring_gcd(x_poly(s:R ring)) (p,q)` THEN case `ring_unit(x_poly(s:R ring)) g` THENL [ have `g IN ring_carrier(x_poly(s:R ring))` [RING_GCD] THEN choose `ginv:(1->num)->R` `ginv IN ring_carrier(x_poly(s:R ring)) /\ ring_mul(x_poly s) g ginv = ring_1(x_poly s)` [ring_unit] THEN have `ring_polynomial(s:R ring) (ginv:(1->num)->R)` [x_poly_use] THEN have `ring_polynomial(s:R ring) (g:(1->num)->R)` [x_poly_use] THEN specialize[`s:R ring`;`p:(1->num)->R`;`q:(1->num)->R`]gcd_poly_linear_combination THEN choose2 `u:(1->num)->R` `v:(1->num)->R` `ring_polynomial s u /\ ring_polynomial s v /\ poly_add s (poly_mul s p u) (poly_mul s q v) = g:(1->num)->R` [] THEN have `ring_polynomial(s:R ring) (poly_mul s (p:(1->num)->R) (u:(1->num)->R))` [RING_POLYNOMIAL_MUL] THEN have `ring_polynomial(s:R ring) (poly_mul s (q:(1->num)->R) (v:(1->num)->R))` [RING_POLYNOMIAL_MUL] THEN have `ring_polynomial(r:R ring) (g:(1->num)->R)` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (ginv:(1->num)->R)` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (p:(1->num)->R)` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (q:(1->num)->R)` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (u:(1->num)->R)` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (v:(1->num)->R)` [ring_polynomial_if_subring] THEN have `poly_mul(s:R ring) g ginv:(1->num)->R = poly_1 s` [x_poly_use] THEN have `poly_mul(r:R ring) g ginv:(1->num)->R = poly_1 r` [poly_mul_subring;ring_polynomial;poly_1_subring] THEN specialize[`r:R ring`;`g:(1->num)->R`;`ginv:(1->num)->R`;`z:R`]POLY_EVAL_MUL THEN specialize[`r:R ring`;`p:(1->num)->R`;`u:(1->num)->R`;`z:R`]POLY_EVAL_MUL THEN specialize[`r:R ring`;`q:(1->num)->R`;`v:(1->num)->R`;`z:R`]POLY_EVAL_MUL THEN have `ring_mul(r:R ring) (poly_eval r g z) (poly_eval r ginv z) = ring_1 r` [POLY_EVAL_1] THEN have `poly_eval(r:R ring) (poly_mul r p u) z = ring_0 r` [RING_MUL_LZERO;POLY_EVAL] THEN have `poly_eval(r:R ring) (poly_mul r q v) z = ring_0 r` [RING_MUL_LZERO;POLY_EVAL] THEN have `poly_mul(r:R ring) p u:(1->num)->R = poly_mul s p u` [poly_mul_subring;ring_polynomial] THEN have `poly_mul(r:R ring) q v:(1->num)->R = poly_mul s q v` [poly_mul_subring;ring_polynomial] THEN have `ring_polynomial(r:R ring) (poly_mul r (p:(1->num)->R) (u:(1->num)->R))` [ring_polynomial_if_subring] THEN have `ring_polynomial(r:R ring) (poly_mul r (q:(1->num)->R) (v:(1->num)->R))` [ring_polynomial_if_subring] THEN have `poly_add(r:R ring) (poly_mul r p u) (poly_mul r q v:(1->num)->R) = poly_add s (poly_mul s p u) (poly_mul s q v)` [poly_add_subring;ring_polynomial] THEN have `poly_add(r:R ring) (poly_mul r p u) (poly_mul r q v:(1->num)->R) = g` [] THEN specialize[`r:R ring`;`poly_mul r (p:(1->num)->R) (u:(1->num)->R)`;`poly_mul r (q:(1->num)->R) (v:(1->num)->R)`;`z:R`]POLY_EVAL_ADD THEN have `poly_eval(r:R ring) g z = ring_add r (ring_0 r) (ring_0 r)` [] THEN have `poly_eval(r:R ring) g z = ring_0 r` [RING_ADD_LZERO;RING_0] THEN have `poly_eval(r:R ring) (poly_1 r) z = ring_mul r (ring_0 r) (poly_eval r ginv z)` [] THEN have `ring_1(r:R ring) = ring_0 r` [POLY_EVAL_1;RING_MUL_LZERO;POLY_EVAL] THEN have `ring_1(s:R ring) = ring_0 s` [SUBRING_GENERATED] THEN qed[field] ; pass ] THEN have `ring_divides(x_poly s) (g:(1->num)->R) p` [RING_GCD_DIVIDES;x_poly_use] THEN have `ring_divides(x_poly s) (g:(1->num)->R) q` [RING_GCD_DIVIDES;x_poly_use] THEN have `ring_associates(x_poly s) (g:(1->num)->R) p` [RING_NONUNIT_DIVIDES_IRREDUCIBLE] THEN have `ring_associates(x_poly s) (g:(1->num)->R) q` [RING_NONUNIT_DIVIDES_IRREDUCIBLE] THEN qed[RING_ASSOCIATES_SYM;RING_ASSOCIATES_TRANS] );; let algebraic_number_is_root_unique_monic_irreducible_QinC_poly_lemma = prove(` !(z:complex) p q. ring_polynomial QinC_ring p /\ ring_polynomial QinC_ring q /\ ~(p = poly_0 QinC_ring) /\ ~(q = poly_0 QinC_ring) /\ ring_irreducible(x_poly QinC_ring) p /\ ring_irreducible(x_poly QinC_ring) q /\ poly_eval complex_ring p z = Cx(&0) /\ poly_eval complex_ring q z = Cx(&0) ==> ring_associates(x_poly QinC_ring) p q `, intro THEN have `QinC_ring = subring_generated complex_ring QinC` [subring_complex_QinC] THEN have `field QinC_ring` [field_QinC] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `poly_eval complex_ring p z = ring_0 complex_ring` [complex_ring_clauses] THEN have `poly_eval complex_ring q z = ring_0 complex_ring` [complex_ring_clauses] THEN specialize[`complex_ring`;`QinC_ring`;`QinC`;`p:(1->num)->complex`;`q:(1->num)->complex`;`z:complex`]coprimes_sharing_root THEN qed[] );; let algebraic_number_is_root_unique_monic_irreducible_QinC_poly_simple = prove(` !(z:complex) p q. ring_polynomial QinC_ring p /\ ring_polynomial QinC_ring q /\ ~(p = poly_0 QinC_ring) /\ ~(q = poly_0 QinC_ring) /\ monic QinC_ring p /\ monic QinC_ring q /\ ring_irreducible(x_poly QinC_ring) p /\ ring_irreducible(x_poly QinC_ring) q /\ poly_eval complex_ring p z = Cx(&0) /\ poly_eval complex_ring q z = Cx(&0) ==> p = q `, intro THEN have `ring_associates(x_poly QinC_ring) p q` [algebraic_number_is_root_unique_monic_irreducible_QinC_poly_lemma] THEN qed[monic_associates;field_QinC] );; let algebraic_number_is_root_unique_monic_irreducible_QinC_poly = prove(` !z:complex. algebraic_number z ==> ?!p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ monic QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) ) `, qed[EXISTS_UNIQUE_THM; algebraic_number_is_root_monic_irreducible_QinC_poly; algebraic_number_is_root_unique_monic_irreducible_QinC_poly_simple ] );; let algebraic_number_QinC_explicit = prove(` !z. z IN QinC ==> ( ring_polynomial QinC_ring (x_minus_const QinC_ring z) /\ ~((x_minus_const QinC_ring z) = poly_0 QinC_ring) /\ monic QinC_ring (x_minus_const QinC_ring z) /\ ring_irreducible (x_poly QinC_ring) (x_minus_const QinC_ring z) /\ poly_eval complex_ring (x_minus_const QinC_ring z) z = Cx(&0) ) `, intro_gendisch THEN have `z IN ring_carrier(QinC_ring)` [QinC_ring_clauses] THEN intro THENL [ qed[x_minus_const_poly] ; have `coeff 1 (x_minus_const QinC_ring z) = ring_1 QinC_ring` [coeff_x_minus_const] THEN have `coeff 1 (poly_0 QinC_ring) = ring_0 QinC_ring` [coeff_poly_0] THEN have `ring_1 QinC_ring = Cx(&1)` [QinC_ring_clauses] THEN have `ring_0 QinC_ring = Cx(&0)` [QinC_ring_clauses] THEN have `Cx(&1) = Cx(&0)` [] THEN qed[CX_INJ;REAL_ARITH `~(&1 = &0:real)`] ; qed[monic_x_minus_const] ; qed[irred_x_minus_const;field_QinC] ; have `QinC subring_of complex_ring` [QinC_subring_complex] THEN have `z IN ring_carrier(QinC_ring)` [QinC_ring_clauses] THEN have `ring_powerseries(subring_generated complex_ring QinC) (x_minus_const QinC_ring z)` [x_minus_const_series;subring_complex_QinC] THEN specialize[`complex_ring`;`QinC`;`x_minus_const QinC_ring z`;`z:complex`]poly_eval_subring THEN simp[] THEN qed[eval_x_minus_const_refl;subring_complex_QinC;QinC_0;QinC_ring_clauses] ] );; let algebraic_number_QinC = prove(` !z. z IN QinC ==> algebraic_number z `, rw[algebraic_number_root_QinC_poly] THEN qed[algebraic_number_QinC_explicit] );; (* ===== minimal polynomials *) let minimal_polynomial = new_definition ` minimal_polynomial (z:complex) = if algebraic_number z then (@p. (ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ monic QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p z = Cx(&0) )) else poly_0 QinC_ring `;; let algebraic_has_minimal_polynomial = prove(` !z:complex. algebraic_number z ==> ring_polynomial QinC_ring (minimal_polynomial z) /\ ~((minimal_polynomial z) = poly_0 QinC_ring) /\ monic QinC_ring (minimal_polynomial z) /\ ring_irreducible(x_poly QinC_ring) (minimal_polynomial z) /\ poly_eval complex_ring (minimal_polynomial z) z = Cx(&0) `, qed[minimal_polynomial;algebraic_number_is_root_unique_monic_irreducible_QinC_poly] );; let maybe_algebraic_minimal_polynomial = prove(` !z:complex. ring_polynomial QinC_ring (minimal_polynomial z) /\ poly_eval complex_ring (minimal_polynomial z) z = Cx(&0) `, intro THEN case `algebraic_number z` THENL [ qed[algebraic_has_minimal_polynomial] ; have `minimal_polynomial z = poly_0 QinC_ring` [minimal_polynomial] THEN qed[RING_POLYNOMIAL_0] ; qed[algebraic_has_minimal_polynomial] ; have `minimal_polynomial z = poly_0 QinC_ring` [minimal_polynomial] THEN qed[POLY_EVAL_0;complex_ring_clauses;poly_0_QinC_eq_poly_0_complex] ] );; let minimal_polynomial_QinC = prove(` !z:complex. z IN QinC ==> minimal_polynomial z = x_minus_const QinC_ring z `, intro THEN rw[minimal_polynomial] THEN simp[algebraic_number_QinC] THEN specialize[`z:complex`]algebraic_number_QinC_explicit THEN sufficesby SELECT_UNIQUE THEN qed[algebraic_number_is_root_unique_monic_irreducible_QinC_poly_simple] );; (* ===== using fundamental theorem of algebra *) let poly_eval_vsum_lemma = prove(` !(p:(1->num)->complex) z n. vsum (0..n) (\i. coeff i p * z pow i) = ring_sum complex_ring (0..n) (\d. ring_mul complex_ring (coeff d p) (ring_pow complex_ring z d)) `, intro THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN simp[vsum_ring_sum_complex] THEN sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[complex_ring_clauses;ring_pow_complex] );; let poly_eval_vsum = prove(` !(p:(1->num)->complex) z n. ring_polynomial complex_ring p ==> poly_deg complex_ring p <= n ==> vsum (0..n) (\i. coeff i p * z pow i) = poly_eval complex_ring p z `, have `!y. y IN ring_carrier complex_ring` [in_complex_ring] THEN simp[poly_eval_expand_coeff] THEN qed[poly_eval_vsum_lemma] );; let nonconstant_complex_root = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> 1 <= poly_deg complex_ring p ==> ?z. poly_eval complex_ring p z = Cx(&0) `, intro THEN def `n:num` `poly_deg complex_ring (p:(1->num)->complex)` THEN subgoal `~(!i. i IN 1..n ==> coeff i p = Cx(&0))` THENL [ rw[NOT_FORALL_THM] THEN witness `n:num` THEN have `n IN 1..n` [IN_NUMSEG;ARITH_RULE `n <= n:num`] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [POLY_DEG_0;ARITH_RULE `~(1 <= 0)`] THEN have `~(coeff n p = ring_0 complex_ring)` [topcoeff_nonzero] THEN qed[complex_ring_clauses] ; specialize_assuming[`\i. coeff i p:complex`;`n:num`]FTA THEN choose `z:complex` `vsum(0..n) (\i. coeff i p * z pow i) = Cx(&0)` [] THEN witness `z:complex` THEN qed[poly_eval_vsum;ARITH_RULE `n <= n:num`] ] );; let nonconstant_complex_x_minus_root = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> 1 <= poly_deg complex_ring p ==> ?z q. ring_polynomial complex_ring q /\ p = poly_mul complex_ring (x_minus_const complex_ring z) q `, intro THEN choose `z:complex` `poly_eval complex_ring p z = Cx(&0)` [nonconstant_complex_root] THEN have `poly_eval complex_ring p z = ring_0 complex_ring` [complex_ring_clauses] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `p IN ring_carrier(poly_ring complex_ring (:1))` [x_poly;x_poly_use] THEN specialize[`complex_ring`;`z:complex`;`p:(1->num)->complex`]POLY_DIVIDES_X_MINUS_ROOT THEN have `poly_sub complex_ring (poly_var complex_ring one) (poly_const complex_ring z) = x_minus_const complex_ring z` [x_minus_const;x_pow_1] THEN have `ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) p` [x_poly] THEN choose `q:(1->num)->complex` `q IN ring_carrier(x_poly complex_ring) /\ (p:(1->num)->complex) = ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) q` [ring_divides] THEN witness `z:complex` THEN witness `q:(1->num)->complex` THEN qed[x_poly_use] );; let associates_monic_vanishing_at_if_complex_lemma = prove(` !n (p:(1->num)->complex). ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> poly_deg complex_ring p = n ==> ?c. ring_associates(x_poly complex_ring) p (monic_vanishing_at complex_ring {i:num | i < n} c) `, INDUCT_TAC THENL [ intro THEN witness `\i:num. Cx(&0)` THEN rw[NUMSEG_LT] THEN rw[monic_vanishing_at_empty] THEN choose `c:complex` `c IN ring_carrier complex_ring /\ p = poly_const complex_ring c:(1->num)->complex` [POLY_DEG_EQ_0] THEN have `~(c = ring_0 complex_ring)` [poly_0] THEN have `field complex_ring` [field_complex] THEN have `ring_unit complex_ring c` [FIELD_UNIT] THEN have `ring_unit(x_poly complex_ring) p` [RING_UNIT_POLY_CONST;x_poly] THEN have `ring_associates(x_poly complex_ring) p (ring_1(x_poly complex_ring))` [RING_ASSOCIATES_1] THEN qed[x_poly_use] ; intro THEN num_linear_fact `poly_deg complex_ring (p:(1->num)->complex) = SUC n ==> 1 <= poly_deg complex_ring p` THEN choose2 `z:complex` `q:(1->num)->complex` `ring_polynomial complex_ring q /\ p = poly_mul complex_ring (x_minus_const complex_ring z) q` [nonconstant_complex_x_minus_root] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly] THEN have `integral_domain complex_ring` [integral_domain_complex] THEN have `~(x_minus_const complex_ring z = poly_0 complex_ring)` [x_minus_const_nonzero;integral_domain] THEN have `~(q = poly_0 complex_ring:(1->num)->complex)` [POLY_MUL_0;x_poly_use;x_minus_const_poly] THEN specialize_assuming[`complex_ring`;`x_minus_const complex_ring z`;`q:(1->num)->complex`]POLY_DEG_MUL THEN have `poly_deg complex_ring (p:(1->num)->complex) = poly_deg complex_ring (x_minus_const complex_ring z) + poly_deg complex_ring (q:(1->num)->complex)` [integral_domain_complex;x_poly_use] THEN have `SUC n = 1 + poly_deg complex_ring (q:(1->num)->complex)` [deg_x_minus_const;integral_domain] THEN num_linear_fact `SUC n = 1 + poly_deg complex_ring (q:(1->num)->complex) ==> poly_deg complex_ring q = n` THEN have `x_minus_const complex_ring z IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN choose `d:num->complex` `ring_associates(x_poly complex_ring) q (monic_vanishing_at complex_ring {i:num | i < n} d)` [] THEN def `c:num->complex` `\i:num. if i = n then z:complex else d i` THEN have `c(n:num) IN ring_carrier complex_ring` [in_complex_ring] THEN have `monic_vanishing_at complex_ring {i:num | i < n+1} c = poly_mul complex_ring (x_minus_const complex_ring (c n)) (monic_vanishing_at complex_ring {i:num | i < n} c)` [monic_vanishing_at_plus1] THEN set_fact_using `!i:num. i IN {i | i < n} ==> ~(i = n)` [ARITH_RULE `~(n < n:num)`] THEN have `!i:num. i IN {i | i < n} ==> c i = d i:complex` [] THEN have `monic_vanishing_at complex_ring {i:num | i < n} c = monic_vanishing_at complex_ring {i:num | i < n} d` [monic_vanishing_at_eq] THEN have `ring_associates(x_poly complex_ring) (x_minus_const complex_ring z) (x_minus_const complex_ring z)` [RING_ASSOCIATES_REFL;x_minus_const_poly;x_poly_use] THEN have `ring_associates(x_poly complex_ring) (ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) q) (ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) (monic_vanishing_at complex_ring {i:num | i < n} d))` [RING_ASSOCIATES_MUL] THEN have `ring_associates(x_poly complex_ring) (poly_mul complex_ring (x_minus_const complex_ring z) q) (poly_mul complex_ring (x_minus_const complex_ring z) (monic_vanishing_at complex_ring {i:num | i < n} d))` [x_poly_use] THEN have `ring_associates(x_poly complex_ring) p (poly_mul complex_ring (x_minus_const complex_ring z) (monic_vanishing_at complex_ring {i:num | i < n} d))` [RING_ASSOCIATES_MUL;x_poly_use] THEN have `ring_associates(x_poly complex_ring) p (poly_mul complex_ring (x_minus_const complex_ring z) (monic_vanishing_at complex_ring {i:num | i < n} c))` [] THEN have `ring_associates(x_poly complex_ring) p (monic_vanishing_at complex_ring {i:num | i < n+1} c)` [] THEN witness `c:num->complex` THEN qed[ARITH_RULE `SUC n = n+1`] ] );; let associates_monic_vanishing_at_if_complex = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> ?c. ring_associates(x_poly complex_ring) p (monic_vanishing_at complex_ring {i:num | i < poly_deg complex_ring p} c) `, intro THEN specialize[`poly_deg complex_ring (p:(1->num)->complex)`;`p:(1->num)->complex`]associates_monic_vanishing_at_if_complex_lemma THEN qed[] );; let monic_vanishing_at_if_monic_complex = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> monic complex_ring p ==> ?c. p = monic_vanishing_at complex_ring {i | i < poly_deg complex_ring p} c `, intro THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [monic_poly_0;field_complex;field] THEN specialize[`p:(1->num)->complex`]associates_monic_vanishing_at_if_complex THEN choose `c:num->complex` `ring_associates (x_poly complex_ring) p (monic_vanishing_at complex_ring {i | i < poly_deg complex_ring p} c)` [] THEN have `!i. i IN {i | i < poly_deg complex_ring (p:(1->num)->complex)} ==> c i IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`poly_deg complex_ring (p:(1->num)->complex)`]FINITE_NUMSEG_LT THEN have `monic complex_ring (monic_vanishing_at complex_ring {i | i < poly_deg complex_ring (p:(1->num)->complex)} c)` [monic_vanishing_at_monic] THEN qed[monic_associates;field_complex] );; let monic_QinC_squarefree_complex_squarefree_lemma = prove(` !(r:R ring) C p qpd qp'd D:(1->num)->R. ring_polynomial r C ==> ring_polynomial r p ==> ring_polynomial r qpd ==> ring_polynomial r qp'd ==> ring_polynomial r D ==> p = poly_mul r C qpd ==> poly_mul r C qp'd = poly_add r qpd (poly_mul r C D) ==> p = poly_mul r (poly_mul r C C) (poly_sub r qp'd D) `, rw[x_poly_use;x_poly_sub_use] THEN intro THEN specialize[`x_poly(r:R ring)` ;`C:(1->num)->R` ;`D:(1->num)->R` ;`p:(1->num)->R` ;`qp'd:(1->num)->R` ;`qpd:(1->num)->R`]( GENL[`r:R ring`;`C:R`;`D:R`;`p:R`;`qp'd:R`;`qpd:R`]( RING_RULE ` p = ring_mul(r:R ring) C qpd ==> ring_mul r C qp'd = ring_add r qpd (ring_mul r C D) ==> p = ring_mul r (ring_mul r C C) (ring_sub r qp'd D) ` )) THEN qed[] );; (* XXX: should write this for more general field extensions *) (* XXX: and should do more elementary proof via gcd computation being preserved by field extensions *) (* XXX: e.g., formulate via resultants *) let monic_QinC_squarefree_complex_squarefree = prove(` !p:(1->num)->complex. ring_polynomial QinC_ring p ==> monic QinC_ring p ==> (ring_squarefree(x_poly QinC_ring) p <=> ring_squarefree(x_poly complex_ring) p) `, intro THEN recall subring_complex_QinC THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring] THEN splitiff THENL [ intro THEN recall field_QinC THEN recall ring_char_QinC THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;field] THEN specialize[`QinC_ring`;`p:(1->num)->complex`]coprime_derivative_if_squarefree THEN specialize[`QinC_ring`;`p:(1->num)->complex`;`x_derivative QinC_ring p`]linear_combination_if_coprime_poly THEN choose2 `x:(1->num)->complex` `y:(1->num)->complex` `ring_polynomial QinC_ring x /\ ring_polynomial QinC_ring y /\ poly_add QinC_ring (poly_mul QinC_ring p x) (poly_mul QinC_ring (x_derivative QinC_ring p) y) = poly_1 QinC_ring` [] THEN have `poly_mul QinC_ring p x = poly_mul complex_ring p x:(1->num)->complex` [poly_mul_subring;ring_polynomial] THEN have `ring_polynomial QinC_ring (x_derivative QinC_ring p)` [x_derivative_polynomial] THEN have `poly_mul QinC_ring (x_derivative QinC_ring p) y = poly_mul complex_ring (x_derivative QinC_ring p) y:(1->num)->complex` [poly_mul_subring;ring_polynomial] THEN have `poly_add QinC_ring (poly_mul complex_ring p x) (poly_mul complex_ring (x_derivative QinC_ring p) y) = poly_1 complex_ring` [poly_1_subring] THEN have `poly_add complex_ring (poly_mul complex_ring p x) (poly_mul complex_ring (x_derivative QinC_ring p) y) = poly_1 complex_ring` [poly_add_subring] THEN have `poly_add complex_ring (poly_mul complex_ring p x) (poly_mul complex_ring (x_derivative complex_ring p) y) = poly_1 complex_ring` [x_derivative_subring] THEN specialize_assuming[`complex_ring`;`p:(1->num)->complex`;`x_derivative complex_ring p:(1->num)->complex`;`x:(1->num)->complex`;`y:(1->num)->complex`]coprime_poly_if_linear_combination THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [ring_polynomial_if_subring] THEN have `ring_powerseries complex_ring (p:(1->num)->complex)` [series_complex] THEN have `ring_polynomial complex_ring (x_derivative complex_ring (p:(1->num)->complex))` [x_derivative_polynomial] THEN have `ring_polynomial complex_ring (x:(1->num)->complex)` [ring_polynomial_if_subring] THEN have `ring_polynomial complex_ring (y:(1->num)->complex)` [ring_polynomial_if_subring] THEN have `ring_coprime (x_poly complex_ring) (p,x_derivative complex_ring p)` [] THEN qed[squarefree_if_coprime_derivative;field_complex] ; intro THEN proven_if `ring_coprime(x_poly QinC_ring) (p,x_derivative QinC_ring p)` [squarefree_if_coprime_derivative;field_QinC] THEN have `p IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `x_derivative QinC_ring p IN ring_carrier(x_poly QinC_ring)` [x_derivative_polynomial;x_poly_use] THEN choose `d:(1->num)->complex` `ring_divides(x_poly QinC_ring) d p /\ ring_divides(x_poly QinC_ring) d (x_derivative QinC_ring p) /\ ~(ring_unit(x_poly QinC_ring) d)` [ring_coprime] THEN case `poly_deg QinC_ring (d:(1->num)->complex) = 0` THENL [ have `ring_polynomial QinC_ring (d:(1->num)->complex)` [ring_divides;x_poly_use] THEN choose `c:complex` `c IN ring_carrier(QinC_ring) /\ d = poly_const QinC_ring c:(1->num)->complex` [POLY_DEG_EQ_0] THEN case `c = ring_0 QinC_ring` THENL [ have `d = poly_0 QinC_ring:(1->num)->complex` [poly_0] THEN have `p = poly_0 QinC_ring:(1->num)->complex` [x_poly_use;RING_DIVIDES_ZERO] THEN qed[monic_poly_0;field;field_QinC] ; have `c IN QinC` [QinC_ring_clauses] THEN choose `d:complex` `d IN ring_carrier QinC_ring /\ ring_mul QinC_ring c d = ring_1 QinC_ring` [field;field_QinC] THEN have `ring_unit QinC_ring c` [ring_unit] THEN qed[RING_UNIT_POLY_CONST;x_poly] ] ; pass ] THEN num_linear_fact `~(poly_deg QinC_ring (d:(1->num)->complex) = 0) ==> 1 <= poly_deg QinC_ring (d:(1->num)->complex)` THEN have `1 <= poly_deg complex_ring (d:(1->num)->complex)` [POLY_DEG_SUBRING_GENERATED] THEN have `d IN ring_carrier(x_poly QinC_ring)` [ring_divides] THEN have `ring_polynomial QinC_ring (d:(1->num)->complex)` [x_poly_use] THEN have `ring_polynomial complex_ring (d:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_powerseries QinC_ring (d:(1->num)->complex)` [ring_polynomial] THEN have `d IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN choose2 `z:complex` `q:(1->num)->complex` `ring_polynomial complex_ring q /\ d = poly_mul complex_ring (x_minus_const complex_ring z) q` [nonconstant_complex_x_minus_root] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly] THEN have `ring_powerseries complex_ring (x_minus_const complex_ring z)` [series_complex] THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_powerseries complex_ring (q:(1->num)->complex)` [series_complex] THEN have `ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) d` [x_poly_use;ring_divides] THEN have `p IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN choose `pd:(1->num)->complex` `pd IN ring_carrier(x_poly QinC_ring) /\ p = ring_mul(x_poly QinC_ring) d pd` [ring_divides] THEN have `p = poly_mul QinC_ring d pd:(1->num)->complex` [x_poly_use] THEN have `ring_polynomial QinC_ring (pd:(1->num)->complex)` [x_poly_use] THEN have `ring_powerseries QinC_ring (pd:(1->num)->complex)` [ring_polynomial] THEN have `ring_polynomial complex_ring (pd:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_powerseries complex_ring (pd:(1->num)->complex)` [series_complex] THEN specialize_assuming[`complex_ring`;`QinC`;`d:(1->num)->complex`;`pd:(1->num)->complex`]poly_mul_subring THEN have `p = poly_mul complex_ring d pd:(1->num)->complex` [ring_polynomial] THEN def `qpd:(1->num)->complex` `poly_mul complex_ring q pd:(1->num)->complex` THEN have `ring_polynomial complex_ring (qpd:(1->num)->complex)` [RING_POLYNOMIAL_MUL] THEN have `ring_powerseries complex_ring (qpd:(1->num)->complex)` [series_complex] THEN have `ring_polynomial complex_ring (x_derivative complex_ring qpd)` [x_derivative_polynomial] THEN have `p = poly_mul complex_ring (x_minus_const complex_ring z) qpd` [POLY_MUL_ASSOC;ring_polynomial;RING_MUL] THEN have `x_derivative complex_ring p = poly_add complex_ring (poly_mul complex_ring (x_derivative complex_ring (x_minus_const complex_ring z)) qpd) (poly_mul complex_ring (x_minus_const complex_ring z) (x_derivative complex_ring qpd))` [x_derivative_mul] THEN choose `p'd:(1->num)->complex` `p'd IN ring_carrier(x_poly QinC_ring) /\ x_derivative QinC_ring p = ring_mul(x_poly QinC_ring) d p'd` [ring_divides] THEN have `p'd IN ring_carrier(x_poly complex_ring)` [poly_complex_if_poly_QinC;x_poly_use] THEN have `x_derivative complex_ring p = x_derivative QinC_ring p` [x_derivative_subring] THEN have `x_derivative complex_ring p = ring_mul(x_poly QinC_ring) d p'd` [] THEN have `x_derivative complex_ring p = poly_mul QinC_ring d p'd` [x_poly_use] THEN have `ring_polynomial complex_ring (p'd:(1->num)->complex)` [x_poly_use] THEN have `ring_powerseries complex_ring (p'd:(1->num)->complex)` [series_complex] THEN have `ring_powerseries QinC_ring (p'd:(1->num)->complex)` [x_poly_use;ring_polynomial] THEN have `x_derivative complex_ring p = poly_mul complex_ring d p'd` [poly_mul_subring] THEN def `qp'd:(1->num)->complex` `poly_mul complex_ring q p'd:(1->num)->complex` THEN have `ring_polynomial complex_ring (qp'd:(1->num)->complex)` [RING_POLYNOMIAL_MUL] THEN have `ring_powerseries complex_ring (qp'd:(1->num)->complex)` [series_complex] THEN have `x_derivative complex_ring p = poly_mul complex_ring (x_minus_const complex_ring z) qp'd` [POLY_MUL_ASSOC;ring_polynomial;RING_MUL] THEN have `poly_mul complex_ring (x_minus_const complex_ring z) qp'd = poly_add complex_ring (poly_mul complex_ring (x_derivative complex_ring (x_minus_const complex_ring z)) qpd) (poly_mul complex_ring (x_minus_const complex_ring z) (x_derivative complex_ring qpd))` [] THEN have `poly_mul complex_ring (x_minus_const complex_ring z) qp'd = poly_add complex_ring (poly_mul complex_ring (poly_1 complex_ring) qpd) (poly_mul complex_ring (x_minus_const complex_ring z) (x_derivative complex_ring qpd))` [x_derivative_x_minus_const] THEN have `poly_mul complex_ring (x_minus_const complex_ring z) qp'd = poly_add complex_ring qpd (poly_mul complex_ring (x_minus_const complex_ring z) (x_derivative complex_ring qpd))` [POLY_MUL_LID] THEN specialize[`complex_ring`;`x_minus_const complex_ring z`;`p:(1->num)->complex`;`qpd:(1->num)->complex`;`qp'd:(1->num)->complex`;`x_derivative complex_ring qpd:(1->num)->complex`]monic_QinC_squarefree_complex_squarefree_lemma THEN have `pd IN ring_carrier(x_poly complex_ring)` [poly_complex_if_poly_QinC;x_poly_use] THEN have `p = ring_mul(x_poly complex_ring) d pd` [ring_polynomial;x_poly_use;poly_mul_subring] THEN have `ring_divides(x_poly complex_ring) d p` [ring_divides] THEN have `ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) p` [RING_DIVIDES_TRANS] THEN have `x_derivative QinC_ring p IN ring_carrier(x_poly complex_ring)` [poly_complex_if_poly_QinC;x_poly_use] THEN have `x_derivative QinC_ring p = ring_mul(x_poly complex_ring) d p'd` [ring_polynomial;x_poly_use;poly_mul_subring] THEN have `ring_divides(x_poly complex_ring) d (x_derivative QinC_ring p)` [ring_divides] THEN have `ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) (x_derivative QinC_ring p)` [RING_DIVIDES_TRANS] THEN have `integral_domain(x_poly complex_ring)` [integral_domain_x_poly_field;field_complex] THEN have `~(monic complex_ring (poly_0 complex_ring))` [monic_poly_0;field;field_complex] THEN have `~(p = ring_0(x_poly complex_ring))` [x_poly_use] THEN have `x_minus_const complex_ring z IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN have `~ring_unit(x_poly complex_ring) (x_minus_const complex_ring z)` [x_minus_const_not_unit;integral_domain_complex] THEN specialize_assuming[`x_poly complex_ring`;`p:(1->num)->complex`;`x_minus_const complex_ring z:(1->num)->complex`]not_squarefree_if_divisible_by_square THEN have `p = ring_mul(x_poly complex_ring) (poly_mul complex_ring (x_minus_const complex_ring z) (x_minus_const complex_ring z)) (poly_sub complex_ring qp'd (x_derivative complex_ring qpd))` [x_poly_use] THEN have `p = ring_mul(x_poly complex_ring) (ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) (x_minus_const complex_ring z)) (poly_sub complex_ring qp'd (x_derivative complex_ring qpd))` [x_poly_use] THEN have `p = ring_mul(x_poly complex_ring) (ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) (x_minus_const complex_ring z)) (ring_sub(x_poly complex_ring) qp'd (x_derivative complex_ring qpd))` [x_poly_sub_use] THEN subgoal `ring_divides (x_poly complex_ring) (ring_mul (x_poly complex_ring) (x_minus_const complex_ring z) (x_minus_const complex_ring z)) p` THENL [ rw[ring_divides] THEN intro THENL [ qed[RING_MUL;x_minus_const_poly;x_poly_use] ; qed[] ; witness `ring_sub (x_poly complex_ring) qp'd (x_derivative complex_ring qpd)` THEN qed[RING_SUB;x_poly_use] ] ; qed[] ] ] );; let monic_squarefree_complex_roots = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> ring_squarefree(x_poly complex_ring) p ==> monic complex_ring p ==> ?S. FINITE S /\ p = monic_vanishing_at complex_ring S I `, intro THEN choose `c:num->complex` `p = monic_vanishing_at complex_ring {i | i < poly_deg complex_ring p} c` [monic_vanishing_at_if_monic_complex] THEN def `S:num->bool` `{i | i < poly_deg complex_ring (p:(1->num)->complex)}` THEN have `p = monic_vanishing_at complex_ring (S:num->bool) c` [] THEN witness `IMAGE (c:num->complex) S` THEN subgoal `!e f:num. e IN S ==> f IN S ==> c e = (c f:complex) ==> e = f` THENL [ intro THEN def `d:num->(1->num)->complex` `\j:num. if j = e then x_minus_const complex_ring (c e) else if j = f then x_minus_const complex_ring (c e) else ring_1(x_poly complex_ring)` THEN specialize[`poly_deg complex_ring (p:(1->num)->complex)`]FINITE_NUMSEG_LT THEN have `FINITE (S:num->bool)` [] THEN have `c(e:num) IN ring_carrier complex_ring` [in_complex_ring] THEN have `x_minus_const complex_ring (c(e:num)) IN ring_carrier(x_poly complex_ring)` [x_poly_use;x_minus_const_poly] THEN specialize[`x_poly complex_ring`;`S:num->bool`;`e:num`;`f:num`;`x_minus_const complex_ring (c(e:num))`;`x_minus_const complex_ring (c(e:num))`]ring_product_delta_delta THEN have `ring_product (x_poly complex_ring) (S:num->bool) d = (if e = f:num then x_minus_const complex_ring (c e) else ring_mul (x_poly complex_ring) (x_minus_const complex_ring (c e)) (x_minus_const complex_ring (c e)))` [] THEN subgoal `!s:num. s IN S ==> ring_divides(x_poly complex_ring) (d s) (x_minus_const complex_ring (c s))` THENL [ intro THEN have `c(s:num) IN ring_carrier complex_ring` [in_complex_ring] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring (c(s:num)))` [x_minus_const_poly] THEN have `x_minus_const complex_ring (c(s:num)) IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN case `s = e:num \/ s = f:num` THENL [ have `d(s:num) = x_minus_const complex_ring (c s)` [] THEN specialize[`x_poly complex_ring`;`x_minus_const complex_ring (c(s:num))`]RING_DIVIDES_REFL THEN qed[] ; pass ] THEN have `d(s:num) = ring_1(x_poly complex_ring)` [] THEN specialize[`x_poly complex_ring`;`x_minus_const complex_ring (c(s:num))`]RING_DIVIDES_1 THEN qed[] ; pass ] THEN specialize[`x_poly complex_ring`;`d:num->(1->num)->complex`;`\i:num. x_minus_const complex_ring (c i):(1->num)->complex`;`S:num->bool`]ring_product_divides_factor_by_factor THEN have `integral_domain(x_poly complex_ring)` [integral_domain_x_poly_field;field_complex] THEN have `~(monic complex_ring (poly_0 complex_ring))` [monic_poly_0;field;field_complex] THEN have `~(p = ring_0(x_poly complex_ring))` [x_poly_use] THEN have `~ring_unit(x_poly complex_ring) (x_minus_const complex_ring (c(e:num)))` [x_minus_const_not_unit;integral_domain_complex] THEN proven_if `e = f:num` [] THEN have `ring_product (x_poly complex_ring) (S:num->bool) d = ring_mul (x_poly complex_ring) (x_minus_const complex_ring (c e)) (x_minus_const complex_ring (c(e:num)))` [] THEN have `!s:num. s IN S ==> c s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!s:num. s IN S ==> ring_polynomial complex_ring (x_minus_const complex_ring (c s))` [x_minus_const_poly] THEN specialize[`complex_ring`;`\i:num. x_minus_const complex_ring (c i)`;`S:num->bool`]poly_product_ring_product_x_poly THEN specialize_assuming[`S:num->bool`;`complex_ring`;`c:num->complex`]monic_vanishing_at THEN have `p = poly_product complex_ring S (\s:num. x_minus_const complex_ring (c s))` [] THEN have `ring_product (x_poly complex_ring) (S:num->bool) (\i. x_minus_const complex_ring (c i)) = p` [monic_vanishing_at] THEN have `ring_divides(x_poly complex_ring) (ring_mul (x_poly complex_ring) (x_minus_const complex_ring (c e)) (x_minus_const complex_ring (c(e:num)))) p` [] THEN specialize[`x_poly complex_ring`;`p:(1->num)->complex`;`x_minus_const complex_ring (c(e:num))`]not_squarefree_if_divisible_by_square THEN qed[] ; pass ] THEN specialize[`poly_deg complex_ring (p:(1->num)->complex)`]FINITE_NUMSEG_LT THEN have `FINITE (S:num->bool)` [] THEN have `FINITE (IMAGE (c:num->complex) S)` [FINITE_IMAGE] THEN specialize[`complex_ring`;`S:num->bool`;`c:num->complex`;`I:complex->complex`]monic_vanishing_at_image THEN qed[monic_vanishing_at_image;I_O_ID] );; let QinC_monic_irreducible_complex_roots = prove(` !p:(1->num)->complex. ring_polynomial QinC_ring p ==> ring_irreducible(x_poly QinC_ring) p ==> monic QinC_ring p ==> ?S. FINITE S /\ p = monic_vanishing_at complex_ring S I `, intro THEN have `UFD(x_poly QinC_ring)` [UFD_x_poly_QinC] THEN have `ring_prime(x_poly QinC_ring) p` [UFD_IRREDUCIBLE_EQ_PRIME] THEN have `ring_squarefree(x_poly QinC_ring) p` [ring_squarefree_if_prime] THEN have `ring_squarefree(x_poly complex_ring) p` [monic_QinC_squarefree_complex_squarefree] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring;subring_complex_QinC] THEN qed[monic_squarefree_complex_roots] );; (* ===== complex_root *) let complex_root = new_definition ` complex_root p (z:complex) <=> poly_eval complex_ring p z = Cx(&0) `;; let complex_root_ring = prove(` !p z. complex_root p z <=> poly_eval complex_ring p z = ring_0 complex_ring `, qed[complex_root;complex_ring_clauses] );; let complex_root_divides = prove(` !p q. ring_divides(x_poly complex_ring) p q ==> complex_root p SUBSET complex_root q `, rw[SUBSET;IN;complex_root_ring] THEN intro THEN choose `u:(1->num)->complex` `u IN ring_carrier(x_poly complex_ring) /\ q = ring_mul(x_poly complex_ring) p u` [ring_divides] THEN have `p IN ring_carrier(x_poly complex_ring)` [ring_divides] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [x_poly_use] THEN have `ring_polynomial complex_ring (u:(1->num)->complex)` [x_poly_use] THEN have `x IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`p:(1->num)->complex`;`u:(1->num)->complex`;`x:complex`]POLY_EVAL_MUL THEN simp[] THEN rw[GSYM x_poly_use] THEN simp[] THEN qed[RING_MUL_LZERO;POLY_EVAL] );; let complex_root_associates = prove(` !p q. ring_associates(x_poly complex_ring) p q ==> complex_root p = complex_root q `, qed[ring_associates;complex_root_divides;SUBSET_ANTISYM] );; let complex_root_le_deg = prove(` !p. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> (FINITE(complex_root p) /\ CARD(complex_root p) <= poly_deg complex_ring p ) `, GEN_TAC THEN REPEAT DISCH_TAC THEN recall integral_domain_complex THEN have `p IN ring_carrier(poly_ring complex_ring (:1))` [x_poly_use;x_poly] THEN have `~(p = ring_0(poly_ring complex_ring (:1)))` [x_poly_use;x_poly] THEN specialize[`complex_ring`;`p:(1->num)->complex`]POLY_ROOT_BOUND THEN subgoal `complex_root p = { x | x IN ring_carrier complex_ring /\ poly_eval complex_ring p x = ring_0 complex_ring}` THENL [ rw[EXTENSION;IN;IN_ELIM_THM;complex_root_ring] THEN qed[complex_ring_clauses;UNIV] ; qed[] ] );; let complex_root_monic_vanishing_at = prove(` !S. FINITE S ==> complex_root (monic_vanishing_at complex_ring S I) = S `, rw[FUN_EQ_THM] THEN rw[complex_root_ring] THEN intro THEN splitiff THENL [ intro THEN have `x IN ring_carrier complex_ring` [in_complex_ring] THEN have `!s. s IN S ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`S:complex->bool`;`I:complex->complex`;`x:complex`]eval_monic_vanishing_at THEN recall integral_domain_complex THEN have `ring_product complex_ring S (\s. ring_sub complex_ring x (I s)) = ring_0 complex_ring` [] THEN choose `s:complex` `s IN S /\ ring_sub complex_ring x (I s) = ring_0 complex_ring` [INTEGRAL_DOMAIN_PRODUCT_EQ_0] THEN have `x = I s:complex` [RING_SUB_EQ_0;in_complex_ring] THEN qed[IN;I_DEF] ; pass ] THEN intro THEN have `x:complex IN S` [IN] THEN have `!s. s IN S ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`S:complex->bool`;`I:complex->complex`;`x:complex`]eval_monic_vanishing_at_refl THEN qed[I_DEF] );; let monic_vanishing_at_complex_root = prove(` !p. ring_polynomial complex_ring p ==> ring_squarefree(x_poly complex_ring) p ==> monic complex_ring p ==> monic_vanishing_at complex_ring (complex_root p) I = p `, intro THEN choose `S:complex->bool` `FINITE S /\ p = monic_vanishing_at complex_ring S I` [monic_squarefree_complex_roots] THEN qed[complex_root_monic_vanishing_at] );; let complex_root_if_x_minus_const_divides = prove(` !p z. ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) p ==> complex_root p z `, intro THEN have `p IN ring_carrier(x_poly complex_ring)` [ring_divides] THEN choose `q:(1->num)->complex` `q IN ring_carrier(x_poly complex_ring) /\ p = ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) q` [ring_divides] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [x_poly_use] THEN specialize_assuming[`complex_ring`;`x_minus_const complex_ring z`;`q:(1->num)->complex`;`z:complex`]POLY_EVAL_MUL THEN have `poly_eval complex_ring p z = ring_mul complex_ring (poly_eval complex_ring (x_minus_const complex_ring z) z) (poly_eval complex_ring q z)` [in_complex_ring;x_poly_use] THEN simp[complex_root] THEN simp[eval_x_minus_const_refl;in_complex_ring] THEN simp[RING_MUL_LZERO;POLY_EVAL] THEN qed[complex_ring_clauses] );; let complex_root_x_pow = prove(` !n. complex_root(x_pow QinC_ring n) = if n = 0 then {} else {Cx(&0)} `, intro THEN rw[x_pow_QinC_eq_x_pow_complex] THEN case `n = 0` THENL [ simp[EXTENSION;EMPTY;IN_ELIM_THM;IN;complex_root] THEN simp[eval_x_pow;in_complex_ring;RING_POW_0] THEN qed[ring_1_0_complex;complex_ring_clauses] ; pass ] THEN simp[EXTENSION;IN_SING] THEN simp[IN;complex_root] THEN intro THEN simp[eval_x_pow;in_complex_ring] THEN qed[COMPLEX_POW_EQ_0;ring_pow_complex] );; let complex_root_x_minus_const = prove(` !c. complex_root(x_minus_const complex_ring c) = {c} `, intro THEN rw[EXTENSION;IN_SING;IN;complex_root] THEN intro THEN simp[eval_x_minus_const;in_complex_ring;complex_ring_clauses;ring_sub_complex] THEN CONV_TAC COMPLEX_FIELD );; let not_coprime_QinC_if_shared_complex_root = prove(` !z:complex p q. complex_root p z ==> complex_root q z ==> ~(ring_coprime(x_poly QinC_ring) (p,q)) `, intro THEN have `PID(x_poly QinC_ring)` [PID_x_poly_QinC] THEN have `bezout_ring(x_poly QinC_ring)` [PID_IMP_BEZOUT_RING] THEN choose2 `x:(1->num)->complex` `y:(1->num)->complex` `x IN ring_carrier(x_poly QinC_ring) /\ y IN ring_carrier(x_poly QinC_ring) /\ ring_add(x_poly QinC_ring) (ring_mul(x_poly QinC_ring) p x) (ring_mul(x_poly QinC_ring) q y) = ring_1(x_poly QinC_ring)` [BEZOUT_RING_COPRIME] THEN have `ring_polynomial QinC_ring (x:(1->num)->complex)` [x_poly_use] THEN have `ring_polynomial QinC_ring (y:(1->num)->complex)` [x_poly_use] THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [ring_coprime;x_poly_use] THEN have `ring_polynomial QinC_ring (q:(1->num)->complex)` [ring_coprime;x_poly_use] THEN have `ring_polynomial complex_ring (x:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (y:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (poly_mul complex_ring p x:(1->num)->complex)` [RING_POLYNOMIAL_MUL] THEN have `ring_polynomial complex_ring (poly_mul complex_ring q y:(1->num)->complex)` [RING_POLYNOMIAL_MUL] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `p:(1->num)->complex`; `x:(1->num)->complex`; `z:complex` ]POLY_EVAL_MUL THEN have `poly_eval complex_ring (poly_mul complex_ring p x) z = Cx(&0)` [complex_root;complex_ring_clauses;COMPLEX_MUL_LZERO] THEN specialize[ `complex_ring`; `q:(1->num)->complex`; `y:(1->num)->complex`; `z:complex` ]POLY_EVAL_MUL THEN have `poly_eval complex_ring (poly_mul complex_ring q y) z = Cx(&0)` [complex_root;complex_ring_clauses;COMPLEX_MUL_LZERO] THEN specialize[ `complex_ring`; `poly_mul complex_ring p x:(1->num)->complex`; `poly_mul complex_ring q y:(1->num)->complex`; `z:complex` ]POLY_EVAL_ADD THEN have `poly_eval complex_ring (poly_add complex_ring (poly_mul complex_ring p x) (poly_mul complex_ring q y)) z = Cx(&0)` [COMPLEX_ADD_LID;complex_ring_clauses] THEN have `poly_mul complex_ring p x = poly_mul QinC_ring p x:(1->num)->complex` [poly_mul_subring;subring_complex_QinC;ring_polynomial] THEN have `poly_mul complex_ring q y = poly_mul QinC_ring q y:(1->num)->complex` [poly_mul_subring;subring_complex_QinC;ring_polynomial] THEN have `poly_add complex_ring (poly_mul complex_ring p x) (poly_mul complex_ring q y) = poly_add QinC_ring (poly_mul QinC_ring p x) (poly_mul QinC_ring q y):(1->num)->complex` [poly_add_subring;subring_complex_QinC;ring_polynomial] THEN subgoal `poly_add QinC_ring (poly_mul QinC_ring p x) (poly_mul QinC_ring q y):(1->num)->complex = poly_1 complex_ring` THENL [ rw[GSYM poly_1_QinC_eq_poly_1_complex] THEN rw[x_poly_use] THEN qed[] ; pass ] THEN have `poly_eval complex_ring (poly_1 complex_ring) z = Cx(&0)` [] THEN have `poly_eval complex_ring (poly_1 complex_ring) z = Cx(&1)` [POLY_EVAL_1;complex_ring_clauses] THEN complex_field_fact `~(Cx(&0) = Cx(&1))` THEN qed[] );; let complex_root_minimal_polynomial_refl = prove(` !z:complex. complex_root (minimal_polynomial z) z `, rw[complex_root] THEN qed[maybe_algebraic_minimal_polynomial] );; let minimal_polynomial_divides = prove(` !z:complex p:(1->num)->complex. algebraic_number z ==> ring_polynomial QinC_ring p ==> ( ring_divides(x_poly QinC_ring) (minimal_polynomial z) p <=> complex_root p z ) `, intro THEN splitiff THENL [ intro THEN specialize[ `minimal_polynomial z`; `p:(1->num)->complex` ]ring_divides_poly_complex_if_ring_divides_poly_QinC THEN have `complex_root (minimal_polynomial z) SUBSET complex_root p` [complex_root_divides] THEN qed[SUBSET;IN;complex_root_minimal_polynomial_refl] ; intro THEN have `ring_polynomial QinC_ring (minimal_polynomial z) /\ ~((minimal_polynomial z) = poly_0 QinC_ring) /\ monic QinC_ring (minimal_polynomial z) /\ ring_irreducible (x_poly QinC_ring) (minimal_polynomial z) /\ poly_eval complex_ring (minimal_polynomial z) z = Cx (&0)` [algebraic_has_minimal_polynomial] THEN have `~ring_coprime (x_poly QinC_ring) (minimal_polynomial z,p)` [not_coprime_QinC_if_shared_complex_root;complex_root] THEN recall integral_domain_x_poly_QinC THEN have `minimal_polynomial z IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `p IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `ring_prime(x_poly QinC_ring) (minimal_polynomial z)` [prime_iff_irreducible_over_field;field_QinC] THEN specialize[ `x_poly QinC_ring`; `minimal_polynomial z`; `p:(1->num)->complex` ]INTEGRAL_DOMAIN_PRIME_DIVIDES_OR_COPRIME THEN qed[] ] );; (* ===== complex_root_powersums: sum_{z:p(z)=0} z^n *) let complex_root_powersums = new_definition ` complex_root_powersums p n = ring_sum complex_ring (complex_root p) (\z. z pow n) `;; let complex_root_powersums_ring = prove(` !p n. complex_root_powersums p n = ring_sum complex_ring (complex_root p) (\z. ring_pow complex_ring z n) `, rw[complex_root_powersums;ring_pow_complex] );; let complex_root_powersums_QinC_monic_irreducible_QinC = prove(` !(p:(1->num)->complex) n. ring_polynomial QinC_ring p ==> ring_irreducible(x_poly QinC_ring) p ==> monic QinC_ring p ==> complex_root_powersums p n IN QinC `, rw[complex_root_powersums_ring] THEN intro THEN recall subring_complex_QinC THEN have `!n. coeff n p IN ring_carrier(QinC_ring)` [coeff_poly_in_ring] THEN choose `S:complex->bool` `FINITE S /\ p = monic_vanishing_at complex_ring S I` [QinC_monic_irreducible_complex_roots] THEN have `complex_root p = S` [complex_root_monic_vanishing_at] THEN have `FINITE (complex_root p)` [] THEN have `!s. s IN complex_root p ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!n. coeff n (monic_vanishing_at complex_ring (complex_root p) I) IN ring_carrier(subring_generated complex_ring QinC)` [coeff_poly_in_ring] THEN specialize[`complex_ring`;`QinC`;`complex_root p`;`I:complex->complex`;`n:num`]powersums_subring_if_poly_subring THEN have `!z. z IN complex_root p ==> ring_pow complex_ring z n = ring_pow complex_ring (I z) n` [I_DEF] THEN specialize[`complex_ring`;`\z. ring_pow complex_ring z n`;`\z. ring_pow complex_ring (I z) n`;`complex_root p`]RING_SUM_EQ THEN qed[QinC_ring_clauses] );; (* Another way to prove the following: prove complex_root_powersums_ZinC_monic_irreducible_ZinC (as in complex_root_powersums_ZinC_monic_irreducible_ZinC, plus relating ZinC irreducibles to QinC irreducibles) and then multiply each root by D to reduce to that case. Alternatively, after proving this, can specialize to Z=1 to prove complex_root_powersums_ZinC_monic_irreducible_ZinC. Yet another way, after showing that algebraic integers form a ring: Dz is an algebraic integer for each z so sum_z (Dz)^n is an algebraic integer but D^n sum_z z^n is rational, ergo integer. This again boils down to Newton's identities. *) let complex_root_powersums_QinC_monic_irreducible_ZinC = prove(` !(p:(1->num)->complex) D n. ring_polynomial QinC_ring p ==> ring_irreducible(x_poly QinC_ring) p ==> monic QinC_ring p ==> (!i. D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> D pow n * complex_root_powersums p n IN ZinC `, rw[complex_root_powersums_ring] THEN intro THEN recall subring_complex_empty THEN choose `S:complex->bool` `FINITE S /\ p = monic_vanishing_at complex_ring S I` [QinC_monic_irreducible_complex_roots] THEN have `complex_root p = S` [complex_root_monic_vanishing_at] THEN have `~(ring_1 complex_ring = ring_0 complex_ring)` [field_complex;field] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN recall subring_complex_QinC THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring] THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [monic_poly_0;field_complex;field] THEN have `FINITE (complex_root p)` [complex_root_le_deg] THEN have `poly_deg complex_ring (p:(1->num)->complex) = CARD(S:complex->bool)` [deg_monic_vanishing_at;in_complex_ring] THEN have `!s. s IN complex_root p ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN subgoal `!n. ring_mul complex_ring (ring_pow complex_ring D n) (coeff (CARD(complex_root p) - n) (monic_vanishing_at complex_ring (complex_root p) I)) IN ring_carrier(subring_generated complex_ring {})` THENL [ simp[complex_ring_clauses;ring_pow_complex;ZinC_ring_clauses] THEN qed[] ; pass ] THEN have `D IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`{}:complex->bool`;`complex_root p`;`I:complex->complex`;`D:complex`;`n:num`]powersums_subring_if_poly_subring_denominators THEN have `(D pow n) * (ring_sum complex_ring (complex_root p) (\s. ring_pow complex_ring (I s) n)) IN ring_carrier (subring_generated complex_ring {})` [complex_ring_clauses;ring_pow_complex] THEN have `(D pow n) * (ring_sum complex_ring (complex_root p) (\s. ring_pow complex_ring (I s) n)) IN ZinC` [ZinC_ring_clauses] THEN have `!s. ring_pow complex_ring (I s) n = ring_pow complex_ring s n` [I_DEF;ring_pow_complex] THEN qed[RING_SUM_EQ] );; (* ===== denominators of monic polys *) (* meaning: D where all D^i p_(deg p - i) are integers *) (* equivalently: Dz is algebraic integer for each root z *) let denominator_if_monic_QinC = prove(` !(p:(1->num)->complex). ring_polynomial QinC_ring p ==> monic QinC_ring p ==> ?D. D IN ZinC /\ ~(D = Cx(&0)) /\ (!i. D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) `, intro THEN def `n:num` `poly_deg complex_ring (p:(1->num)->complex)` THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN have `!i. i IN 0..n ==> coeff (n-i) p IN QinC` [coeff_poly_in_ring;QinC_ring_clauses] THEN specialize[`0..n`;`\i. coeff (n-i) p:complex`]multi_QinC_to_ZinC THEN choose `d:num` `~(d = 0) /\ !i. i IN 0..n ==> Cx(&d) * (coeff (n-i) p) IN ZinC` [] THEN witness `Cx(&d)` THEN have `Cx(&d) IN ZinC` [num_in_ZinC] THEN intro THENL [ qed[] ; have `~(&d = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&d) = Cx(&0))` [CX_INJ] THEN qed[] ; pass ] THEN case `i = 0` THENL [ simp[complex_pow;COMPLEX_MUL_LID;ARITH_RULE `d - 0 = d`] THEN recall subring_complex_QinC THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring] THEN have `coeff (poly_deg complex_ring p) p = ring_1 complex_ring` [monic] THEN qed[complex_ring_clauses;ZinC_1] ; pass ] THEN num_linear_fact `~(i = 0) ==> (i-1)+1 = i` THEN have `Cx(&d) pow i = Cx(&d) pow (i-1) * Cx(&d)` [COMPLEX_POW_ADD;COMPLEX_POW_1] THEN simp[GSYM COMPLEX_MUL_ASSOC] THEN have `Cx(&d) pow (i-1) IN ZinC` [ring_pow_ZinC;RING_POW;ring_pow_complex;ZinC_ring_clauses] THEN subgoal `Cx(&d) * coeff(n-i) p IN ZinC` THENL [ proven_if `i <= n:num` [IN_NUMSEG_0] THEN num_linear_fact `~(i <= n) ==> n - i = n - n:num` THEN num_linear_fact `n <= n:num` THEN have `n IN 0..n` [IN_NUMSEG_0] THEN qed[] ; pass ] THEN qed[ZinC_ring_clauses;RING_MUL;complex_ring_clauses] );; let denominator_reverse = prove(` !(p:(1->num)->complex). ring_polynomial QinC_ring p ==> monic QinC_ring p ==> (!i. D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> (!i. D pow i * (coeff i (x_truncreverse QinC_ring (poly_deg complex_ring p) p)) IN ZinC) `, intro THEN rw[coeff_x_truncreverse] THEN qed[QinC_ring_clauses;COMPLEX_MUL_RZERO;ZinC_0] );; let denominator_reverse_mul = prove(` !(p:(1->num)->complex) q D. ring_powerseries QinC_ring p ==> ring_powerseries QinC_ring q ==> D IN ZinC ==> (!i. D pow i * (coeff i p) IN ZinC) ==> (!i. D pow i * (coeff i q) IN ZinC) ==> (!i. D pow i * (coeff i (poly_mul QinC_ring p q)) IN ZinC) `, intro THEN rw[coeff_poly_mul_oneindex] THEN have `D pow i IN ZinC` [RING_POW;ring_pow_ZinC;ZinC_ring_clauses] THEN have `D pow i IN QinC` [ZinC_in_QinC] THEN have `D pow i IN ring_carrier QinC_ring` [QinC_ring_clauses] THEN have `FINITE (0..i)` [FINITE_NUMSEG] THEN have `!a. a IN 0..i ==> ring_mul QinC_ring (coeff a p) (coeff (i - a) q) IN ring_carrier QinC_ring` [coeff_series_in_ring;RING_MUL] THEN specialize[`QinC_ring`;`(\a. ring_mul QinC_ring (coeff a p) (coeff (i - a) q))`;`(D:complex) pow i`;`0..i`](GSYM RING_SUM_LMUL) THEN simp[GSYM QinC_ring_clauses] THEN subgoal `!s. s IN 0..i ==> ring_mul QinC_ring (D pow i) (ring_mul QinC_ring (coeff s p) (coeff (i - s) q)) IN ring_carrier (subring_generated QinC_ring {})` THENL [ intro THEN rw[subring_QinC_empty;QinC_ring_clauses;ZinC_ring_clauses] THEN have `s <= i:num` [IN_NUMSEG_0] THEN num_linear_fact `s <= i:num ==> i = s + (i-s)` THEN have `(D:complex) pow i = D pow s * D pow (i-s)` [COMPLEX_POW_ADD] THEN complex_field_fact `((D:complex) pow s * D pow (i-s)) * coeff s p * coeff (i - s) q = (D pow s * coeff s p) * (D pow (i-s) * coeff (i-s) q)` THEN qed[mul_in_ZinC] ; pass ] THEN specialize[`QinC_ring`;`{}:complex->bool`;`0..i`;`\x. ring_mul QinC_ring (D pow i) (ring_mul QinC_ring (coeff x p) (coeff (i - x) q))`]ring_sum_in_subring THEN qed[subring_QinC_empty;ZinC_ring_clauses] );; let denominator_reverse_product = prove(` !p D (S:X->bool). FINITE S ==> (!s. s IN S ==> ring_powerseries QinC_ring (p s)) ==> D IN ZinC ==> (!s i. s IN S ==> D pow i * (coeff i (p s)) IN ZinC) ==> (!i. D pow i * (coeff i (poly_product QinC_ring S p)) IN ZinC) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty] THEN rw[coeff_poly_1] THEN rw[QinC_ring_clauses] THEN have `D pow i * Cx(&1) IN ZinC` [COMPLEX_MUL_RID;RING_POW;ring_pow_ZinC;ZinC_ring_clauses] THEN have `D pow i * Cx(&0) IN ZinC` [COMPLEX_MUL_RZERO;ZinC_0] THEN qed[] ; set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[poly_product_insert] THEN have `ring_powerseries QinC_ring (p(x:X):(1->num)->complex)` [] THEN have `ring_powerseries QinC_ring (poly_product QinC_ring (S:X->bool) p)` [poly_product_series] THEN have `!i. D pow i * coeff i (p(x:X)) IN ZinC` [] THEN have `!s:X. s IN S ==> ring_powerseries QinC_ring (p s:(1->num)->complex)` [] THEN have `!s:X i. s IN S ==> D pow i * coeff i (p s) IN ZinC` [] THEN have `!i. D pow i * coeff i (poly_product QinC_ring (S:X->bool) p) IN ZinC` [] THEN specialize[`p(x:X):(1->num)->complex`;`poly_product QinC_ring (S:X->bool) p`;`D:complex`]denominator_reverse_mul THEN qed[] ] );; let denominator_reverse_pow = prove(` !p D n. ring_powerseries QinC_ring p ==> D IN ZinC ==> (!i. D pow i * (coeff i p) IN ZinC) ==> (!i. D pow i * (coeff i (poly_pow QinC_ring p n)) IN ZinC) `, intro THEN simp[poly_pow_is_product] THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN specialize_assuming[`\i:num. p:(1->num)->complex`;`D:complex`;`1..n`]denominator_reverse_product THEN qed[] );; (* ===== distinct minpolys *) let distinct_minpolys = new_definition ` distinct_minpolys P <=> !p. p IN P ==> (ring_polynomial QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ monic QinC_ring p ) `;; let polynomial_product_distinct_minpolys = prove(` !P. FINITE P ==> distinct_minpolys P ==> ring_polynomial QinC_ring (poly_product QinC_ring P I) `, rw[distinct_minpolys] THEN intro THEN have `!p. p IN P ==> ring_polynomial QinC_ring (I p:(1->num)->complex)` [I_DEF] THEN qed[poly_product_poly] );; let monic_product_distinct_minpolys = prove(` !P. FINITE P ==> distinct_minpolys P ==> monic QinC_ring (poly_product QinC_ring P I) `, rw[distinct_minpolys] THEN intro THEN have `!p. p IN P ==> ring_polynomial QinC_ring (I p:(1->num)->complex)` [I_DEF] THEN have `!p. p IN P ==> monic QinC_ring (I p)` [I_DEF] THEN qed[monic_poly_product] );; let distinct_minpolys_nonzero = prove(` !P p. distinct_minpolys P ==> p IN P ==> ~(p = poly_0 QinC_ring) `, qed[distinct_minpolys;monic_poly_0;field_QinC;field] );; let distinct_minpolys_reverse_nonzero = prove(` !P p. distinct_minpolys P ==> p IN P ==> ~(x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_0 complex_ring) `, intro THEN recall integral_domain_QinC THEN have `~(ring_1 QinC_ring = ring_0 QinC_ring)` [integral_domain] THEN have `monic QinC_ring p` [distinct_minpolys] THEN have `poly_deg QinC_ring (p:(1->num)->complex) = poly_deg complex_ring p` [poly_deg_subring;subring_complex_QinC] THEN have `x_truncreverse QinC_ring (poly_deg QinC_ring p) p = poly_0 QinC_ring` [poly_0_QinC_eq_poly_0_complex] THEN qed[nonzero_if_x_truncreverse_monic;integral_domain] );; let distinct_minpolys_total_deg = prove(` !P. FINITE P ==> distinct_minpolys P ==> (P = {} <=> nsum P (\p. poly_deg complex_ring p) = 0) `, intro THEN splitiff THENL [ simp[NSUM_CLAUSES] ; rw[EXTENSION;EMPTY;IN_ELIM_THM] THEN intro THEN have `poly_deg complex_ring (x:(1->num)->complex) = 0` [NSUM_EQ_0_IFF] THEN have `poly_deg QinC_ring (x:(1->num)->complex) = 0` [poly_deg_subring;subring_complex_QinC] THEN recall field_QinC THEN have `~(ring_prime(x_poly QinC_ring) x)` [deg_prime] THEN have `~(ring_irreducible(x_poly QinC_ring) x)` [prime_iff_irreducible_over_field] THEN qed[distinct_minpolys] ] );; let distinct_minpolys_monic_vanishing_at_simple = prove(` !p. ring_polynomial QinC_ring p ==> monic QinC_ring p ==> ring_irreducible(x_poly QinC_ring) p ==> monic_vanishing_at complex_ring (complex_root p) I = p `, intro THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN recall subring_complex_QinC THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring] THEN have `ring_squarefree(x_poly QinC_ring) p` [squarefree_if_irreducible_over_field;field_QinC] THEN have `ring_squarefree(x_poly complex_ring) p` [monic_QinC_squarefree_complex_squarefree] THEN qed[monic_vanishing_at_complex_root] );; let distinct_minpolys_monic_vanishing_at = prove(` !P p. distinct_minpolys P ==> p IN P ==> monic_vanishing_at complex_ring (complex_root p) I = p `, qed[distinct_minpolys;distinct_minpolys_monic_vanishing_at_simple] );; let distinct_minpolys_deg_nonzero = prove(` !P p. distinct_minpolys P ==> p IN P ==> ~(poly_deg complex_ring p = 0) `, intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_irreducible(x_poly QinC_ring) (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_prime(x_poly QinC_ring) (p:(1->num)->complex)` [prime_iff_irreducible_over_field;field_QinC] THEN have `~(poly_deg QinC_ring (p:(1->num)->complex) = 0)` [deg_prime;field_QinC] THEN qed[poly_deg_subring;subring_complex_QinC] );; let distinct_minpolys_finite_root_simple = prove(` !p. ring_polynomial QinC_ring p ==> monic QinC_ring p ==> ring_irreducible(x_poly QinC_ring) p ==> FINITE(complex_root p) `, intro THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p:(1->num)->complex = poly_0 QinC_ring)` [monic_poly_0;field_QinC;field] THEN recall subring_complex_QinC THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [poly_0_subring] THEN qed[complex_root_le_deg] );; let distinct_minpolys_finite_root = prove(` !P p. distinct_minpolys P ==> p IN P ==> FINITE(complex_root p) `, qed[distinct_minpolys;distinct_minpolys_finite_root_simple] );; let distinct_minpolys_card_root = prove(` !P p. distinct_minpolys P ==> p IN P ==> CARD(complex_root p) = poly_deg complex_ring p `, intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p:(1->num)->complex = poly_0 QinC_ring)` [distinct_minpolys_nonzero] THEN recall subring_complex_QinC THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [poly_0_subring] THEN have `monic_vanishing_at complex_ring (complex_root p) I = p` [distinct_minpolys_monic_vanishing_at] THEN have `!c. c IN complex_root p ==> I c IN ring_carrier complex_ring` [in_complex_ring] THEN have `FINITE (complex_root p)` [distinct_minpolys_finite_root] THEN specialize[`complex_ring`;`I:complex->complex`;`complex_root p`]deg_monic_vanishing_at THEN qed[deg_monic_vanishing_at;field_complex;field] );; let distinct_minpolys_denominator = prove(` !P. FINITE P ==> distinct_minpolys P ==> ?D. D IN ZinC /\ ~(D = Cx(&0)) /\ (!p. p IN P ==> !i. D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC ) `, rw[distinct_minpolys] THEN intro THEN def `denom:((1->num)->complex)->complex` `\p. @d. d IN ZinC /\ ~(d = Cx(&0)) /\ (!i. d pow i * (coeff(poly_deg complex_ring p - i) p) IN ZinC)` THEN subgoal `!p:(1->num)->complex. p IN P ==> denom p IN ZinC /\ ~(denom p = Cx(&0)) /\ (!i. (denom p) pow i * (coeff(poly_deg complex_ring p - i) p) IN ZinC)` THENL [ intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [] THEN have `monic QinC_ring (p:(1->num)->complex)` [] THEN specialize[`p:(1->num)->complex`]denominator_if_monic_QinC THEN qed[] ; pass ] THEN def `D:complex` `ring_product ZinC_ring P (denom:((1->num)->complex)->complex)` THEN have `D IN ZinC` [RING_PRODUCT;ZinC_ring_clauses] THEN witness `D:complex` THEN intro THENL [ qed[] ; qed[INTEGRAL_DOMAIN_PRODUCT_EQ_0;integral_domain_ZinC;ZinC_ring_clauses] ; pass ] THEN def `q:complex` `ring_product ZinC_ring (P DELETE (p:(1->num)->complex)) denom` THEN have `denom (p:(1->num)->complex) IN ring_carrier ZinC_ring` [ZinC_ring_clauses] THEN specialize[`ZinC_ring`;`P:((1->num)->complex)->bool`;`p:(1->num)->complex`;`denom:((1->num)->complex)->complex`]ring_product_delete THEN have `D = (denom (p:(1->num)->complex)) * q:complex` [ZinC_ring_clauses] THEN have `D pow i = (denom (p:(1->num)->complex) pow i) * (q pow i):complex` [COMPLEX_POW_MUL] THEN have `q pow i IN ZinC` [RING_PRODUCT;RING_POW;ZinC_ring_clauses;ring_pow_ZinC] THEN have `q pow i * denom p pow i * coeff (poly_deg complex_ring p - i) p IN ZinC` [RING_MUL;ZinC_ring_clauses] THEN have `(q pow i * denom p pow i) * coeff (poly_deg complex_ring p - i) p IN ZinC` [COMPLEX_MUL_ASSOC] THEN have `(denom p pow i * q pow i) * coeff (poly_deg complex_ring p - i) p IN ZinC` [COMPLEX_MUL_SYM] THEN qed[] );; let weighted_powersums_distinct_minpolys = prove(` !P D B n. FINITE P ==> distinct_minpolys P ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> D pow n * (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n)) IN ZinC `, intro THEN subgoal `!p. p IN P ==> D pow n * complex_root_powersums p n IN ZinC` THENL [ intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_irreducible(x_poly QinC_ring) p` [distinct_minpolys] THEN have `monic QinC_ring p` [distinct_minpolys] THEN have `!i. D pow i * coeff(poly_deg complex_ring p - i) p IN ZinC` [] THEN specialize[`p:(1->num)->complex`;`D:complex`;`n:num`]complex_root_powersums_QinC_monic_irreducible_ZinC THEN qed[] ; pass ] THEN have `!p. p IN P ==> (complex_of_int(B p)) * (D pow n * complex_root_powersums p n) IN ZinC` [ZinC_ring_clauses;int_in_ZinC;RING_MUL] THEN have `!p. p IN P ==> ((complex_of_int(B p)) * D pow n) * complex_root_powersums p n IN ZinC` [COMPLEX_MUL_ASSOC] THEN have `!p. p IN P ==> (D pow n * (complex_of_int(B p))) * complex_root_powersums p n IN ZinC` [COMPLEX_MUL_SYM] THEN have `!p. p IN P ==> D pow n * ((complex_of_int(B p)) * complex_root_powersums p n) IN ZinC` [COMPLEX_MUL_ASSOC] THEN have `D pow n IN ring_carrier complex_ring` [in_complex_ring] THEN have `!p. p IN P ==> (complex_of_int(B p)) * complex_root_powersums p n IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`\p. (complex_of_int(B p)) * complex_root_powersums p n`;`(D pow n):complex`;`P:((1->num)->complex)->bool`](GSYM RING_SUM_LMUL) THEN have `(D pow n) * (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n)) = ring_sum complex_ring P (\x. (D pow n) * ((complex_of_int(B x)) * complex_root_powersums x n))` [complex_ring_clauses;RING_SUM_EQ] THEN simp[] THEN recall subring_complex_empty THEN have `!p. p IN P ==> D pow n * (complex_of_int(B p)) * complex_root_powersums p n IN ring_carrier(subring_generated complex_ring {})` [ZinC_ring_clauses] THEN specialize[`complex_ring`;`{}:complex->bool`;`P:((1->num)->complex)->bool`;`\x. D pow n * (complex_of_int(B x)) * complex_root_powersums x n`]ring_sum_subring_generated_v2 THEN have `ring_sum complex_ring P (\x. D pow n * (complex_of_int(B x)) * complex_root_powersums x n) = ring_sum ZinC_ring P (\x. D pow n * (complex_of_int(B x)) * complex_root_powersums x n)` [] THEN qed[RING_SUM;ZinC_ring_clauses] );; let distinct_minpolys_distinct_roots = prove(` !P p q z. distinct_minpolys P ==> p IN P ==> q IN P ==> ~(p = q) ==> complex_root p z ==> ~(complex_root q z) `, intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_polynomial QinC_ring (q:(1->num)->complex)` [distinct_minpolys] THEN have `monic QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `monic QinC_ring (q:(1->num)->complex)` [distinct_minpolys] THEN have `ring_irreducible(x_poly QinC_ring) p` [distinct_minpolys] THEN have `ring_irreducible(x_poly QinC_ring) q` [distinct_minpolys] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [ring_irreducible;x_poly_use] THEN have `~(q = poly_0 QinC_ring:(1->num)->complex)` [ring_irreducible;x_poly_use] THEN have `ring_irreducible(x_poly QinC_ring) q` [distinct_minpolys] THEN have `poly_eval complex_ring p z = Cx(&0)` [complex_root] THEN have `poly_eval complex_ring q z = Cx(&0)` [complex_root] THEN qed[algebraic_number_is_root_unique_monic_irreducible_QinC_poly_simple] );; let distinct_minpolys_zero_root = prove(` !P q. distinct_minpolys P ==> q IN P ==> complex_root q (Cx(&0)) ==> q = x_pow QinC_ring 1 `, intro THEN def `S:((1->num)->complex)->bool` `\p. p = q \/ p = x_pow QinC_ring 1` THEN have `(q:(1->num)->complex) IN S` [IN] THEN have `x_pow QinC_ring 1 IN S` [IN] THEN subgoal `distinct_minpolys S` THENL [ rw[distinct_minpolys] THEN GEN_TAC THEN DISCH_TAC THEN proven_if `p = q:(1->num)->complex` [distinct_minpolys] THEN have `p = x_pow QinC_ring 1` [IN] THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [x_pow_poly] THEN have `poly_deg QinC_ring (x_pow QinC_ring 1) = 1` [deg_x_pow;ring_1_0_QinC] THEN have `ring_prime(x_poly QinC_ring) p` [prime_if_deg_1;field_QinC] THEN have `ring_irreducible(x_poly QinC_ring) p` [prime_iff_irreducible_over_field;field_QinC] THEN qed[monic_x_pow] ; pass ] THEN have `x_pow QinC_ring 1 = x_pow complex_ring 1` [x_pow_QinC_eq_x_pow_complex] THEN subgoal `complex_root (x_pow QinC_ring 1) (Cx(&0))` THENL [ rw[complex_root] THEN simp[eval_x_pow;RING_POW_1;in_complex_ring] ; pass ] THEN qed[distinct_minpolys_distinct_roots] );; let monic_factorization_distinct_minpolys = prove(` !f. ring_polynomial QinC_ring f ==> monic QinC_ring f ==> ?P e. ( FINITE P /\ distinct_minpolys P /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = f ) `, intro THEN rw[distinct_minpolys] THEN recall field_QinC THEN specialize[`QinC_ring`;`f:(1->num)->complex`]monic_factorization_exponents THEN choose2 `P:((1->num)->complex)->bool` `e:((1->num)->complex)->num` `FINITE P /\ (!p. p IN P ==> ring_polynomial QinC_ring p) /\ (!p. p IN P ==> monic QinC_ring p) /\ (!p. p IN P ==> ring_irreducible (x_poly QinC_ring) p) /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = f` [] THEN witness `P:((1->num)->complex)->bool` THEN witness `e:((1->num)->complex)->num` THEN qed[] );; let multi_monic_factorization_distinct_minpolys = prove(` !S f. FINITE S ==> (!s:X. s IN S ==> ring_polynomial QinC_ring (f s)) ==> (!s:X. s IN S ==> monic QinC_ring (f s)) ==> ?P e. ( FINITE P /\ distinct_minpolys P /\ (!s. s IN S ==> poly_product QinC_ring P (\p. poly_pow QinC_ring p (e s p)) = f s ) ) `, intro THEN subgoal `!s:X. s IN S ==> ?P e. FINITE P /\ distinct_minpolys P /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = f s` THENL [ intro THEN specialize_assuming[`f(s:X):(1->num)->complex`]monic_factorization_distinct_minpolys THEN qed[] ; pass ] THEN def `Q:X->((1->num)->complex)->bool` `\s:X. @P. ?e. FINITE P /\ distinct_minpolys P /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = f s` THEN subgoal `!s:X. s IN S ==> ?d. FINITE (Q s) /\ distinct_minpolys (Q s) /\ (!p. p IN (Q s) ==> ~(d p = 0)) /\ poly_product QinC_ring (Q s) (\p. poly_pow QinC_ring p (d p)) = f s` THENL [ specialize[ `S:X->bool`; `\s:X P. (?d. FINITE P /\ distinct_minpolys P /\ (!p. p IN P ==> ~(d p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (d p)) = f s)` ]select_foreach THEN simp[] ; pass ] THEN have `!s:X. s IN S ==> FINITE (Q s:((1->num)->complex)->bool)` [] THEN have `!s:X. s IN S ==> distinct_minpolys (Q s:((1->num)->complex)->bool)` [] THEN def `P:((1->num)->complex)->bool` `UNIONS { Q s:((1->num)->complex)->bool| s:X IN S }` THEN subgoal `!s:X. s IN S ==> ?e. poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = f s:(1->num)->complex` THENL [ have `!s:X. s IN S ==> ?d. ((!p. p IN (Q s) ==> ~(d p = 0)) /\ poly_product QinC_ring (Q s) (\p. poly_pow QinC_ring p (d p)) = f s)` [] THEN specialize[ `S:X->bool`; `\s:X d. (!p. p IN (Q s) ==> ~(d p = 0)) /\ poly_product QinC_ring (Q s) (\p. poly_pow QinC_ring p (d p)) = f s` ]select_foreach THEN intro THEN choose `d:((1->num)->complex)->num` `poly_product QinC_ring (Q s) (\p. poly_pow QinC_ring p (d p)) = f(s:X):(1->num)->complex` [] THEN witness `\p:(1->num)->complex. if p IN Q(s:X) then d p else 0` THEN rw[BETA_THM] THEN subgoal `poly_product QinC_ring P (\p. poly_pow QinC_ring p (if p IN Q(s:X) then d p else 0)) = poly_product QinC_ring P (\p. if p IN Q s then poly_pow QinC_ring p (d p) else poly_1 QinC_ring)` THENL [ sufficesby poly_product_eq THEN qed[poly_pow_0] ; pass ] THEN rw[know `poly_product QinC_ring P (\p. poly_pow QinC_ring p (if p IN Q(s:X) then d p else 0)) = poly_product QinC_ring P (\p. if p IN Q s then poly_pow QinC_ring p (d p) else poly_1 QinC_ring)`] THEN set_fact_using `Q(s:X) SUBSET P:((1->num)->complex)->bool` [know `s:X IN S`;know `P = UNIONS {Q s:((1->num)->complex)->bool | s:X IN S}`] THEN specialize[ `QinC_ring`; `Q(s:X):((1->num)->complex)->bool`; `P:((1->num)->complex)->bool`; `\p:(1->num)->complex. poly_pow QinC_ring p (d p)` ]poly_product_restrict_subset THEN qed[] ; pass ] THEN def `e:X->((1->num)->complex)->num` `\s:X. @d. poly_product QinC_ring P (\p. poly_pow QinC_ring p (d p)) = f s:(1->num)->complex` THEN witness `P:((1->num)->complex)->bool` THEN witness `e:X->((1->num)->complex)->num` THEN intro THENL [ rw[know `P = UNIONS {Q s:((1->num)->complex)->bool | s:X IN S}`] THEN rw[FINITE_UNIONS;IN_ELIM_THM] THEN once_rw[SIMPLE_IMAGE_GEN] THEN rw[IN_GSPEC] THEN specialize[ `Q:X->((1->num)->complex)->bool`; `S:X->bool` ]FINITE_IMAGE THEN qed[] ; rw[distinct_minpolys] THEN rw[know `P = UNIONS {Q s:((1->num)->complex)->bool | s:X IN S}`] THEN rw[IN_UNIONS;IN_ELIM_THM] THEN qed[distinct_minpolys] ; specialize[ `S:X->bool`; `\s:X d. poly_product QinC_ring P (\p. poly_pow QinC_ring p (d p)) = f s` ]select_foreach THEN rw[know `e = (\s:X. @d. poly_product QinC_ring P (\p. poly_pow QinC_ring p (d p)) = f s)`] THEN qed[] ] );; (* ===== more on binomial coefficients *) let binom_coeff = prove(` !n d. coeff d ( poly_pow integer_ring ( poly_add integer_ring (x_pow integer_ring 0) (x_pow integer_ring 1) ) n ) = &(binom(n,d)):int `, INDUCT_TAC THENL [ rw[poly_pow_0] THEN rw[BINOM_0;coeff_poly_1] THEN rw[INTEGER_RING] THEN qed[] ; intro THEN rw[ARITH_RULE `SUC n = 1+n`] THEN simp[poly_pow_add;RING_POWERSERIES_ADD;x_pow_series] THEN rw[coeff_poly_mul_oneindex] THEN simp[poly_pow_1;RING_POWERSERIES_ADD;x_pow_series] THEN rw[coeff_poly_add;coeff_x_pow] THEN subgoal `ring_sum integer_ring (0..d) (\a. ring_mul integer_ring (ring_add integer_ring (if a = 0 then ring_1 integer_ring else ring_0 integer_ring) (if a = 1 then ring_1 integer_ring else ring_0 integer_ring)) (&(binom (n,d - a)))) = ring_sum integer_ring (0..d) (\a. if a = 0 then &(binom(n,d)) else if a = 1 then &(binom(n,d-1)) else ring_0 integer_ring)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN case `a = 0` THENL [ simp[ARITH_RULE `~(0 = 1)`] THEN simp[RING_ADD_RZERO;RING_1] THEN qed[RING_MUL_LID;RING_OF_NUM;INTEGER_RING_OF_NUM;ARITH_RULE `d - 0 = d`] ; pass ] THEN case `a = 1` THENL [ simp[RING_ADD_LZERO;RING_1] THEN qed[RING_MUL_LID;RING_OF_NUM;INTEGER_RING_OF_NUM] ; pass ] THEN simp[RING_ADD_LZERO;RING_0] THEN qed[RING_MUL_LZERO;RING_OF_NUM;INTEGER_RING_OF_NUM] ; pass ] THEN simp[] THEN case `d = 0` THENL [ subgoal `ring_sum integer_ring (0..d) (\a. if a = 0 then &(binom (n,d)) else if a = 1 then &(binom (n,d - 1)) else ring_0 integer_ring) = ring_sum integer_ring (0..d) (\a. if a = 0 then &(binom (n,d)) else ring_0 integer_ring)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN have `a <= d:num` [IN_NUMSEG] THEN num_linear_fact `a <= d ==> d = 0 ==> ~(a = 1)` THEN qed[] ; pass ] THEN have `0 IN 0..0` [IN_NUMSEG_0;ARITH_RULE `0 <= 0`] THEN have `&(binom (n,0)) IN ring_carrier integer_ring` [RING_OF_NUM;INTEGER_RING_OF_NUM] THEN simp[RING_SUM_DELTA] THEN qed[binom] ; pass ] THEN num_linear_fact `0 <= d` THEN have `0 IN 0..d` [IN_NUMSEG_0] THEN num_linear_fact `~(d = 0) ==> 1 <= d` THEN have `FINITE (0..d)` [FINITE_NUMSEG] THEN have `1 IN 0..d` [IN_NUMSEG_0] THEN have `&(binom (n,d)) IN ring_carrier integer_ring` [RING_OF_NUM;INTEGER_RING_OF_NUM] THEN have `&(binom (n,d-1)) IN ring_carrier integer_ring` [RING_OF_NUM;INTEGER_RING_OF_NUM] THEN simp[ring_sum_delta_delta;ARITH_RULE `~(0 = 1)`] THEN rw[INTEGER_RING] THEN num_linear_fact `1+n = SUC n` THEN def `d1:num` `d-1` THEN num_linear_fact `~(d = 0) ==> d1 = d-1 ==> d = SUC d1` THEN rw[know `d = SUC d1`] THEN rw[know `1 + n = SUC n`] THEN rw[binom] THEN rw[ARITH_RULE `SUC d1 - 1 = d1`] THEN qed[INT_OF_NUM_ADD] ] );; let binom_sum = prove(` !m n d. binom(m+n,d) = nsum(0..d) (\i. binom(m,i) * binom(n,d-i)) `, intro THEN rw[GSYM INT_OF_NUM_EQ] THEN rw[GSYM binom_coeff] THEN simp[poly_pow_add;RING_POWERSERIES_ADD;x_pow_series] THEN rw[coeff_poly_mul_oneindex] THEN have `FINITE (0..d)` [FINITE_NUMSEG] THEN simp[int_of_num_sum] THEN simp[isum_integer_sum] THEN sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[GSYM INT_OF_NUM_MUL] THEN rw[GSYM binom_coeff] THEN rw[INTEGER_RING] );; let binom_rowsum = prove(` !m. nsum(0..m) (\i. binom(m,i)) = 2 EXP m `, intro THEN specialize[`m:num`;`1`;`1`]BINOMIAL_THEOREM THEN subgoal `nsum (0..m) (\k. binom (m,k) * 1 EXP k * 1 EXP (m - k)) = nsum (0..m) (\k. binom (m,k))` THENL [ sufficesby NSUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[EXP_ONE] THEN rw[MULT_CLAUSES] ; pass ] THEN qed[ARITH_RULE `1+1 = 2`] );; let binom_rowsum_partial = prove(` !m k. nsum(0..k) (\i. binom(m,i)) <= 2 EXP m `, intro THEN have `2 EXP m = nsum(0..m) (\i. binom(m,i))` [binom_rowsum] THEN case `k <= m:num` THENL [ have `(0..k) SUBSET (0..m)` [SUBSET_NUMSEG;ARITH_RULE `0 <= 0`] THEN qed[NSUM_SUBSET_SIMPLE;FINITE_NUMSEG] ; simp[] THEN sufficesby NSUM_SUBSET THEN simp[FINITE_NUMSEG] THEN rw[IN_DIFF;IN_NUMSEG] THEN intro THEN num_linear_fact `~(0 <= x /\ x <= m) ==> m < x` THEN qed[BINOM_LT] ] );; (* ===== bounds on (reversed) coefficients in terms of roots *) let coeff_root_bound_1 = prove(` !A:real d. norm(coeff d (poly_1 complex_ring)) <= A pow d * &(binom(0,d)) `, intro THEN rw[coeff_poly_1] THEN case `d = 0` THENL [ simp[complex_ring_clauses] THEN rw[real_pow;binom;COMPLEX_NORM_CX] THEN real_field_fact `abs(&1) <= &1 * &1` THEN qed[] ; pass ] THEN simp[complex_ring_clauses] THEN rw[COMPLEX_NORM_CX] THEN num_linear_fact `~(d = 0) ==> 0 < d` THEN simp[BINOM_LT] THEN real_field_fact `abs(&0):real <= A pow d * &0` THEN qed[] );; let coeff_root_bound_one_minus_constx = prove(` !c:complex A:real d:num. norm(c) <= A ==> norm(coeff d (one_minus_constx complex_ring c)) <= A pow d * &(binom(1,d)) `, intro THEN simp[coeff_one_minus_constx;in_complex_ring] THEN case `d = 0` THENL [ simp[complex_ring_clauses] THEN rw[real_pow;binom;COMPLEX_NORM_CX] THEN real_field_fact `abs(&1) <= &1 * &1` THEN qed[] ; pass ] THEN case `d = 1` THENL [ simp[complex_ring_clauses] THEN rw[REAL_POW_1;BINOM_REFL;NORM_NEG] THEN real_field_fact `A = A * &1:real` THEN qed[] ; pass ] THEN simp[complex_ring_clauses] THEN rw[COMPLEX_NORM_CX] THEN num_linear_fact `~(d = 0) ==> ~(d = 1) ==> 1 < d` THEN simp[BINOM_LT] THEN real_field_fact `abs(&0):real <= A pow d * &0` THEN qed[] );; let coeff_root_bound_mul = prove(` !p q (A:real) m n. (!d. norm(coeff d p) <= A pow d * &(binom(m,d))) ==> (!d. norm(coeff d q) <= A pow d * &(binom(n,d))) ==> (!d. norm(coeff d (poly_mul complex_ring p q)) <= A pow d * &(binom(m+n,d))) `, intro THEN rw[coeff_poly_mul_oneindex] THEN rw[complex_ring_clauses] THEN simp[GSYM vsum_ring_sum_complex;FINITE_NUMSEG] THEN subgoal `!a. a IN (0..d) ==> norm(coeff a p * coeff (d-a) q) <= A pow d * &(binom(m,a)*binom(n,d-a))` THENL [ intro THEN rw[GSYM REAL_OF_NUM_MUL] THEN have `norm(coeff a (p:(1->num)->complex) * coeff (d-a) (q:(1->num)->complex)) = norm(coeff a p) * norm(coeff (d-a) q)` [COMPLEX_NORM_MUL] THEN have `norm(coeff a (p:(1->num)->complex)) * norm(coeff (d-a) (q:(1->num)->complex)) <= (A pow a * &(binom(m,a))) * (A pow (d-a) * &(binom(n,d-a)))` [REAL_LE_MUL2;NORM_POS_LE] THEN have `a <= d:num` [IN_NUMSEG] THEN num_linear_fact `a <= d:num ==> d = a + (d-a)` THEN have `(A:real) pow d = A pow a * A pow (d-a)` [REAL_POW_ADD] THEN real_linear_fact `((A:real) pow a * &(binom (m,a))) * A pow (d - a) * &(binom (n,d - a)) = (A pow a * A pow (d-a)) * &(binom (m,a)) * &(binom (n,d - a))` THEN qed[] ; pass ] THEN subgoal `norm(vsum(0..d) (\a. coeff a p * coeff (d - a) q)) <= sum(0..d) (\a. A pow d * &(binom (m,a) * binom (n,d - a)))` THENL [ sufficesby VSUM_NORM_LE THEN qed[FINITE_NUMSEG] ; pass ] THEN specialize[`\a. &(binom (m,a) * binom (n,d - a)):real`;`(A:real) pow d`;`0..d`]SUM_LMUL THEN have `norm(vsum(0..d) (\a. coeff a p * coeff (d - a) q)) <= (A:real) pow d * sum(0..d) (\a. &(binom (m,a) * binom (n,d - a)))` [] THEN rw[binom_sum] THEN simp[REAL_OF_NUM_SUM;FINITE_NUMSEG] );; let coeff_root_bound_product = prove(` !p (A:real) m S. FINITE S ==> (!s:X d. s IN S ==> norm(coeff d (p s)) <= A pow d * &(binom(m s,d))) ==> (!d. norm(coeff d (poly_product complex_ring S p)) <= A pow d * &(binom(nsum S m,d))) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;NSUM_CLAUSES] THEN qed[coeff_root_bound_1] ; have `(!s:X. s IN x INSERT S ==> ring_powerseries complex_ring (p s:(1->num)->complex))` [series_complex] THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[poly_product_insert;NSUM_CLAUSES] THEN qed[coeff_root_bound_mul] ] );; (* or can prove this via coeff_root_bound_product *) let coeff_root_bound_pow = prove(` !p (A:real) m n. ring_powerseries complex_ring p ==> (!d. norm(coeff d p) <= A pow d * &(binom(m,d))) ==> (!d. norm(coeff d (poly_pow complex_ring p n)) <= A pow d * &(binom(m*n,d))) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0] THEN rw[ARITH_RULE `m * 0 = 0`] THEN qed[coeff_root_bound_1] ; intro THEN rw[ARITH_RULE `SUC n = n + 1`] THEN simp[poly_pow_add;poly_pow_1] THEN specialize_assuming[`poly_pow complex_ring (p:(1->num)->complex) n`;`p:(1->num)->complex`;`A:real`;`m*n:num`;`m:num`]coeff_root_bound_mul THEN qed[ARITH_RULE `m*n+m = m*(n+1)`] ] );; (* ===== first-order Stirling lower bound *) let factorial_lower_bound = prove(` !n. (&n / exp(&1)) pow n <= &(FACT n) `, INDUCT_TAC THENL [ rw[FACT] THEN CONV_TAC REAL_FIELD ; pass ] THEN rw[FACT;real_pow] THEN rw[ARITH_RULE `SUC n = n+1`] THEN case `n = 0` THENL [ simp[real_pow;ARITH_RULE `0+1 = 1`] THEN rw[FACT;ARITH_RULE `1*1 = 1`] THEN real_linear_fact `&0 <= &1:real` THEN have `exp(&0) <= exp(&1)` [REAL_EXP_MONO_LE] THEN have `&1 <= exp(&1)` [REAL_EXP_0] THEN have `inv(exp(&1)) <= &1` [REAL_INV_LE_1] THEN real_linear_fact `inv(exp(&1)) <= &1 ==> &1 / exp(&1) * &1 <= &1` THEN qed[] ; pass ] THEN have `&1 + &1 / &n <= exp(&1 / &n)` [REAL_EXP_LE_X] THEN have `~(&n = &0:real)` [REAL_OF_NUM_EQ] THEN real_field_fact `~(&n = &0) ==> (&n + &1) / &n:real = &1 + &1 / &n` THEN have `(&n + &1) / &n <= exp(&1 / &n)` [] THEN have `&0 <= &n:real` [REAL_OF_NUM_LE;ARITH_RULE `0 <= n`] THEN real_linear_fact `&0 <= &n:real ==> &0 <= &n + &1` THEN have `&0 <= (&n + &1) / &n` [REAL_LE_DIV] THEN have `&0 <= exp(&1)` [REAL_EXP_POS_LE] THEN have `((&n + &1) / &n) pow n <= exp(&1 / &n) pow n` [REAL_POW_LE2] THEN subgoal `exp(&1 / &n) pow n = exp(&1)` THENL [ rw[GSYM REAL_EXP_N] THEN real_field_fact `~(&n = &0:real) ==> &n * &1 / &n = &1` THEN qed[] ; pass ] THEN have `((&n + &1) / &n) pow n <= exp(&1)` [] THEN have `&0 <= &n / exp(&1)` [REAL_LE_DIV] THEN subgoal `(&n / exp (&1)) pow n * ((&n + &1) / &n) pow n <= &(FACT n) * exp(&1)` THENL [ have `&0 <= (&n / exp(&1)) pow n` [REAL_POW_LE] THEN have `&0 <= ((&n + &1) / &n) pow n` [REAL_POW_LE] THEN qed[REAL_LE_MUL2] ; pass ] THEN subgoal `(&n / exp (&1)) pow n * ((&n + &1) / &n) pow n = ((&n + &1) / exp(&1)) pow n` THENL [ rw[GSYM REAL_POW_MUL] THEN real_field_fact `~(&n = &0:real) ==> (&n / exp (&1) * (&n + &1) / &n) = ((&n + &1) / exp (&1))` THEN qed[] ; pass ] THEN have `((&n + &1) / exp (&1)) pow n <= &(FACT n) * exp(&1)` [] THEN have `&0 <= (&n + &1) / exp(&1)` [REAL_LE_DIV] THEN have `(&n + &1) / exp (&1) * ((&n + &1) / exp (&1)) pow n <= (&n + &1) / exp (&1) * &(FACT n) * exp(&1)` [REAL_LE_LMUL] THEN have `&0 < exp(&1)` [REAL_EXP_POS_LT] THEN real_field_fact `&0 < exp(&1) ==> (&n + &1) / exp (&1) * &(FACT n) * exp (&1) = (&n + &1) * &(FACT n)` THEN have `(&n + &1) / exp (&1) * ((&n + &1) / exp (&1)) pow n <= (&n + &1) * &(FACT n)` [] THEN simp[GSYM REAL_OF_NUM_ADD;GSYM REAL_OF_NUM_MUL] );; (* ===== poly_ord p z: order of vanishing of polynomial p at z *) let poly_ord = new_definition ` poly_ord (p:(1->num)->complex) (z:complex) = @e:num. ( ?q:(1->num)->complex. ( ring_polynomial complex_ring q /\ ~(complex_root q z) /\ p = poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) e ) ( q ) ) ) `;; let poly_ord_exists_lemma = prove(` !n p:(1->num)->complex. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> poly_deg complex_ring p = n ==> ?e:num q:(1->num)->complex. ( ring_polynomial complex_ring q /\ ~(complex_root q z) /\ p = poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) e ) ( q ) ) `, INDUCT_TAC THENL [ intro THEN witness `0` THEN witness `p:(1->num)->complex` THEN rw[poly_pow_0] THEN simp[POLY_MUL_LID;series_complex] THEN choose `c:complex` `(p:(1->num)->complex) = poly_const complex_ring c` [POLY_DEG_EQ_0] THEN have `poly_eval complex_ring p z = c` [POLY_EVAL_CONST;in_complex_ring] THEN have `~(c = ring_0 complex_ring)` [poly_0] THEN have `~(c = Cx(&0))` [complex_ring_clauses] THEN qed[complex_root] ; pass ] THEN intro THEN case `~(complex_root p z)` THENL [ witness `0` THEN witness `p:(1->num)->complex` THEN rw[poly_pow_0] THEN simp[POLY_MUL_LID;series_complex] ; pass ] THEN have `poly_eval complex_ring p z = ring_0 complex_ring` [complex_root;complex_ring_clauses] THEN have `z IN ring_carrier complex_ring` [in_complex_ring] THEN have `p IN ring_carrier(poly_ring complex_ring (:1))` [x_poly;x_poly_use] THEN specialize[`complex_ring`;`z:complex`;`p:(1->num)->complex`]POLY_DIVIDES_X_MINUS_ROOT THEN have `poly_sub complex_ring (poly_var complex_ring one) (poly_const complex_ring z) = x_minus_const complex_ring z` [x_minus_const;x_pow_1] THEN have `ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) p` [x_poly] THEN choose `q:(1->num)->complex` `q IN ring_carrier(x_poly complex_ring) /\ (p:(1->num)->complex) = ring_mul(x_poly complex_ring) (x_minus_const complex_ring z) q` [ring_divides] THEN have `integral_domain complex_ring` [integral_domain_complex] THEN have `poly_deg complex_ring (x_minus_const complex_ring z) = 1` [deg_x_minus_const;integral_domain] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly] THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [x_poly_use] THEN have `p = poly_mul complex_ring (x_minus_const complex_ring z) q` [x_poly_use] THEN have `~(q = poly_0 complex_ring:(1->num)->complex)` [POLY_MUL_0] THEN have `~(x_minus_const complex_ring z = poly_0 complex_ring)` [POLY_DEG_0;ARITH_RULE `~(0 = 1)`] THEN have `(x_minus_const complex_ring z = poly_0 complex_ring) <=> (q = poly_0 complex_ring:(1->num)->complex)` [] THEN specialize[`complex_ring`;`x_minus_const complex_ring z`;`q:(1->num)->complex`]POLY_DEG_MUL THEN have `poly_deg complex_ring (p:(1->num)->complex) = poly_deg complex_ring (x_minus_const complex_ring z) + poly_deg complex_ring (q:(1->num)->complex)` [] THEN have `SUC n = 1 + poly_deg complex_ring (q:(1->num)->complex)` [] THEN num_linear_fact `SUC n = 1 + poly_deg complex_ring (q:(1->num)->complex) ==> poly_deg complex_ring q = n` THEN specialize[`q:(1->num)->complex`](know(`!p. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> poly_deg complex_ring p = n ==> (?e q. ring_polynomial complex_ring q /\ ~complex_root q z /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) e) q)`)) THEN choose2 `e:num` `r:(1->num)->complex` `ring_polynomial complex_ring r /\ ~complex_root r z /\ q = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) e) r` [] THEN witness `1+e` THEN witness `r:(1->num)->complex` THEN SIMP_TAC[poly_pow_add;poly_pow_1;series_complex] THEN SIMP_TAC[GSYM POLY_MUL_ASSOC;series_complex] THEN qed[] );; let poly_ord_exists_lemma2 = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> ?e q:(1->num)->complex. ( ring_polynomial complex_ring q /\ ~(complex_root q z) /\ p = poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) e ) ( q ) ) `, intro THEN specialize[`poly_deg complex_ring (p:(1->num)->complex)`;`p:(1->num)->complex`]poly_ord_exists_lemma THEN qed[] );; let poly_ord_exists = prove(` !p:(1->num)->complex. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> ?q:(1->num)->complex. ( ring_polynomial complex_ring q /\ ~(complex_root q z) /\ p = poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) (poly_ord p z) ) ( q ) ) `, rw[poly_ord] THEN qed[poly_ord_exists_lemma2] );; let poly_ord_unique = prove(` !p:(1->num)->complex q e. ring_polynomial complex_ring q ==> ~(complex_root q z) ==> p = poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) e ) ( q ) ==> ( ring_polynomial complex_ring p /\ ~(p = poly_0 complex_ring) /\ e = poly_ord p z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN subgoal `~((q:(1->num)->complex) = poly_0 complex_ring)` THENL [ intro THEN have `poly_eval complex_ring q z = Cx(&0)` [POLY_EVAL_0;complex_ring_clauses] THEN qed[complex_root] ; pass ] THEN subgoal `~((p:(1->num)->complex) = poly_0 complex_ring)` THENL [ intro THEN have `p = ring_mul(x_poly complex_ring) (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) e) q` [x_poly_use;x_poly_use_pow] THEN have `integral_domain(x_poly complex_ring)` [integral_domain_x_poly_field;field_complex] THEN have `p = ring_0(x_poly complex_ring)` [x_poly_use] THEN have `~(q = ring_0(x_poly complex_ring))` [x_poly_use] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) e)` [x_poly_use_pow;poly_pow_poly] THEN have `ring_mul (x_poly complex_ring) (ring_pow (x_poly complex_ring) (x_minus_const complex_ring z) e) q = ring_0(x_poly complex_ring)` [] THEN have `ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) e = ring_0(x_poly complex_ring)` [integral_domain;x_poly_use] THEN have `x_minus_const complex_ring z = ring_0(x_poly complex_ring)` [INTEGRAL_DOMAIN_POW_EQ_0;x_poly_use] THEN have `x_minus_const complex_ring z = poly_0 complex_ring` [x_poly_use] THEN have `poly_deg complex_ring (x_minus_const complex_ring z) = 1` [deg_x_minus_const;integral_domain;integral_domain_complex;in_complex_ring] THEN have `~(x_minus_const complex_ring z = poly_0 complex_ring)` [POLY_DEG_0;ARITH_RULE `~(0 = 1)`] THEN qed[] ; pass ] THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) e)` [poly_pow_poly] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [RING_POLYNOMIAL_MUL] THEN specialize[`p:(1->num)->complex`]poly_ord_exists THEN choose `f:(1->num)->complex` `ring_polynomial complex_ring f /\ ~(complex_root f z) /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (f)` [] THEN recall field_complex THEN have `poly_deg complex_ring (x_minus_const complex_ring z) = 1` [deg_x_minus_const;field;in_complex_ring] THEN have `ring_prime(x_poly complex_ring) (x_minus_const complex_ring z)` [prime_if_deg_1] THEN have `integral_domain(x_poly complex_ring)` [integral_domain_x_poly_field] THEN have `f IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN have `q IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN have `~(ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) q)` [complex_root_if_x_minus_const_divides] THEN have `~(ring_divides(x_poly complex_ring) (x_minus_const complex_ring z) f)` [complex_root_if_x_minus_const_divides] THEN have `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) e) q = ring_mul(x_poly complex_ring) (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) (poly_ord p z)) f` [x_poly_use;x_poly_use_pow] THEN have `poly_mul complex_ring (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) e) q = ring_mul(x_poly complex_ring) (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) (poly_ord p z)) f` [x_poly_use;x_poly_use_pow] THEN have `ring_mul(x_poly complex_ring) (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) e) q = ring_mul(x_poly complex_ring) (ring_pow(x_poly complex_ring) (x_minus_const complex_ring z) (poly_ord p z)) f` [x_poly_use;x_poly_use_pow] THEN qed[unique_prime_valuation] );; let poly_ord_unique_0 = prove(` !p:(1->num)->complex z. ring_polynomial complex_ring p ==> ~(complex_root p z) ==> ( ~(p = poly_0 complex_ring) /\ poly_ord p z = 0 ) `, intro THEN have `p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) 0) (p)` [poly_pow_0;POLY_MUL_LID;ring_polynomial] THEN qed[poly_ord_unique] );; let poly_ord_unique_1 = prove(` !p:(1->num)->complex q z. ring_polynomial complex_ring q ==> ~(complex_root q z) ==> p = poly_mul complex_ring ( x_minus_const complex_ring z ) ( q ) ==> ( ring_polynomial complex_ring p /\ ~(p = poly_0 complex_ring) /\ poly_ord p z = 1 ) `, intro THEN have `p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) 1) (q)` [poly_pow_1;series_complex] THEN qed[poly_ord_unique] );; let poly_ord_const = prove(` !c z. ~(c = Cx(&0)) ==> ( ring_polynomial complex_ring (poly_const complex_ring c:(1->num)->complex) /\ ~(poly_const complex_ring c = poly_0 complex_ring:(1->num)->complex) /\ poly_ord(poly_const complex_ring c) z = 0 ) `, REPEAT GEN_TAC THEN DISCH_TAC THEN have `ring_polynomial complex_ring (poly_const complex_ring c:(1->num)->complex)` [RING_POLYNOMIAL_CONST;in_complex_ring] THEN have `~complex_root (poly_const complex_ring c) z` [complex_root;POLY_EVAL_CONST;in_complex_ring] THEN qed[poly_ord_unique_0] );; let poly_ord_1 = prove(` !z. ring_polynomial complex_ring (poly_1 complex_ring:(1->num)->complex) /\ ~(poly_1 complex_ring = poly_0 complex_ring:(1->num)->complex) /\ poly_ord(poly_1 complex_ring) z = 0 `, rw[poly_1] THEN complex_field_fact `~(Cx(&1) = Cx(&0))` THEN have `~(ring_1 complex_ring = Cx(&0))` [complex_ring_clauses] THEN specialize[`ring_1 complex_ring`]poly_ord_const THEN qed[poly_1_0_complex;complex_ring_clauses] );; let poly_ord_x_minus_const = prove(` !c z. ring_polynomial complex_ring (x_minus_const complex_ring c) /\ ~(x_minus_const complex_ring c = poly_0 complex_ring) /\ poly_ord (x_minus_const complex_ring c) z = if z = c then 1 else 0 `, REPEAT GEN_TAC THEN have `ring_polynomial complex_ring (x_minus_const complex_ring c)` [x_minus_const_poly;in_complex_ring] THEN case `z = c:complex` THENL [ specialize_assuming[`x_minus_const complex_ring c`;`poly_1 complex_ring:(1->num)->complex`;`z:complex`]poly_ord_unique_1 THEN have `~complex_root (poly_1 complex_ring) z` [complex_root;POLY_EVAL_1;field_complex;field;complex_ring_clauses] THEN have `x_minus_const complex_ring c = poly_mul complex_ring (x_minus_const complex_ring c) (poly_1 complex_ring)` [RING_POLYNOMIAL_1;complex_root;POLY_MUL_SYM;POLY_MUL_LID;series_complex] THEN qed[RING_POLYNOMIAL_1] ; pass ] THEN have `poly_eval complex_ring (x_minus_const complex_ring c) z = z - c` [eval_x_minus_const;in_complex_ring;ring_sub_complex] THEN complex_field_fact `~(z = c:complex) ==> ~(z - c = Cx(&0))` THEN have `~(poly_eval complex_ring (x_minus_const complex_ring c) z = Cx(&0))` [] THEN have `~(complex_root(x_minus_const complex_ring c) z)` [complex_root] THEN qed[poly_ord_unique_0] );; let poly_ord_mul = prove(` !p:(1->num)->complex q z. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> ~(p = poly_0 complex_ring) ==> ~(q = poly_0 complex_ring) ==> ( ring_polynomial complex_ring (poly_mul complex_ring p q) /\ ~(poly_mul complex_ring p q = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring p q) z = poly_ord p z + poly_ord q z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN choose `f:(1->num)->complex` `ring_polynomial complex_ring f /\ ~complex_root f z /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) f` [poly_ord_exists] THEN choose `g:(1->num)->complex` `ring_polynomial complex_ring g /\ ~complex_root g z /\ q = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) g` [poly_ord_exists] THEN subgoal `~(complex_root (poly_mul complex_ring f g) z)` THENL [ rw[complex_root] THEN simp[POLY_EVAL_MUL;in_complex_ring] THEN have `~(poly_eval complex_ring f z = Cx(&0))` [complex_root] THEN have `~(poly_eval complex_ring g z = Cx(&0))` [complex_root] THEN qed[COMPLEX_ENTIRE;complex_ring_clauses] ; pass ] THEN have `p = ring_mul(x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) f` [x_poly_use] THEN have `q = ring_mul(x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) g` [x_poly_use] THEN specialize_assuming[`x_poly(complex_ring)`; `poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)`; `poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)`; `f:(1->num)->complex`; `g:(1->num)->complex`; `p:(1->num)->complex`; `q:(1->num)->complex`; ](GENL[ `r:R ring`;`d:R`;`e:R`;`f:R`;`g:R`;`p:R`;`q:R` ]( RING_RULE `p = ring_mul(r:R ring) d f ==> q = ring_mul r e g ==> ring_mul r p q = ring_mul r (ring_mul r d e) (ring_mul r f g)` )) THEN have `ring_mul (x_poly complex_ring) p q = ring_mul (x_poly complex_ring) (ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (ring_mul (x_poly complex_ring) f g)` [x_poly_use;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `ring_mul (x_poly complex_ring) p q = ring_mul (x_poly complex_ring) (ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring f g)` [x_poly_use] THEN have `ring_mul (x_poly complex_ring) p q = ring_mul (x_poly complex_ring) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring f g)` [x_poly_use] THEN have `ring_mul (x_poly complex_ring) p q = poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring f g)` [x_poly_use] THEN have `poly_mul complex_ring p q = poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring f g)` [x_poly_use] THEN have `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) = poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)` [poly_pow_add;series_complex] THEN have `poly_mul complex_ring p q = poly_mul complex_ring ((poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z))) (poly_mul complex_ring f g)` [] THEN specialize[`poly_mul complex_ring p q:(1->num)->complex`;`poly_mul complex_ring f g:(1->num)->complex`]poly_ord_unique THEN qed[RING_POLYNOMIAL_MUL] );; let poly_ord_product = prove(` !p z S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial complex_ring (p s)) ==> (!s:X. s IN S ==> ~(p s = poly_0 complex_ring)) ==> ( ring_polynomial complex_ring (poly_product complex_ring S p) /\ ~(poly_product complex_ring S p = poly_0 complex_ring) /\ poly_ord (poly_product complex_ring S p) z = nsum S (\s. poly_ord (p s) z) ) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN STRIP_TAC THEN STRIP_TAC THEN STRIP_TAC THENL [ rw[NSUM_CLAUSES;poly_product_empty] THEN rw[poly_ord_1;poly_1_0_complex] ; STRIP_TAC THEN STRIP_TAC THEN STRIP_TAC THEN have `(!s:X. s IN x INSERT S ==> ring_powerseries complex_ring (p s:(1->num)->complex))` [series_complex] THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN set_fact `(x:X) IN x INSERT S` THEN simp[poly_product_insert;NSUM_CLAUSES] THEN have `~(poly_product complex_ring (S:X->bool) p = poly_0(complex_ring))` [poly_product_ring_product_x_poly;x_poly_use] THEN have `~(poly_mul complex_ring (p(x:X)) (poly_product complex_ring S p) = poly_0 complex_ring)` [nonzero_poly_mul;series_complex;integral_domain_complex] THEN specialize_assuming[`p(x:X):(1->num)->complex`;`poly_product complex_ring (S:X->bool) p`;`z:complex`]poly_ord_mul THEN qed[poly_product_poly] ] );; let poly_ord_monic_vanishing_at = prove(` !S c z. FINITE S ==> ( ring_polynomial complex_ring (monic_vanishing_at complex_ring S c) /\ ~(monic_vanishing_at complex_ring S c = poly_0 complex_ring) /\ poly_ord (monic_vanishing_at complex_ring S c) z = CARD {s:X | s IN S /\ z = c s} ) `, intro THENL [ qed[monic_vanishing_at_poly;in_complex_ring] ; have `monic complex_ring (monic_vanishing_at complex_ring S (c:X->complex))` [monic_vanishing_at_monic;in_complex_ring] THEN have `~(monic complex_ring (poly_0 complex_ring))` [monic_poly_0;ring_1_0_complex] THEN qed[] ; subgoal `poly_ord (monic_vanishing_at complex_ring S c) z = nsum S (\s:X. poly_ord (x_minus_const complex_ring (c s)) z)` THENL [ rw[monic_vanishing_at] THEN specialize_assuming[ `\s:X. x_minus_const complex_ring (c s)`; `z:complex`; `S:X->bool` ]poly_ord_product THEN qed[x_minus_const_poly;in_complex_ring;monic_x_minus_const;monic_poly_0;ring_1_0_complex] ; pass ] THEN subgoal `nsum S (\s. poly_ord (x_minus_const complex_ring (c s)) z) = nsum S (\s:X. if z = c s then 1 else 0)` THENL [ sufficesby NSUM_EQ THEN qed[poly_ord_x_minus_const] ; pass ] THEN simp[] THEN rw[GSYM NSUM_RESTRICT_SET] THEN specialize[ `S:X->bool`; `\s:X. z = c s:complex` ]FINITE_RESTRICT THEN qed[CARD_EQ_NSUM] ] );; let poly_ord_monic_vanishing_at_numpreimages = prove(` !S:X->bool c z. FINITE S ==> ( ring_polynomial complex_ring (monic_vanishing_at complex_ring S c) /\ ~(monic_vanishing_at complex_ring S c = poly_0 complex_ring) /\ poly_ord (monic_vanishing_at complex_ring S c) z = numpreimages c S z ) `, simp[poly_ord_monic_vanishing_at] THEN rw[numpreimages] THEN intro THEN subgoal `{s:X | s IN S /\ z = c s:complex} = {x | x IN S /\ c x = z}` THENL [ rw[EXTENSION;IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[] );; let poly_ord_one_minus_constx = prove(` !c z. ring_polynomial complex_ring (one_minus_constx complex_ring c) /\ ~(one_minus_constx complex_ring c = poly_0 complex_ring) /\ poly_ord (one_minus_constx complex_ring c) z = if c * z = Cx(&1) then 1 else 0 `, REPEAT GEN_TAC THEN have `ring_polynomial complex_ring (one_minus_constx complex_ring c)` [one_minus_constx_poly;in_complex_ring] THEN case `c * z = Cx(&1)` THENL [ specialize_assuming[`one_minus_constx complex_ring c`;`poly_1 complex_ring:(1->num)->complex`;`z:complex`]poly_ord_unique_1 THEN have `~complex_root (poly_1 complex_ring) z` [complex_root;POLY_EVAL_1;field_complex;field;complex_ring_clauses] THEN subgoal `one_minus_constx complex_ring c = poly_mul complex_ring (x_minus_const complex_ring z) (poly_const complex_ring (--c))` THENL [ sufficesby eq_coeff THEN intro THEN simp[coeff_times_poly_const;series_complex;in_complex_ring] THEN simp[coeff_one_minus_constx;in_complex_ring] THEN simp[coeff_x_minus_const;in_complex_ring] THEN case `d = 0` THENL [ simp[ARITH_RULE `~(0 = 1)`] THEN rw[complex_ring_clauses] THEN complex_field_fact `c * z = Cx(&1) ==> Cx(&1) = --c * --z` THEN qed[] ; pass ] THEN case `d = 1` THENL [ simp[] THEN rw[complex_ring_clauses] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN qed[RING_MUL_RZERO;in_complex_ring] ; pass ] THEN subgoal `~complex_root (poly_const complex_ring (--c)) z` THENL [ rw[complex_root] THEN simp[POLY_EVAL_CONST;in_complex_ring] THEN complex_field_fact `c * z = Cx(&1) ==> ~(--c = Cx(&0))` THEN qed[] ; pass ] THEN specialize_assuming[`one_minus_constx complex_ring c`;`poly_const complex_ring (--c):(1->num)->complex`;`z:complex`]poly_ord_unique_1 THEN qed[one_minus_constx_poly;in_complex_ring;RING_POLYNOMIAL_CONST] ; pass ] THEN have `poly_eval complex_ring (one_minus_constx complex_ring c) z = Cx(&1) - c * z` [eval_one_minus_constx;in_complex_ring;ring_sub_complex;complex_ring_clauses] THEN complex_field_fact `~(c * z = Cx(&1)) ==> ~(Cx(&1) - c * z = Cx(&0))` THEN have `~(poly_eval complex_ring (one_minus_constx complex_ring c) z = Cx(&0))` [] THEN have `~(complex_root(one_minus_constx complex_ring c) z)` [complex_root] THEN qed[poly_ord_unique_0] );; let poly_ord_pow = prove(` !p:(1->num)->complex n z. ring_polynomial complex_ring p ==> ~(p = poly_0 complex_ring) ==> ( ring_polynomial complex_ring (poly_pow complex_ring p n) /\ ~(poly_pow complex_ring p n = poly_0 complex_ring) /\ poly_ord (poly_pow complex_ring p n) z = n * poly_ord p z ) `, GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;poly_1_0_complex;ARITH_RULE `0 * e = 0`] THEN qed[poly_ord_1] ; rw[ARITH_RULE `SUC n = n + 1`;RIGHT_ADD_DISTRIB;ARITH_RULE `1 * e = e`] THEN simp[poly_pow_add;poly_pow_1;series_complex] THEN GEN_TAC THEN REPEAT DISCH_TAC THEN have `ring_polynomial complex_ring (poly_pow complex_ring p n:(1->num)->complex)` [poly_pow_poly] THEN have `~(poly_pow complex_ring p n = poly_0 complex_ring:(1->num)->complex)` [] THEN specialize[`poly_pow complex_ring p n:(1->num)->complex`;`p:(1->num)->complex`;`z:complex`]poly_ord_mul THEN qed[] ] );; (* strict special case of poly_ord_x_pow *) let poly_ord_x_pow_1 = prove(` !z. ring_polynomial complex_ring (x_pow complex_ring 1) /\ ~(x_pow complex_ring 1 = poly_0 complex_ring) /\ poly_ord (x_pow complex_ring 1) z = if z = Cx(&0) then 1 else 0 `, rw[GSYM x_minus_const_0] THEN qed[poly_ord_x_minus_const;complex_ring_clauses] );; let poly_ord_x_pow = prove(` !n z. ring_polynomial complex_ring (x_pow complex_ring n) /\ ~(x_pow complex_ring n = poly_0 complex_ring) /\ poly_ord (x_pow complex_ring n) z = if z = Cx(&0) then n else 0 `, REPEAT GEN_TAC THEN have `x_pow complex_ring n = poly_pow complex_ring (x_pow complex_ring 1) n` [x_pow_pow;ARITH_RULE `1*n = n`] THEN simp[] THEN have `ring_polynomial complex_ring (x_pow complex_ring 1) /\ ~(x_pow complex_ring 1 = poly_0 complex_ring) /\ poly_ord (x_pow complex_ring 1) z = if z = Cx(&0) then 1 else 0` [poly_ord_x_pow_1] THEN specialize_assuming[ `x_pow complex_ring 1`; `n:num`; `z:complex` ]poly_ord_pow THEN qed[ARITH_RULE `n*1 = n`;ARITH_RULE `n*0 = 0`] );; let poly_ord_mul_0 = prove(` !p:(1->num)->complex q z. ( ring_polynomial complex_ring p /\ ~(p = poly_0 complex_ring) /\ poly_ord p z = 0 ) ==> ( ring_polynomial complex_ring q /\ ~(q = poly_0 complex_ring) /\ poly_ord q z = 0 ) ==> ( ring_polynomial complex_ring (poly_mul complex_ring p q) /\ ~(poly_mul complex_ring p q = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring p q) z = 0 ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN specialize_assuming[`p:(1->num)->complex`;`q:(1->num)->complex`;`z:complex`]poly_ord_mul THEN qed[ARITH_RULE `0+0 = 0`] );; let poly_ord_ge_if = prove(` !p:(1->num)->complex f n. ring_polynomial complex_ring f ==> p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) n) f ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) `, intro THEN case `f = poly_0 complex_ring:(1->num)->complex` THENL [ qed[POLY_MUL_0;poly_pow_poly;x_minus_const_poly;in_complex_ring] ; pass ] THEN choose `g:(1->num)->complex` `ring_polynomial complex_ring g /\ ~complex_root g z /\ f = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord f z)) g` [poly_ord_exists] THEN have `p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) n) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord f z)) g)` [] THEN have `p = poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) n) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord f z))) g` [POLY_MUL_ASSOC;series_complex] THEN have `p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (n + poly_ord f z)) g` [poly_pow_add;series_complex] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [x_minus_const_poly;poly_pow_poly;RING_POLYNOMIAL_MUL;in_complex_ring] THEN have `n + poly_ord f z = poly_ord p z` [poly_ord_unique] THEN ASM_ARITH_TAC );; let poly_ord_ge = prove(` !p:(1->num)->complex n. ring_polynomial complex_ring p ==> ((p = poly_0 complex_ring \/ n <= poly_ord p z) <=> ?f:(1->num)->complex. ( ring_polynomial complex_ring f /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) n) f ) ) `, intro THEN splitiff THENL [ case `p = poly_0 complex_ring:(1->num)->complex` THENL [ intro THEN witness `poly_0 complex_ring:(1->num)->complex` THEN have `ring_polynomial complex_ring (x_minus_const complex_ring z)` [x_minus_const_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) n)` [poly_pow_poly] THEN qed[POLY_MUL_0;RING_POLYNOMIAL_0] ; pass ] THEN intro THENL [ qed[] ; pass ] THEN choose `f:(1->num)->complex` `ring_polynomial complex_ring f /\ ~complex_root f z /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) f` [poly_ord_exists] THEN num_linear_fact `n <= poly_ord p z ==> poly_ord p z = n + (poly_ord p z - n)` THEN have `poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z) = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (n)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z - n))` [poly_pow_add;series_complex] THEN have `p = poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (n)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z - n))) f` [] THEN have `p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (n)) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z - n)) f)` [POLY_MUL_ASSOC;series_complex] THEN witness `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z - n)) f` THEN qed[RING_POLYNOMIAL_MUL;poly_pow_poly;x_minus_const_poly;in_complex_ring] ; qed[poly_ord_ge_if] ] );; let poly_ord_mul_ge_ge = prove(` !p:(1->num)->complex q z m n. ring_polynomial complex_ring p ==> ( p = poly_0 complex_ring \/ m <= poly_ord p z ) ==> ring_polynomial complex_ring q ==> ( q = poly_0 complex_ring \/ n <= poly_ord q z ) ==> ( ring_polynomial complex_ring (poly_mul complex_ring p q) /\ ( poly_mul complex_ring p q = poly_0 complex_ring \/ m + n <= poly_ord (poly_mul complex_ring p q) z ) ) `, REPEAT GEN_TAC THEN simp[poly_ord_ge;RING_POLYNOMIAL_MUL] THEN intro THEN witness `poly_mul complex_ring f f':(1->num)->complex` THEN have `poly_pow complex_ring (x_minus_const complex_ring z) (m + n) = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) m) (poly_pow complex_ring (x_minus_const complex_ring z) n)` [poly_pow_add;series_complex] THEN simp[RING_POLYNOMIAL_MUL] THEN rw[x_series_use] THEN specialize_assuming[ `x_series complex_ring`; `poly_pow complex_ring (x_minus_const complex_ring z) m`; `poly_pow complex_ring (x_minus_const complex_ring z) n`; ](GENL[ `r:R ring`; `M:R`; `N:R`; ](RING_RULE `ring_mul(r:R ring) (ring_mul r M f) (ring_mul r N f') = ring_mul r (ring_mul r M N) (ring_mul r f f')` )) THEN qed[x_series_use;series_complex] );; let poly_ord_mul_ge_ge0 = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ( p = poly_0 complex_ring \/ n <= poly_ord p z ) ==> ring_polynomial complex_ring q ==> ( ring_polynomial complex_ring (poly_mul complex_ring p q) /\ ( poly_mul complex_ring p q = poly_0 complex_ring \/ n <= poly_ord (poly_mul complex_ring p q) z ) ) `, REPEAT GEN_TAC THEN simp[poly_ord_ge;RING_POLYNOMIAL_MUL] THEN intro THEN witness `poly_mul complex_ring f q:(1->num)->complex` THEN simp[RING_POLYNOMIAL_MUL] THEN qed[POLY_MUL_ASSOC;series_complex] );; let poly_ord_mul_ge0_ge = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> ( q = poly_0 complex_ring \/ n <= poly_ord q z ) ==> ( ring_polynomial complex_ring (poly_mul complex_ring p q) /\ ( poly_mul complex_ring p q = poly_0 complex_ring \/ n <= poly_ord (poly_mul complex_ring p q) z ) ) `, qed[poly_ord_mul_ge_ge0;POLY_MUL_SYM;series_complex] );; let poly_ord_neg_ge = prove(` !p:(1->num)->complex z n. ring_polynomial complex_ring p ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) ==> (poly_neg complex_ring p = poly_0 complex_ring \/ n <= poly_ord(poly_neg complex_ring p) z) `, REPEAT GEN_TAC THEN simp[poly_ord_ge;RING_POLYNOMIAL_NEG] THEN intro THEN witness `poly_neg complex_ring f:(1->num)->complex` THEN simp[RING_POLYNOMIAL_NEG] THEN rw[x_series_use] THEN qed[RING_MUL_RNEG;series_complex;x_series_use] );; let poly_ord_add_ge = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) ==> (q = poly_0 complex_ring \/ n <= poly_ord q z) ==> (poly_add complex_ring p q = poly_0 complex_ring \/ n <= poly_ord(poly_add complex_ring p q) z) `, REPEAT GEN_TAC THEN simp[poly_ord_ge;RING_POLYNOMIAL_ADD] THEN intro THEN witness `poly_add complex_ring f f':(1->num)->complex` THEN simp[RING_POLYNOMIAL_ADD] THEN simp[POLY_ADD_LDISTRIB;series_complex] );; let poly_ord_sum_ge = prove(` !p z n S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial complex_ring (p s)) ==> (!s:X. s IN S ==> (p s = poly_0 complex_ring \/ n <= poly_ord (p s) z)) ==> (poly_sum complex_ring S p = poly_0 complex_ring \/ n <= poly_ord(poly_sum complex_ring S p) z) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_sum_empty] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `ring_powerseries complex_ring (p(x:X):(1->num)->complex)` [x_poly_use;ring_polynomial] THEN simp[poly_sum_insert] THEN have `ring_polynomial complex_ring (poly_sum complex_ring (S:X->bool) p)` [poly_sum_poly] THEN qed[poly_ord_add_ge] ] );; let poly_ord_sub_ge = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) ==> (q = poly_0 complex_ring \/ n <= poly_ord q z) ==> (poly_sub complex_ring p q = poly_0 complex_ring \/ n <= poly_ord(poly_sub complex_ring p q) z) `, rw[POLY_SUB] THEN qed[poly_ord_add_ge;poly_ord_neg_ge;RING_POLYNOMIAL_NEG] );; let poly_ord_add_dominant = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) ==> ~(q = poly_0 complex_ring) ==> poly_ord q z < n ==> ( ~(poly_add complex_ring p q = poly_0 complex_ring) /\ poly_ord(poly_add complex_ring p q) z = poly_ord q z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN case `poly_add complex_ring p q = poly_0 complex_ring:(1->num)->complex` THENL [ have `ring_add(x_series complex_ring) p q = ring_0(x_series complex_ring)` [x_series_use] THEN have `q = ring_neg(x_series complex_ring) p` [RING_ADD_EQ_0;series_complex;x_series_use] THEN specialize[`p:(1->num)->complex`;`z:complex`;`n:num`]poly_ord_neg_ge THEN have `q = poly_0 complex_ring \/ n <= poly_ord q z` [x_series_use] THEN qed[ARITH_RULE `n <= d:num ==> ~(d < n)`] ; pass ] THEN case `poly_ord q z < poly_ord(poly_add complex_ring p q) z` THENL [ num_linear_fact `poly_ord q z < poly_ord(poly_add complex_ring p q) z ==> 1 + poly_ord q z <= poly_ord(poly_add complex_ring p q) z` THEN have `ring_polynomial complex_ring (poly_add complex_ring p q:(1->num)->complex)` [RING_POLYNOMIAL_ADD] THEN have `poly_add complex_ring p q = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord (poly_add complex_ring p q) z` [] THEN have `p = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord p z` [ARITH_RULE `q < n ==> n <= p ==> 1+q <= p`] THEN specialize[`poly_add complex_ring p q:(1->num)->complex`;`p:(1->num)->complex`;`z:complex`;`1 + poly_ord q z`]poly_ord_sub_ge THEN have `poly_sub complex_ring (poly_add complex_ring p q) p = q:(1->num)->complex` [x_series_use;x_series_sub_use;RING_RULE `ring_sub(r:R ring) (ring_add r p q) p = q`;series_complex] THEN have `q = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord q z` [] THEN num_linear_fact `~(1 + poly_ord q z <= poly_ord q z)` THEN qed[] ; pass ] THEN num_linear_fact `~(poly_ord q z < poly_ord (poly_add complex_ring p q) z) ==> poly_ord(poly_add complex_ring p q) z <= poly_ord q z` THEN specialize_assuming[`poly_ord q z`;`n:num`;`poly_ord p z`]LTE_TRANS THEN have `p = poly_0 complex_ring \/ poly_ord q z <= poly_ord p z` [LT_IMP_LE] THEN have `q = poly_0 complex_ring \/ poly_ord q z <= poly_ord q z` [LE_REFL] THEN specialize[`p:(1->num)->complex`;`q:(1->num)->complex`;`z:complex`;`poly_ord q z`]poly_ord_add_ge THEN have `poly_ord q z <= poly_ord(poly_add complex_ring p q) z` [] THEN num_linear_fact `poly_ord (poly_add complex_ring p q) z <= poly_ord q z ==> poly_ord q z <= poly_ord(poly_add complex_ring p q) z ==> poly_ord (poly_add complex_ring p q) z = poly_ord q z` THEN qed[] );; let poly_ord_sub_dominant = prove(` !p:(1->num)->complex q z n. ring_polynomial complex_ring p ==> ring_polynomial complex_ring q ==> (p = poly_0 complex_ring \/ n <= poly_ord p z) ==> ~(q = poly_0 complex_ring) ==> poly_ord q z < n ==> ( ~(poly_sub complex_ring p q = poly_0 complex_ring) /\ poly_ord(poly_sub complex_ring p q) z = poly_ord q z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN case `poly_sub complex_ring p q = poly_0 complex_ring:(1->num)->complex` THENL [ have `ring_sub(x_series complex_ring) p q = ring_0(x_series complex_ring)` [x_series_sub_use;x_series_use] THEN have `q = p:(1->num)->complex` [RING_SUB_EQ_0;series_complex;x_series_use] THEN qed[ARITH_RULE `n <= d:num ==> ~(d < n)`] ; pass ] THEN case `poly_ord q z < poly_ord(poly_sub complex_ring p q) z` THENL [ num_linear_fact `poly_ord q z < poly_ord(poly_sub complex_ring p q) z ==> 1 + poly_ord q z <= poly_ord(poly_sub complex_ring p q) z` THEN have `ring_polynomial complex_ring (poly_sub complex_ring p q:(1->num)->complex)` [RING_POLYNOMIAL_SUB] THEN have `poly_sub complex_ring p q = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord (poly_sub complex_ring p q) z` [] THEN have `p = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord p z` [ARITH_RULE `q < n ==> n <= p ==> 1+q <= p`] THEN specialize[`p:(1->num)->complex`;`poly_sub complex_ring p q:(1->num)->complex`;`z:complex`;`1 + poly_ord q z`]poly_ord_sub_ge THEN have `poly_sub complex_ring p (poly_sub complex_ring p q) = q:(1->num)->complex` [x_series_use;x_series_sub_use;RING_RULE `ring_sub(r:R ring) p (ring_sub r p q) = q`;series_complex] THEN have `q = poly_0 complex_ring \/ 1 + poly_ord q z <= poly_ord q z` [] THEN num_linear_fact `~(1 + poly_ord q z <= poly_ord q z)` THEN qed[] ; pass ] THEN num_linear_fact `~(poly_ord q z < poly_ord (poly_sub complex_ring p q) z) ==> poly_ord(poly_sub complex_ring p q) z <= poly_ord q z` THEN specialize_assuming[`poly_ord q z`;`n:num`;`poly_ord p z`]LTE_TRANS THEN have `p = poly_0 complex_ring \/ poly_ord q z <= poly_ord p z` [LT_IMP_LE] THEN have `q = poly_0 complex_ring \/ poly_ord q z <= poly_ord q z` [LE_REFL] THEN specialize[`p:(1->num)->complex`;`q:(1->num)->complex`;`z:complex`;`poly_ord q z`]poly_ord_sub_ge THEN have `poly_ord q z <= poly_ord(poly_sub complex_ring p q) z` [] THEN num_linear_fact `poly_ord (poly_sub complex_ring p q) z <= poly_ord q z ==> poly_ord q z <= poly_ord(poly_sub complex_ring p q) z ==> poly_ord (poly_sub complex_ring p q) z = poly_ord q z` THEN qed[] );; let poly_ord_sum_dominant = prove(` !p z n S t. FINITE S ==> (!s:X. s IN S ==> ring_polynomial complex_ring (p s)) ==> (!s:X. s IN S ==> ~(s = t) ==> (p s = poly_0 complex_ring) \/ n <= poly_ord (p s) z) ==> t IN S ==> ~(p t = poly_0 complex_ring) ==> poly_ord (p t) z < n ==> ( ring_polynomial complex_ring (poly_sum complex_ring S p) /\ ~(poly_sum complex_ring S p = poly_0 complex_ring) /\ poly_ord(poly_sum complex_ring S p) z = poly_ord (p t) z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN have `poly_sum complex_ring S p = poly_add complex_ring (p(t:X)) (poly_sum complex_ring (S DELETE t) p)` [poly_sum_delete2;series_complex] THEN have `!s:X. s IN S DELETE t ==> s IN S` [IN_DELETE] THEN have `!s:X. s IN S DELETE t ==> ~(s = t)` [IN_DELETE] THEN subgoal `poly_sum complex_ring (S DELETE (t:X)) p = poly_0 complex_ring \/ n <= poly_ord (poly_sum complex_ring (S DELETE t) p) z` THENL [ specialize_assuming[`p:X->((1->num)->complex)`;`z:complex`;`n:num`;`S DELETE(t:X)`]poly_ord_sum_ge THEN qed[FINITE_DELETE] ; pass ] THEN have `ring_polynomial complex_ring (p(t:X):(1->num)->complex)` [] THEN have `ring_polynomial complex_ring (poly_sum complex_ring (S DELETE(t:X)) p)` [poly_sum_poly;FINITE_DELETE] THEN specialize[ `poly_sum complex_ring (S DELETE(t:X)) p`; `p(t:X):(1->num)->complex`; `z:complex`;`n:num`]poly_ord_add_dominant THEN qed[POLY_ADD_SYM;series_complex;poly_sum_poly] );; (* order of derivative of p/q when p/q has no pole at z *) let poly_ord_derivative_ge = prove(` !p:(1->num)->complex q PqpQ z. ring_polynomial complex_ring q ==> ring_polynomial complex_ring p ==> ~(q = poly_0 complex_ring) ==> (p = poly_0 complex_ring \/ poly_ord q z <= poly_ord p z) ==> PqpQ = poly_sub complex_ring ( poly_mul complex_ring ( x_derivative complex_ring p ) ( q ) ) ( poly_mul complex_ring ( p ) ( x_derivative complex_ring q ) ) ==> ( PqpQ = poly_0 complex_ring \/ 2 * poly_ord q z <= poly_ord PqpQ z ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN choose `f:(1->num)->complex` `ring_polynomial complex_ring f /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) f` [poly_ord_ge] THEN choose `g:(1->num)->complex` `ring_polynomial complex_ring g /\ q = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) g` [poly_ord_exists] THEN specialize_assuming[`complex_ring`;`f:(1->num)->complex`;`g:(1->num)->complex`;`poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)`]x_derivative_ratio_scaling THEN have `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) f)) q) (poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) f) (x_derivative complex_ring q)) = poly_mul complex_ring (poly_pow complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))` [series_complex] THEN simp[poly_pow_mul;series_complex] THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))` [RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB;x_derivative_polynomial;x_minus_const_poly;in_complex_ring] THEN specialize[`poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z * 2)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))`;`poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))`;`poly_ord q z * 2`]poly_ord_ge_if THEN qed[MULT_SYM] );; let poly_ord_derivative_pole_lemma = prove(` !z:complex m n. poly_sub complex_ring ( poly_mul complex_ring ( x_derivative complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) m ) ) ( poly_pow complex_ring ( x_minus_const complex_ring z ) n ) ) ( poly_mul complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) m ) ( x_derivative complex_ring ( poly_pow complex_ring ( x_minus_const complex_ring z ) n ) ) ) = poly_mul complex_ring ( poly_sub complex_ring ( poly_const complex_ring (ring_of_num complex_ring m) ) ( poly_const complex_ring (ring_of_num complex_ring n) ) ) ( poly_pow complex_ring ( x_minus_const complex_ring z ) ((m+n)-1) ) `, intro THEN simp[x_derivative_pow;series_complex] THEN recall in_complex_ring THEN simp[x_derivative_x_minus_const] THEN recall series_complex THEN case `m = 0` THENL [ simp[poly_pow_0;ARITH_RULE `0 - 1 = 0`;POLY_MUL_LID;RING_OF_NUM_0;GSYM poly_0;POWSER_MUL_0;ARITH_RULE `0+n = n`] THEN rw[x_series_use;x_series_sub_use] THEN qed[RING_RULE `ring_sub(r:R ring) (ring_0 r) (ring_mul r C P) = ring_mul r (ring_sub r (ring_0 r) C) P`;x_series_use] ; pass ] THEN case `n = 0` THENL [ simp[poly_pow_0;ARITH_RULE `0 - 1 = 0`;POLY_MUL_LID;RING_OF_NUM_0;GSYM poly_0;POWSER_MUL_0;ARITH_RULE `m+0 = m`] THEN rw[x_series_use;x_series_sub_use;RING_MUL_RID] THEN qed[RING_RULE `ring_sub(r:R ring) (ring_mul r (ring_mul r C P) (ring_1 r)) (ring_0 r) = ring_mul r (ring_sub r C (ring_0 r)) P`;x_series_use] ; pass ] THEN num_linear_fact `~(m = 0) ==> (m-1)+n = (m+n)-1` THEN have `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (m-1)) (poly_pow complex_ring (x_minus_const complex_ring z) (n)) = poly_pow complex_ring (x_minus_const complex_ring z) ((m+n)-1)` [poly_pow_add;series_complex] THEN num_linear_fact `~(n = 0) ==> m+(n-1) = (m+n)-1` THEN have `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (m)) (poly_pow complex_ring (x_minus_const complex_ring z) (n-1)) = poly_pow complex_ring (x_minus_const complex_ring z) ((m+n)-1)` [poly_pow_add;series_complex] THEN rw[x_series_use;x_series_sub_use] THEN specialize_assuming[ `x_series complex_ring` ;`poly_const complex_ring (ring_of_num complex_ring m):(1->num)->complex` ;`poly_const complex_ring (ring_of_num complex_ring n):(1->num)->complex` ;`poly_pow complex_ring (x_minus_const complex_ring z) m` ;`poly_pow complex_ring (x_minus_const complex_ring z) (m-1)` ;`poly_pow complex_ring (x_minus_const complex_ring z) n` ;`poly_pow complex_ring (x_minus_const complex_ring z) (n-1)` ;`poly_pow complex_ring (x_minus_const complex_ring z) ((m+n)-1)` ](GENL[ `r:R ring`; `M:R`; `N:R`; `PM:R`; `PM1:R`; `PN:R`; `PN1:R`; `PMN1:R` ](RING_RULE `ring_mul (r:R ring) PM PN1 = PMN1 ==> ring_mul r PM1 PN = PMN1 ==> ring_sub r (ring_mul r (ring_mul r M (ring_mul r (ring_1 r) PM1)) PN) (ring_mul r PM (ring_mul r N (ring_mul r (ring_1 r) PN1))) = ring_mul r (ring_sub r M N) PMN1`)) THEN have `poly_const complex_ring (ring_of_num complex_ring m) IN ring_carrier (x_series complex_ring)` [POWSER_CONST;x_series_use] THEN have `poly_const complex_ring (ring_of_num complex_ring n) IN ring_carrier (x_series complex_ring)` [POWSER_CONST;x_series_use] THEN qed[x_series_use;x_series_sub_use] );; (* order of derivative of p/q when p/q has a pole at z *) let poly_ord_derivative_pole = prove(` !p:(1->num)->complex q PqpQ z. ring_polynomial complex_ring q ==> ring_polynomial complex_ring p ==> ~(q = poly_0 complex_ring) ==> ~(p = poly_0 complex_ring) ==> poly_ord p z < poly_ord q z ==> PqpQ = poly_sub complex_ring ( poly_mul complex_ring ( x_derivative complex_ring p ) ( q ) ) ( poly_mul complex_ring ( p ) ( x_derivative complex_ring q ) ) ==> ( ring_polynomial complex_ring PqpQ /\ ~(PqpQ = poly_0 complex_ring) /\ poly_ord PqpQ z = (poly_ord p z + poly_ord q z) - 1 ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN choose `f:(1->num)->complex` `ring_polynomial complex_ring f /\ ~complex_root f z /\ p = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) f` [poly_ord_exists] THEN choose `g:(1->num)->complex` `ring_polynomial complex_ring g /\ ~complex_root g z /\ q = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) g` [poly_ord_exists] THEN subgoal `PqpQ = ring_sub (x_poly complex_ring) (ring_mul (x_poly complex_ring) (x_derivative complex_ring p) q) (ring_mul (x_poly complex_ring) p (x_derivative complex_ring q))` THENL [ qed[x_poly_use;x_poly_sub_use] ; pass ] THEN subgoal `PqpQ = poly_add complex_ring (poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) (poly_mul complex_ring (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))))) (poly_mul complex_ring f g))` THENL [ have `x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `x_derivative complex_ring f IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial] THEN have `x_derivative complex_ring g IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial] THEN have `x_derivative complex_ring p IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial] THEN have `x_derivative complex_ring q IN ring_carrier (x_poly complex_ring)` [x_poly_use;x_derivative_polynomial] THEN have `poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z) IN ring_carrier (x_poly complex_ring)` [x_poly_use;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z) IN ring_carrier (x_poly complex_ring)` [x_poly_use;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `f IN ring_carrier (x_poly complex_ring)` [x_poly_use] THEN have `g IN ring_carrier (x_poly complex_ring)` [x_poly_use] THEN have `p IN ring_carrier (x_poly complex_ring)` [x_poly_use] THEN have `q IN ring_carrier (x_poly complex_ring)` [x_poly_use] THEN have `p = ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) f` [x_poly_use] THEN have `q = ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) g` [x_poly_use] THEN subgoal `x_derivative complex_ring p = ring_add (x_poly complex_ring) (ring_mul (x_poly complex_ring) (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))) f) (ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (x_derivative complex_ring f))` THENL [ rw[GSYM x_poly_use] THEN rw[GSYM x_poly_sub_use] THEN qed[x_derivative_mul;series_complex] ; pass ] THEN subgoal `x_derivative complex_ring q = ring_add (x_poly complex_ring) (ring_mul (x_poly complex_ring) (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) g) (ring_mul (x_poly complex_ring) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) (x_derivative complex_ring g))` THENL [ rw[GSYM x_poly_use] THEN rw[GSYM x_poly_sub_use] THEN qed[x_derivative_mul;series_complex] ; pass ] THEN specialize[ `x_poly complex_ring` ;`x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))` ;`x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))` ;`x_derivative complex_ring f` ;`x_derivative complex_ring g` ;`x_derivative complex_ring p` ;`x_derivative complex_ring q` ;`(poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))` ;`(poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))` ;`f:(1->num)->complex` ;`g:(1->num)->complex` ;`p:(1->num)->complex` ;`q:(1->num)->complex` ](GENL [ `r:R ring`;`D:R`;`E:R`;`FF:R`;`G:R`;`P:R`;`Q:R`;`d:R`;`e:R`;`f:R`;`g:R`;`p:R`;`q:R` ](RING_RULE ` p = ring_mul(r:R ring) d f ==> q = ring_mul r e g ==> P = ring_add r (ring_mul r D f) (ring_mul r d FF) ==> Q = ring_add r (ring_mul r E g) (ring_mul r e G) ==> ring_sub r (ring_mul r P q) (ring_mul r p Q) = ring_add r (ring_mul r (ring_mul r d e) (ring_sub r (ring_mul r FF g) (ring_mul r f G))) (ring_mul r (ring_sub r (ring_mul r D e) (ring_mul r d E)) (ring_mul r f g)) `)) THEN rw[x_poly_use;x_poly_sub_use] THEN qed[] ; pass ] THEN have `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)) = poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)` [poly_pow_add;series_complex] THEN have `PqpQ = poly_add complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) (poly_mul complex_ring (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))))) (poly_mul complex_ring f g))` [] THEN have `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z))) (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z))) (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z)) (x_derivative complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord q z)))) = poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))` [poly_ord_derivative_pole_lemma] THEN have `PqpQ = poly_add complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) (poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g))` [] THEN subgoal `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) = poly_0 complex_ring \/ poly_ord p z + poly_ord q z <= poly_ord(poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) z` THENL [ specialize_assuming[`poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))`;`poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))`;`poly_ord p z + poly_ord q z`]poly_ord_ge_if THEN qed[x_derivative_polynomial;RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB] ; pass ] THEN subgoal `~(poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g) = poly_0 complex_ring) /\ poly_ord(poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g)) z = (poly_ord p z + poly_ord q z) - 1` THENL [ have `ring_polynomial complex_ring (poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g))` [RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST;poly_pow_poly;x_minus_const_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_mul complex_ring f g:(1->num)->complex))` [RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST;in_complex_ring] THEN subgoal `~complex_root (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_mul complex_ring f g)) z` THENL [ rw[complex_root] THEN simp[POLY_EVAL_MUL;POLY_EVAL_SUB;POLY_EVAL_CONST;in_complex_ring;RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST] THEN recall integral_domain_complex THEN have `~(poly_eval complex_ring f z = Cx(&0))` [complex_root] THEN have `~(poly_eval complex_ring g z = Cx(&0))` [complex_root] THEN have `~(ring_mul complex_ring (poly_eval complex_ring f z) (poly_eval complex_ring g z) = Cx(&0))` [complex_ring_clauses;COMPLEX_ENTIRE] THEN num_linear_fact `poly_ord p z < poly_ord q z ==> ~(poly_ord p z = poly_ord q z)` THEN have `~(ring_of_num complex_ring (poly_ord p z) = ring_of_num complex_ring (poly_ord q z))` [ring_of_num_injective;ring_char_complex] THEN complex_field_fact `~(ring_of_num complex_ring (poly_ord p z) = ring_of_num complex_ring (poly_ord q z)) ==> ~(ring_of_num complex_ring (poly_ord p z) - ring_of_num complex_ring (poly_ord q z) = Cx(&0))` THEN have `~(ring_sub complex_ring (ring_of_num complex_ring (poly_ord p z)) (ring_of_num complex_ring (poly_ord q z)) = Cx(&0))` [ring_sub_complex] THEN qed[complex_ring_clauses;COMPLEX_ENTIRE] ; pass ] THEN subgoal `poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g) = poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1)) (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_mul complex_ring f g))` THENL [ rw[x_series_use;x_series_sub_use] THEN qed[x_series_use;series_complex;RING_RULE `ring_mul(r:R ring) (ring_mul r (ring_sub r P Q) MN1) (ring_mul r f g) = ring_mul r MN1 (ring_mul r (ring_sub r P Q) (ring_mul r f g))`] ; pass ] THEN specialize[`poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g)`;`poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_mul complex_ring f g:(1->num)->complex)`;`(poly_ord p z + poly_ord q z) - 1`]poly_ord_unique THEN qed[] ; pass ] THEN num_linear_fact `poly_ord p z < poly_ord q z ==> (poly_ord p z + poly_ord q z) - 1 < poly_ord p z + poly_ord q z` THEN have `poly_ord (poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g)) z < poly_ord p z + poly_ord q z` [] THEN recall in_complex_ring THEN have `ring_polynomial complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z))` [poly_pow_poly;x_minus_const_poly] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g)` [RING_POLYNOMIAL_MUL;x_derivative_polynomial] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))))` [RING_POLYNOMIAL_MUL;poly_pow_poly;x_minus_const_poly;RING_POLYNOMIAL_SUB;x_derivative_polynomial] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))` [poly_pow_poly;x_minus_const_poly] THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)):(1->num)->complex))` [RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g))` [RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_CONST;x_minus_const_poly;poly_pow_poly] THEN have `~(poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g) = poly_0 complex_ring)` [] THEN specialize[ `poly_mul complex_ring (poly_pow complex_ring (x_minus_const complex_ring z) (poly_ord p z + poly_ord q z)) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))`; `poly_mul complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_const complex_ring (ring_of_num complex_ring (poly_ord p z))) (poly_const complex_ring (ring_of_num complex_ring (poly_ord q z)))) (poly_pow complex_ring (x_minus_const complex_ring z) ((poly_ord p z + poly_ord q z) - 1))) (poly_mul complex_ring f g)`; `z:complex`; `poly_ord p z + poly_ord q z` ]poly_ord_add_dominant THEN subgoal `ring_polynomial complex_ring (PqpQ:(1->num)->complex)` THENL [ simp[] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (x_derivative complex_ring p) q)` [RING_POLYNOMIAL_MUL;x_derivative_polynomial] THEN have `ring_polynomial complex_ring (poly_mul complex_ring p (x_derivative complex_ring q))` [RING_POLYNOMIAL_MUL;x_derivative_polynomial] THEN qed[RING_POLYNOMIAL_SUB] ; pass ] THEN qed[] );; let ord_minpoly = prove(` !P q y. FINITE P ==> distinct_minpolys P ==> q IN P ==> poly_ord q y = (if complex_root q y then 1 else 0) `, intro THEN subgoal `poly_product complex_ring (complex_root q) (\z. x_minus_const complex_ring (I z)) = q` THENL [ rw[GSYM monic_vanishing_at] THEN qed[distinct_minpolys_monic_vanishing_at] ; pass ] THEN have `FINITE(complex_root q)` [distinct_minpolys_finite_root] THEN subgoal `poly_ord (poly_product complex_ring (complex_root q) (\z. x_minus_const complex_ring (I z))) y = nsum (complex_root q) (\s. poly_ord (x_minus_const complex_ring (I s)) y)` THENL [ have `!z. z IN complex_root q ==> ring_polynomial complex_ring (x_minus_const complex_ring (I z))` [x_minus_const_poly;in_complex_ring] THEN have `!z. z IN complex_root q ==> ~(x_minus_const complex_ring (I z) = poly_0 complex_ring)` [x_minus_const_nonzero;ring_1_0_complex;in_complex_ring] THEN specialize[ `\z. x_minus_const complex_ring (I z)`; `y:complex`; `complex_root q` ]poly_ord_product THEN qed[] ; pass ] THEN subgoal `nsum (complex_root q) (\s. poly_ord (x_minus_const complex_ring (I s)) y) = nsum (complex_root q) (\z. if z = y then 1 else 0)` THENL [ sufficesby NSUM_EQ THEN rw[BETA_THM;I_THM] THEN qed[poly_ord_x_minus_const;x_minus_const_poly;in_complex_ring;x_minus_const_nonzero;ring_1_0_complex] ; pass ] THEN have `poly_ord q y = nsum (complex_root q) (\s. poly_ord (x_minus_const complex_ring (I s)) y)` [] THEN have `poly_ord q y = nsum (complex_root q) (\s. if s = y then 1 else 0)` [] THEN specialize[ `complex_root q`; `y:complex`; `1` ]nsum_delta THEN qed[IN] );; let ord_pow_minpoly = prove(` !P q y n. FINITE P ==> distinct_minpolys P ==> q IN P ==> poly_ord (poly_pow QinC_ring q n) y = (if complex_root q y then n else 0) `, intro THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [distinct_minpolys;poly_complex_if_poly_QinC] THEN have `~((q:(1->num)->complex) = poly_0 complex_ring)` [distinct_minpolys_nonzero;poly_0_QinC_eq_poly_0_complex] THEN specialize[ `q:(1->num)->complex`; `n:num`; `y:complex` ]poly_ord_pow THEN specialize[ `P:((1->num)->complex)->bool`; `q:(1->num)->complex`; `y:complex` ]ord_minpoly THEN have `n * poly_ord q y = (if complex_root q y then n else 0)` [ARITH_RULE `n * 1 = n`;ARITH_RULE `n * 0 = 0`] THEN have `poly_pow complex_ring q n = poly_pow QinC_ring q n:(1->num)->complex` [poly_pow_subring;subring_complex_QinC;ring_polynomial;distinct_minpolys] THEN qed[] );; let ord_product_distinct_minpolys = prove(` !P e z. FINITE P ==> distinct_minpolys P ==> poly_ord (poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p))) z = nsum P (\p. if complex_root p z then e p else 0) `, intro THEN subgoal `poly_ord (poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p))) z = nsum P (\p. poly_ord (poly_pow QinC_ring p (e p)) z)` THENL [ subgoal `!p:(1->num)->complex. p IN P ==> ring_polynomial complex_ring (poly_pow QinC_ring p (e p))` THENL [ intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_polynomial QinC_ring (poly_pow QinC_ring (p:(1->num)->complex) (e p))` [poly_pow_poly] THEN qed[poly_complex_if_poly_QinC] ; pass ] THEN subgoal `!p:(1->num)->complex. p IN P ==> ~(poly_pow QinC_ring p (e p) = poly_0 complex_ring)` THENL [ intro THEN have `~(p:(1->num)->complex = poly_0 QinC_ring)` [distinct_minpolys_nonzero] THEN have `~(poly_pow QinC_ring (p:(1->num)->complex) (e p) = poly_0 QinC_ring)` [nonzero_poly_pow;integral_domain_QinC;ring_polynomial;distinct_minpolys] THEN qed[poly_0_QinC_eq_poly_0_complex] ; pass ] THEN specialize[ `\p:(1->num)->complex. poly_pow QinC_ring p (e p)`; `z:complex`; `P:((1->num)->complex)->bool` ]poly_ord_product THEN specialize_assuming[ `complex_ring`; `QinC`; `\p:(1->num)->complex. poly_pow QinC_ring p (e p)`; `P:((1->num)->complex)->bool` ]poly_product_subring THEN qed[ring_polynomial;distinct_minpolys;subring_complex_QinC;poly_pow_poly] ; pass ] THEN simp[] THEN sufficesby NSUM_EQ THEN rw[BETA_THM] THEN intro THEN specialize[ `P:((1->num)->complex)->bool`; `x:(1->num)->complex`; `z:complex`; `e(x:(1->num)->complex):num` ]ord_pow_minpoly THEN qed[] );; let ord_product_distinct_minpolys_root = prove(` !P q e z. FINITE P ==> distinct_minpolys P ==> q IN P ==> complex_root q z ==> poly_ord (poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p))) z = e q `, intro THEN simp[ord_product_distinct_minpolys] THEN subgoal `nsum P (\p. if complex_root p z then e p else 0) = nsum P (\p. if p = q then e q else 0)` THENL [ sufficesby NSUM_EQ THEN rw[BETA_THM] THEN intro THEN qed[distinct_minpolys_distinct_roots] ; pass ] THEN subgoal `nsum P (\p:(1->num)->complex. if p = q then e q else 0) = e q` THENL [ specialize[ `P:((1->num)->complex)->bool`; `q:(1->num)->complex`; `e(q:(1->num)->complex):num` ]nsum_delta THEN qed[] ; pass ] THEN qed[] );; let sum_root_decomposition_if_monic_vanishing_at_factorization_lemma = prove(` !P:((1->num)->complex)->bool S:X->bool e:((1->num)->complex)->num c:X->complex g:complex->complex. FINITE P ==> distinct_minpolys P ==> FINITE S ==> (!p. p IN P ==> ~(e p = 0)) ==> poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = monic_vanishing_at complex_ring S c ==> ring_sum complex_ring P (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g) = ring_sum complex_ring S (g o c) `, intro THEN subgoal `!z:complex. nsum P (\p. if complex_root p z then e p else 0) = numpreimages c (S:X->bool) z` THENL [ intro THEN specialize[ `P:((1->num)->complex)->bool`; `e:((1->num)->complex)->num`; `z:complex` ]ord_product_distinct_minpolys THEN specialize[ `S:X->bool`; `c:X->complex`; `z:complex` ]poly_ord_monic_vanishing_at_numpreimages THEN qed[] ; pass ] THEN have `!y:complex. y IN IMAGE (c:X->complex) S ==> g y IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `S:X->bool`; `c:X->complex`; `g:complex->complex` ]ring_sum_o_v2 THEN simp[] THEN subgoal `ring_sum complex_ring P (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g) = ring_sum complex_ring P (\p. ring_sum complex_ring {y | ~(numpreimages c (S:X->bool) y = 0)} (\y. if y IN complex_root p then Cx(&(e p)) * g y else Cx(&0)))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `FINITE(complex_root a)` [distinct_minpolys_finite_root] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[GSYM VSUM_COMPLEX_LMUL] THEN simp[vsum_ring_sum_complex] THEN subgoal `complex_root a SUBSET {y | ~(numpreimages c (S:X->bool) y = 0)}` THENL [ rw[SUBSET;IN_ELIM_THM] THEN intro THEN have `(if complex_root a x then e a else 0) <= nsum P (\p. if complex_root p x then e p else 0)` [term_le_nsum] THEN have `e a <= nsum P (\p. if complex_root p x then e p else 0)` [IN] THEN have `e(a:(1->num)->complex) <= numpreimages c (S:X->bool) (x:complex)` [] THEN have `~(e(a:(1->num)->complex) = 0)` [] THEN ASM_ARITH_TAC ; pass ] THEN specialize[ `complex_ring`; `complex_root a`; `{y:complex | ~(numpreimages c (S:X->bool) y = 0)}`; `\y:complex. ring_mul complex_ring (Cx(&(e(a:(1->num)->complex)))) (g y)` ]ring_sum_restrict_subset THEN rw[GSYM complex_ring_clauses] THEN qed[] ; pass ] THEN simp[] THEN simp[GSYM image_numpreimages] THEN have `FINITE(IMAGE (c:X->complex) S)` [FINITE_IMAGE] THEN have `!i j. i IN P ==> j IN IMAGE c (S:X->bool) ==> (if j IN complex_root i then Cx (&(e i)) * g j else Cx (&0)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `\p y. if y IN complex_root p then Cx(&(e p)) * g y else Cx(&0)`; `P:((1->num)->complex)->bool`; `IMAGE (c:X->complex) S` ]RING_SUM_SWAP THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[in_image_vw] THEN intro THEN case `!q:(1->num)->complex. q IN P ==> ~complex_root q a` THENL [ subgoal `nsum P (\p. if complex_root p a then e p else 0) = nsum P (\p. 0)` THENL [ sufficesby NSUM_EQ THEN qed[] ; pass ] THEN have `nsum P (\p:(1->num)->complex. 0) = 0` [NSUM_0] THEN have `numpreimages (c:X->complex) S a = 0` [] THEN have `CARD {x:X | x IN S /\ c x = a:complex} = 0` [numpreimages] THEN specialize[`S:X->bool`;`\x:X. c x = a:complex`]FINITE_RESTRICT THEN have `{x:X | x IN S /\ c x = a:complex} = {}` [CARD_EQ_0] THEN ASM SET_TAC[] ; pass ] THEN choose `q:((1->num)->complex)` `(q:(1->num)->complex) IN P /\ complex_root q a` [] THEN subgoal `ring_sum complex_ring P (\p. if a IN complex_root p then Cx (&(e p)) * g a else Cx (&0)) = ring_sum complex_ring P (\p. if p = q then Cx(&(e q)) * g a else Cx(&0))` THENL [ sufficesby RING_SUM_EQ THEN qed[distinct_minpolys_distinct_roots;IN] ; pass ] THEN simp[] THEN rw[GSYM complex_ring_clauses] THEN simp[RING_SUM_DELTA] THEN rw[in_complex_ring] THEN rw[ring_of_num_complex;complex_of_num] THEN subgoal `e(q:(1->num)->complex) = numpreimages c S (c(v:X):complex)` THENL [ rw[GSYM(know `!z. nsum P (\p. if complex_root p z then e p else 0) = numpreimages (c:X->complex) S z`)] THEN subgoal `nsum P (\p. if complex_root p (c(v:X)) then e p else 0) = nsum P (\p. if p = q then e q else 0)` THENL [ sufficesby NSUM_EQ THEN qed[distinct_minpolys_distinct_roots] ; pass ] THEN simp[] THEN simp[NSUM_DELTA] ; pass ] THEN qed[] );; let sum_root_decomposition_if_monic_vanishing_at_factorization = prove(` !P:((1->num)->complex)->bool S:X->bool e:((1->num)->complex)->num c:X->complex g:complex->complex. FINITE P ==> distinct_minpolys P ==> FINITE S ==> poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = monic_vanishing_at complex_ring S c ==> ring_sum complex_ring P (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g) = ring_sum complex_ring S (g o c) `, intro THEN def `Q:((1->num)->complex)->bool` `{p:(1->num)->complex | p IN P /\ ~(e p = 0)}` THEN havetac `Q SUBSET (P:((1->num)->complex)->bool)` (ASM SET_TAC[]) THEN subgoal `FINITE(Q:((1->num)->complex)->bool)` THENL [ specialize[ `P:((1->num)->complex)->bool`; `\p:(1->num)->complex. ~(e p = 0)` ]FINITE_RESTRICT THEN qed[] ; pass ] THEN subgoal `distinct_minpolys Q` THENL [ ASM SET_TAC[distinct_minpolys] ; pass ] THEN subgoal `poly_product QinC_ring Q (\p. poly_pow QinC_ring p (e p)) = monic_vanishing_at complex_ring (S:X->bool) c` THENL [ specialize[ `QinC_ring`; `Q:((1->num)->complex)->bool`; `P:((1->num)->complex)->bool`; `\p:(1->num)->complex. poly_pow QinC_ring p (e p)` ](GSYM poly_product_restrict_subset) THEN simp[] THEN rw[IN_ELIM_THM] THEN subgoal `poly_product QinC_ring P (\s. if s IN P /\ ~(e s = 0) then poly_pow QinC_ring s (e s) else poly_1 QinC_ring) = poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p))` THENL [ sufficesby poly_product_eq THEN qed[poly_pow_0] ; pass ] THEN qed[] ; pass ] THEN subgoal `!p:(1->num)->complex. p IN Q ==> ~(e p = 0)` THENL [ simp[IN_ELIM_THM] ; pass ] THEN specialize[ `Q:((1->num)->complex)->bool`; `S:X->bool`; `e:((1->num)->complex)->num`; `c:X->complex`; `g:complex->complex` ]sum_root_decomposition_if_monic_vanishing_at_factorization_lemma THEN subgoal `ring_sum complex_ring Q (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g) = ring_sum complex_ring P (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g)` THENL [ specialize[ `complex_ring`; `Q:((1->num)->complex)->bool`; `P:((1->num)->complex)->bool`; `\p:(1->num)->complex. Cx(&(e p)) * ring_sum complex_ring (complex_root p) g` ]ring_sum_restrict_subset THEN subgoal `ring_sum complex_ring P (\s. if s IN Q then Cx (&(e s)) * ring_sum complex_ring (complex_root s) g else ring_0 complex_ring) = ring_sum complex_ring P (\p. Cx (&(e p)) * ring_sum complex_ring (complex_root p) g)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN simp[IN_ELIM_THM] THEN case `e(a:(1->num)->complex) = 0` THENL [ rw[complex_ring_clauses] THEN have `Cx (&0) * ring_sum complex_ring (complex_root a) g = Cx(&0)` [COMPLEX_MUL_LZERO] THEN simp[] ; pass ] THEN simp[] ; pass ] THEN qed[] ; pass ] THEN qed[] );; (* ===== core transcendence arguments *) let transcendence_v_denom = prove(` !P D B u. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> !n vn. vn = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) ==> D pow n * vn IN ZinC `, intro THEN have `!n. D pow n * u n IN ZinC` [weighted_powersums_distinct_minpolys] THEN subgoal `D pow n * vn = ring_sum complex_ring (0..n) (\i. D pow n * Cx(&(FACT(n-i) * binom(n,i))) * u i)` THENL [ have `D pow n IN ring_carrier complex_ring` [in_complex_ring] THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN have `!i. i IN (0..n) ==> Cx (&(FACT (n - i) * binom (n,i))) * u i IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`\i. Cx (&(FACT (n - i) * binom (n,i))) * u i`;`(D:complex) pow n`;`0..n`](GSYM RING_SUM_LMUL) THEN have `D pow n * vn = ring_sum complex_ring (0..n) (\x. ring_mul complex_ring (D pow n) (Cx (&(FACT (n - x) * binom (n,x))) * u x))` [complex_ring_clauses] THEN qed[RING_SUM_EQ;complex_ring_clauses] ; pass ] THEN recall subring_complex_empty THEN subgoal `ring_sum complex_ring (0..n) (\i. D pow n * Cx (&(FACT (n - i) * binom (n,i))) * u i) IN ring_carrier(subring_generated complex_ring {})` THENL [ sufficesby ring_sum_in_subring THEN intro THEN have `s <= n:num` [IN_NUMSEG_0] THEN num_linear_fact `s <= n ==> n = s + (n-s):num` THEN have `(D:complex) pow n = D pow s * D pow (n-s)` [COMPLEX_POW_ADD] THEN complex_field_fact `(D:complex) pow n = D pow s * D pow (n-s) ==> D pow n * Cx (&(FACT (n - s) * binom (n,s))) * u s = D pow (n-s) * Cx (&(FACT (n - s) * binom (n,s))) * (D pow s * u s)` THEN have `D pow s * u s IN ZinC` [] THEN have `D pow (n-s) IN ZinC` [ring_pow_ZinC;RING_POW;ZinC_ring_clauses] THEN have `Cx(&(FACT(n-s) * binom(n,s))) IN ZinC` [ZinC_ring_clauses;num_in_ZinC] THEN have `Cx (&(FACT (n - s) * binom (n,s))) * (D pow s * u s) IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `D pow (n-s) * Cx (&(FACT (n - s) * binom (n,s))) * (D pow s * u s) IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `D pow n * Cx (&(FACT (n - s) * binom (n,s))) * u s IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN qed[ZinC_ring_clauses] ; pass ] THEN qed[ZinC_ring_clauses] );; (* could do transcendence_v_denom by specializing to k = 0 *) let transcendence_w_denom = prove(` !P D B u. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> !n k wn. wn = (if k <= n then ring_sum complex_ring (0..n-k) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) else Cx(&0)) ==> (D pow (n-k) * wn) / Cx(&(FACT k)) IN ZinC `, intro THEN case `k <= n:num` THENL [ have `!i. D pow i * u i IN ZinC` [weighted_powersums_distinct_minpolys] THEN rw[COMPLEX_FIELD `(D pow (n-k) * wn) / Cx(&(FACT k)) = (D pow (n-k) / Cx(&(FACT k))) * wn`] THEN subgoal `(D pow (n-k) / Cx(&(FACT k))) * wn = ring_sum complex_ring (0..n-k) (\i. (D pow (n-k) / Cx(&(FACT k))) * Cx(&(FACT(n-i) * binom(n,i))) * u i)` THENL [ have `(D pow (n-k) / Cx(&(FACT k))) IN ring_carrier complex_ring` [in_complex_ring] THEN have `FINITE (0..n-k)` [FINITE_NUMSEG] THEN have `!i. i IN (0..n-k) ==> Cx (&(FACT (n - i) * binom (n,i))) * u i IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`\i. Cx (&(FACT (n - i) * binom (n,i))) * u i`;`(D:complex) pow (n-k) / Cx(&(FACT k))`;`0..n-k`](GSYM RING_SUM_LMUL) THEN have `(D pow (n-k) / Cx(&(FACT k))) * wn = ring_sum complex_ring (0..n-k) (\x. ring_mul complex_ring ((D pow (n-k) / Cx(&(FACT k)))) (Cx (&(FACT (n - x) * binom (n,x))) * u x))` [complex_ring_clauses] THEN qed[RING_SUM_EQ;complex_ring_clauses] ; pass ] THEN recall subring_complex_empty THEN subgoal `ring_sum complex_ring (0..n-k) (\i. D pow (n-k) / Cx(&(FACT k)) * Cx (&(FACT (n - i) * binom (n,i))) * u i) IN ring_carrier(subring_generated complex_ring {})` THENL [ sufficesby ring_sum_in_subring THEN intro THEN have `s <= n-k:num` [IN_NUMSEG_0] THEN num_linear_fact `s <= n-k ==> n-k = s + ((n-k)-s):num` THEN have `(D:complex) pow (n-k) = D pow s * D pow ((n-k)-s)` [COMPLEX_POW_ADD] THEN num_linear_fact `k <= n ==> s <= n-k ==> n-s = ((n-s)-k)+k:num` THEN have `FACT(((n-s)-k)+k) = FACT((n-s)-k) * FACT k * binom(((n-s)-k)+k,k)` [BINOM_FACT] THEN have `FACT(n-s) = FACT((n-s)-k) * FACT k * binom(((n-s)-k)+k,k)` [] THEN have `((n-s)-k)+k,k = n-s,(k:num)` [] THEN have `FACT(n-s) = FACT((n-s)-k) * FACT k * binom(n-s,k)` [] THEN have `(&(FACT(n-s))) = (&(FACT((n-s)-k))) * (&(FACT k)) * (&(binom(n-s,k)))` [REAL_OF_NUM_MUL] THEN have `Cx(&(FACT(n-s))) = Cx(&(FACT((n-s)-k))) * Cx(&(FACT k)) * Cx(&(binom(n-s,k)))` [CX_MUL] THEN have `~(FACT k = 0)` [FACT_NZ] THEN have `~(&(FACT k) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT k)) = Cx(&0))` [CX_INJ] THEN complex_field_fact ` ~(Cx(&(FACT k)) = Cx(&0)) ==> (D:complex) pow (n-k) = D pow s * D pow ((n-k)-s) ==> Cx(&(FACT(n-s))) = Cx(&(FACT((n-s)-k))) * Cx(&(FACT k)) * Cx(&(binom(n-s,k))) ==> D pow (n-k) / Cx(&(FACT k)) * (Cx(&(FACT (n - s))) * Cx(&(binom (n,s)))) * u s = D pow ((n-k)-s) * Cx(&(FACT ((n-s)-k))) * Cx(&(binom(n-s,k))) * Cx(&(binom (n,s))) * D pow s * u s ` THEN have `(&(FACT(n-s) * binom(n,s))) = (&(FACT(n-s))) * (&(binom(n,s)))` [REAL_OF_NUM_MUL] THEN have `Cx(&(FACT(n-s) * binom(n,s))) = Cx(&(FACT(n-s))) * Cx(&(binom(n,s)))` [CX_MUL] THEN have `D pow ((n-k)-s) IN ZinC` [ring_pow_ZinC;RING_POW;ZinC_ring_clauses] THEN have `Cx(&(FACT(n-s-k))) IN ZinC` [ZinC_ring_clauses;num_in_ZinC] THEN have `Cx(&(binom(n-s,k))) IN ZinC` [ZinC_ring_clauses;num_in_ZinC] THEN have `Cx(&(binom(n,s))) IN ZinC` [ZinC_ring_clauses;num_in_ZinC] THEN have `D pow s * u s IN ZinC` [] THEN have `Cx (&(binom (n,s))) * D pow s * u s IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `Cx (&(binom (n - s,k))) * Cx (&(binom (n,s))) * D pow s * u s IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `Cx (&(FACT (n - s - k))) * Cx (&(binom (n - s,k))) * Cx (&(binom (n,s))) * D pow s * u s IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `D pow ((n-k)-s) * Cx (&(FACT (n - s - k))) * Cx (&(binom (n - s,k))) * Cx (&(binom (n,s))) * D pow s * u s IN ZinC` [ZinC_ring_clauses;RING_MUL] THEN have `D pow (n-k) / Cx(&(FACT k)) * (Cx (&(FACT (n - s))) * Cx (&(binom (n,s)))) * u s IN ZinC` [] THEN have `D pow (n-k) / Cx(&(FACT k)) * Cx (&(FACT (n - s) * binom (n,s))) * u s IN ZinC` [] THEN qed[ZinC_ring_clauses] ; pass ] THEN qed[ZinC_ring_clauses] ; have `wn = Cx(&0)` [] THEN qed[COMPLEX_FIELD `(x * Cx(&0)) / y = Cx(&0)`;ZinC_0] ] );; let transcendence_H_poly = prove(` !P H t. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> ring_polynomial complex_ring H `, intro THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring (I p)` [I_THM] THEN have `ring_polynomial QinC_ring (poly_product QinC_ring P I)` [poly_product_poly] THEN have `ring_powerseries QinC_ring (poly_product QinC_ring P I)` [ring_polynomial] THEN qed[x_truncreverse_poly;poly_complex_if_poly_QinC] );; let transcendence_Gp_poly = prove(` !P p. FINITE P ==> distinct_minpolys P ==> p IN P ==> ring_polynomial complex_ring (poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q)) `, intro THEN have `!q:(1->num)->complex. q IN P DELETE p ==> q IN P` [IN_DELETE] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_polynomial QinC_ring q` [distinct_minpolys] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_powerseries QinC_ring q` [ring_polynomial] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_polynomial QinC_ring (x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [x_truncreverse_poly] THEN have `FINITE (P DELETE (p:(1->num)->complex))` [FINITE_DELETE] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [poly_complex_if_poly_QinC] THEN qed[poly_product_poly] );; let transcendence_H_product = prove(` !P H t. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> H = poly_product QinC_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p) `, intro THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring (I p)` [I_THM] THEN have `!p:(1->num)->complex. p IN P ==> poly_deg QinC_ring (I p) <= poly_deg complex_ring p` [I_THM;poly_deg_subring;subring_complex_QinC;ARITH_RULE `d <= d:num`] THEN subgoal `poly_product QinC_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p) = poly_product QinC_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) (I p))` THENL [ sufficesby poly_product_eq THEN qed[I_THM] ; pass ] THEN specialize[`QinC_ring`;`\p:(1->num)->complex. poly_deg complex_ring p`;`I:((1->num)->complex)->(1->num)->complex`;`P:((1->num)->complex)->bool`]x_truncreverse_product THEN qed[] );; let transcendence_H_product_complex_ring = prove(` !P H t. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> H = poly_product complex_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p) `, intro THEN specialize[`P:((1->num)->complex)->bool`;`H:(1->num)->complex`;`t:num`]transcendence_H_product THEN have `!q. q IN P ==> ring_powerseries QinC_ring (q:(1->num)->complex)` [distinct_minpolys;ring_polynomial] THEN have `!q. q IN P ==> ring_powerseries QinC_ring (x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [x_truncreverse_series] THEN specialize_assuming[`complex_ring`;`QinC`;`\q:(1->num)->complex. x_truncreverse QinC_ring (poly_deg complex_ring q) q`;`P:((1->num)->complex)->bool`]poly_product_subring THEN qed[subring_complex_QinC] );; let transcendence_H_botcoeff1 = prove(` !P H t. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> coeff 0 H = Cx(&1) `, intro THEN have `!p:(1->num)->complex. p IN P ==> monic QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> coeff (poly_deg QinC_ring p - 0) p = ring_1 QinC_ring` [monic;ARITH_RULE `x - 0 = x`] THEN have `!p:(1->num)->complex. p IN P ==> coeff 0 (x_truncreverse QinC_ring (poly_deg QinC_ring p) p) = ring_1 QinC_ring` [coeff_x_truncreverse;ARITH_RULE `0 <= x`] THEN have `!p:(1->num)->complex. p IN P ==> coeff 0 (x_truncreverse QinC_ring (poly_deg complex_ring p) p) = ring_1 QinC_ring` [poly_deg_subring;subring_complex_QinC] THEN have `H = poly_product QinC_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [transcendence_H_product] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [x_truncreverse_poly;ring_polynomial] THEN specialize_assuming[`QinC_ring`;`\p:(1->num)->complex. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`P:((1->num)->complex)->bool`]poly_product_botcoeff1 THEN have `coeff 0 (H:(1->num)->complex) = ring_1 QinC_ring` [] THEN qed[QinC_ring_clauses] );; let transcendence_H_denom = prove(` !P H D t. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!i. D pow i * coeff i H IN ZinC) `, intro THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> monic QinC_ring p` [distinct_minpolys] THEN have `!p i. p IN P ==> D pow i * coeff i (x_truncreverse QinC_ring (poly_deg complex_ring p) p) IN ZinC` [denominator_reverse] THEN have `!p:(1->num)->complex. p IN P ==> ring_powerseries QinC_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [ring_polynomial;x_truncreverse_series] THEN specialize[`\p:(1->num)->complex. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`D:complex`;`P:((1->num)->complex)->bool`]denominator_reverse_product THEN qed[transcendence_H_product] );; let transcendence_H_pow_denom = prove(` !P H D t k. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!i. D pow i * coeff i (poly_pow QinC_ring H k) IN ZinC) `, intro THEN have `!i. D pow i * coeff i H IN ZinC` [transcendence_H_denom] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> ring_powerseries QinC_ring (I p)` [ring_polynomial;I_THM] THEN have `ring_powerseries QinC_ring (poly_product QinC_ring P I:(1->num)->complex)` [poly_product_series] THEN have `ring_powerseries QinC_ring (H:(1->num)->complex)` [x_truncreverse_series] THEN qed[denominator_reverse_pow] );; let transcendence_newton_pe = prove(` !P p e. FINITE P ==> distinct_minpolys P ==> p IN P ==> poly_mul complex_ring ( poly_pow complex_ring ( x_truncreverse complex_ring (poly_deg complex_ring p) p ) (e + 1) ) ( poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring ( Cx(&(FACT e * binom (n,e))) ) ( ring_pow complex_ring s (n - e) ) ) ) ) = scaled_pow_newton_rightside complex_ring I (complex_root p) e `, intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_irreducible(x_poly QinC_ring) p` [distinct_minpolys] THEN have `monic QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN recall subring_complex_QinC THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring] THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [monic_poly_0;field_complex;field] THEN have `ring_squarefree(x_poly QinC_ring) p` [squarefree_if_irreducible_over_field;field_QinC] THEN have `ring_squarefree(x_poly complex_ring) p` [monic_QinC_squarefree_complex_squarefree] THEN have `monic_vanishing_at complex_ring (complex_root p) I = p` [monic_vanishing_at_complex_root] THEN have `FINITE (complex_root p)` [distinct_minpolys_finite_root] THEN have `!c. c IN complex_root p ==> I c IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`I:complex->complex`;`complex_root p`]deg_monic_vanishing_at THEN have `poly_deg complex_ring (p:(1->num)->complex) = CARD(complex_root p)` [distinct_minpolys_card_root] THEN specialize[`complex_ring`;`I:complex->complex`;`complex_root p`;`e:num`]scaled_pow_newton_identities_monic_vanishing_at THEN subgoal `poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (ring_of_num complex_ring (FACT e * binom (n,e))) (ring_pow complex_ring (I s) (n - e)))) = poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))` THENL [ sufficesby poly_sum_eq THEN intro THEN sufficesby eq_coeff THEN intro THEN simp[coeff_series_from_coeffs] THEN simp[ring_of_num_complex;complex_of_num;I_THM] ; pass ] THEN qed[] );; let transcendence_newton_He_Ge = prove(` !P t H. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> !p Gp. p IN P ==> Gp = poly_product complex_ring (P DELETE p) (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p) ==> !e. poly_mul complex_ring ( poly_pow complex_ring H (e+1) ) ( poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring ( Cx(&(FACT e * binom (n,e))) ) ( ring_pow complex_ring s (n - e) ) ) ) ) = poly_mul complex_ring ( poly_pow complex_ring Gp (e+1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) `, intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [distinct_minpolys] THEN def `revp:(1->num)->complex` `x_truncreverse QinC_ring (poly_deg complex_ring p) p` THEN have `revp = x_truncreverse complex_ring (poly_deg complex_ring p) p` [x_truncreverse_subring;subring_complex_QinC] THEN specialize[`P:((1->num)->complex)->bool`;`H:(1->num)->complex`;`t:num`]transcendence_H_product_complex_ring THEN subgoal `H = poly_mul complex_ring revp Gp:(1->num)->complex` THENL [ specialize_assuming[`complex_ring`;`P:((1->num)->complex)->bool`;`p:(1->num)->complex`;`\q:(1->num)->complex. x_truncreverse QinC_ring (poly_deg complex_ring q) q`]poly_product_delete THEN qed[series_complex] ; pass ] THEN have `H = poly_mul complex_ring Gp revp:(1->num)->complex` [POLY_MUL_SYM;series_complex] THEN def `Z:(1->num)->complex` `poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx(&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))` THEN have `poly_mul complex_ring (poly_pow complex_ring revp (e + 1)) Z = scaled_pow_newton_rightside complex_ring I (complex_root p) e` [transcendence_newton_pe] THEN have `poly_pow complex_ring H (e+1) = poly_mul complex_ring (poly_pow complex_ring Gp (e+1)) (poly_pow complex_ring revp (e+1)):(1->num)->complex` [poly_mul_pow;series_complex] THEN qed[POLY_MUL_ASSOC;series_complex] );; let transcendence_newton_Hk = prove(` !P t H p Gp e k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> p IN P ==> Gp = poly_product complex_ring (P DELETE p) (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p) ==> e < k ==> poly_mul complex_ring ( poly_pow complex_ring H k ) ( poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring ( Cx(&(FACT e * binom (n,e))) ) ( ring_pow complex_ring s (n - e) ) ) ) ) = poly_mul complex_ring ( poly_pow complex_ring H (k-(e+1)) ) ( poly_mul complex_ring ( poly_pow complex_ring Gp (e+1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) `, intro THEN have `poly_mul complex_ring (poly_pow complex_ring H (e+1)) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx(&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))) = poly_mul complex_ring (poly_pow complex_ring Gp (e+1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [transcendence_newton_He_Ge] THEN num_linear_fact `e < k ==> k = (k-(e+1))+(e+1)` THEN have `poly_pow complex_ring (H:(1->num)->complex) k = poly_mul complex_ring (poly_pow complex_ring H (k-(e+1))) (poly_pow complex_ring H (e+1))` [poly_pow_add;series_complex] THEN qed[POLY_MUL_ASSOC;series_complex] );; let transcendence_newton_Hk_psum = prove(` !P B t H G u e k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!p. p IN P ==> G p = poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q)) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> e < k ==> poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) = poly_sum complex_ring P (\p:(1->num)->complex. poly_mul complex_ring ( poly_const complex_ring ((complex_of_int(B p))) ) ( poly_mul complex_ring ( poly_pow complex_ring H (k-(e+1)) ) ( poly_mul complex_ring ( poly_pow complex_ring (G p) (e+1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) ) ) `, intro THEN subgoal `series_from_coeffs (\n. Cx(&(FACT(e) * binom(n,e))) * u(n-e)) = poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx(&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))))` THENL [ sufficesby eq_coeff THEN intro THEN rw[coeff_series_from_coeffs] THEN simp[coeff_poly_sum;series_complex] THEN simp[coeff_poly_const_times;series_complex;in_complex_ring] THEN rw[complex_ring_clauses] THEN subgoal `ring_sum complex_ring P (\a. (complex_of_int(B a)) * coeff d (poly_sum complex_ring (complex_root a) (\s. series_from_coeffs (\n. Cx (&(FACT e * binom (n,e))) * ring_pow complex_ring s (n - e))))) = ring_sum complex_ring P (\a. (complex_of_int(B a)) * ring_sum complex_ring (complex_root a) (\s. Cx (&(FACT e * binom (d,e))) * ring_pow complex_ring s (d - e)))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN have `FINITE(complex_root a)` [distinct_minpolys_finite_root] THEN simp[coeff_poly_sum;series_complex] THEN rw[coeff_series_from_coeffs] ; pass ] THEN simp[] THEN subgoal `Cx (&(FACT e * binom (d,e))) * ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p (d - e)) = ring_sum complex_ring P (\p. Cx (&(FACT e * binom (d,e))) * (complex_of_int(B p)) * complex_root_powersums p (d - e))` THENL [ rw[GSYM complex_ring_clauses] THEN specialize_assuming[`complex_ring`;`\p. ring_mul complex_ring ((complex_of_int(B p))) (complex_root_powersums p (d - e))`;`Cx(&(FACT e * binom(d,e)))`;`P:((1->num)->complex)->bool`]RING_SUM_LMUL THEN qed[in_complex_ring] ; pass ] THEN simp[] THEN sufficesby RING_SUM_EQ THEN intro THEN rw[complex_root_powersums;ring_pow_complex] THEN subgoal `ring_sum complex_ring (complex_root a) (\s. Cx (&(FACT e * binom (d,e))) * s pow (d - e)) = Cx (&(FACT e * binom (d,e))) * ring_sum complex_ring (complex_root a) (\s. s pow (d - e))` THENL [ rw[GSYM complex_ring_clauses] THEN have `FINITE(complex_root a)` [distinct_minpolys_finite_root] THEN specialize_assuming[`complex_ring`;`\s:complex. s pow (d-e)`;`Cx(&(FACT e * binom(d,e)))`;`complex_root a`]RING_SUM_LMUL THEN qed[in_complex_ring] ; pass ] THEN simp[] THEN complex_field_fact `Cx (&(FACT e * binom (d,e))) * (complex_of_int(B a)) * ring_sum complex_ring (complex_root a) (\z. z pow (d - e)) = (complex_of_int(B a)) * Cx (&(FACT e * binom (d,e))) * ring_sum complex_ring (complex_root a) (\z. z pow (d - e))` THEN qed[] ; pass ] THEN subgoal `poly_mul complex_ring (poly_pow complex_ring H k) (poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))))) = poly_sum complex_ring P (\p. poly_mul complex_ring (poly_pow complex_ring H k) (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))))` THENL [ specialize_assuming[`complex_ring`;`\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))`;`poly_pow complex_ring H k:(1->num)->complex`;`P:((1->num)->complex)->bool`]poly_sum_lmul THEN qed[series_complex] ; pass ] THEN subgoal `poly_sum complex_ring P (\p. poly_mul complex_ring (poly_pow complex_ring H k) (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_sum complex_ring (complex_root p) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))))) = poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))` THENL [ sufficesby poly_sum_eq THEN intro THEN rw[BETA_THM] THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))) = poly_mul complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) (poly_const complex_ring ((complex_of_int(B s))))) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))` [POLY_MUL_ASSOC;series_complex] THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_const complex_ring ((complex_of_int(B(s:(1->num)->complex))))) = poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_pow complex_ring H k):(1->num)->complex` [POLY_MUL_SYM;series_complex] THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))) = poly_mul complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_pow complex_ring H k)) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))` [] THEN have `poly_mul complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_pow complex_ring H k)) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))) = poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_mul complex_ring (poly_pow complex_ring H k) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))))` [POLY_MUL_ASSOC;series_complex] THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e)))))) = poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B s)))) (poly_mul complex_ring (poly_pow complex_ring H k) (poly_sum complex_ring (complex_root s) (\s. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring s (n - e))))))` [] THEN have `G(s:(1->num)->complex) = poly_product complex_ring (P DELETE s) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [] THEN subgoal `poly_product complex_ring (P DELETE s) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q) = poly_product complex_ring (P DELETE s) (\q. x_truncreverse complex_ring (poly_deg complex_ring q) q)` THENL [ sufficesby poly_product_eq THEN qed[x_truncreverse_subring;subring_complex_QinC] ; pass ] THEN have `G(s:(1->num)->complex) = poly_product complex_ring (P DELETE s) (\q. x_truncreverse complex_ring (poly_deg complex_ring q) q)` [] THEN specialize_assuming[`P:((1->num)->complex)->bool`;`t:num`;`H:(1->num)->complex`;`s:(1->num)->complex`;`G(s:(1->num)->complex):(1->num)->complex`;`e:num`;`k:num`]transcendence_newton_Hk THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_sum complex_ring (complex_root s) (\z. series_from_coeffs (\n. ring_mul complex_ring (Cx (&(FACT e * binom (n,e)))) (ring_pow complex_ring z (n - e))))) = poly_mul complex_ring (poly_pow complex_ring H (k-(e+1))) (poly_mul complex_ring (poly_pow complex_ring (G s) (e+1)) (scaled_pow_newton_rightside complex_ring I (complex_root s) e))` [] THEN qed[] ; pass ] THEN qed[] );; let transcendence_Huv_lemma = prove(` !H u k. poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) ) = poly_sum complex_ring {e | e < k} (\e:num. poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. Cx(&(FACT e * binom (n,e))) * u(n-e)) ) ) `, intro THEN have `FINITE {e:num | e < k}` [FINITE_NUMSEG_LT] THEN subgoal `series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e))) = poly_sum complex_ring {e | e < k} (\e:num. series_from_coeffs (\n. Cx(&(FACT e * binom (n,e))) * u(n-e)))` THENL [ sufficesby eq_coeff THEN intro THEN rw[coeff_series_from_coeffs] THEN simp[coeff_poly_sum;series_complex] THEN sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[coeff_series_from_coeffs] ; pass ] THEN specialize[`complex_ring` ;`\e:num. series_from_coeffs (\n. Cx(&(FACT e * binom (n,e))) * u(n-e))` ]poly_sum_lmul THEN have ` poly_sum complex_ring {e | e < k} (\e:num. poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. Cx(&(FACT e * binom (n,e))) * u(n-e)) ) ) = poly_mul complex_ring ( poly_pow complex_ring H k ) ( poly_sum complex_ring {e | e < k} (\e:num. series_from_coeffs (\n. Cx(&(FACT e * binom (n,e))) * u(n-e)) ) ) ` [series_complex] THEN qed[] );; let transcendence_H_main = prove(` !P B t H G u k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!p. p IN P ==> G p = poly_product complex_ring (P DELETE p) (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p)) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) ) = poly_sum complex_ring {e | e < k} (\e:num. poly_sum complex_ring P (\p:(1->num)->complex. poly_mul complex_ring ( poly_const complex_ring ((complex_of_int(B p))) ) ( poly_mul complex_ring ( poly_pow complex_ring H (k-(e+1)) ) ( poly_mul complex_ring ( poly_pow complex_ring (G p) (e+1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) ) ) ) `, intro THEN rw[transcendence_Huv_lemma] THEN sufficesby poly_sum_eq THEN intro THEN rw[BETA_THM] THEN set_fact `s IN {e | e < k} ==> s < k:num` THEN qed[transcendence_newton_Hk_psum] );; let transcendence_Htimes_poly = prove(` !P B t H u k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> ring_polynomial complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) ) ) `, intro THEN def `G:((1->num)->complex)->((1->num)->complex)` `\p. poly_product complex_ring (P DELETE p) (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p)` THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e)))) = poly_sum complex_ring {e | e < k} (\e. poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))))` [transcendence_H_main] THEN have `FINITE {e:num | e < k}` [FINITE_NUMSEG_LT] THEN subgoal `!e. e IN {e | e < k} ==> ring_polynomial complex_ring (poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))))` THENL [ intro THEN subgoal `!p:(1->num)->complex. p IN P ==> ring_polynomial complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))` THENL [ intro THEN have `ring_polynomial complex_ring (poly_const complex_ring ((complex_of_int(B(p:(1->num)->complex)))):(1->num)->complex)` [RING_POLYNOMIAL_CONST;in_complex_ring] THEN have `ring_polynomial complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex)` [transcendence_H_poly;poly_pow_poly;poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex)` [poly_pow_poly;transcendence_Gp_poly;poly_complex_if_poly_QinC] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [poly_scaled_pow_newton_rightside;in_complex_ring] THEN qed[RING_POLYNOMIAL_MUL] ; pass ] THEN qed[poly_sum_poly] ; pass ] THEN qed[transcendence_H_main;poly_sum_poly] );; let transcendence_H_deg_G_lemma = prove(` !p d g s t n:num. ~(p = 0) ==> d <= g + s ==> g <= (e+1)*n ==> s <= e + (p-1)*(e+1) ==> t = p+n ==> d <= e+(t-1)*(e+1) `, intro THEN num_linear_fact `~(p = 0) ==> (p+n)-1 = n+(p-1)` THEN num_linear_fact `d <= g + s ==> g <= (e+1)*n ==> s <= e + (p-1)*(e+1) ==> d <= e+(n+(p-1))*(e+1)` THEN qed[] );; let transcendence_H_deg_G = prove(` !P t p Gp e. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> p IN P ==> Gp = poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q) ==> poly_deg complex_ring ( poly_mul complex_ring ( poly_pow complex_ring Gp (e + 1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) <= e + (t-1)*(e+1) `, intro THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `CARD(complex_root p) = poly_deg complex_ring p` [distinct_minpolys_card_root] THEN have `poly_deg complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e) <= e + (CARD(complex_root p)-1)*(e+1)` [deg_scaled_pow_newton_rightside;in_complex_ring] THEN have `poly_deg complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e) <= e + ((poly_deg complex_ring p)-1)*(e+1)` [deg_scaled_pow_newton_rightside;in_complex_ring] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> poly_deg complex_ring (x_truncreverse complex_ring (poly_deg complex_ring q) q) <= poly_deg complex_ring q` [deg_x_truncreverse_le;series_complex] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_polynomial complex_ring (x_truncreverse complex_ring (poly_deg complex_ring q) q)` [x_truncreverse_poly;series_complex] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> poly_deg complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring q) q) <= poly_deg complex_ring q` [x_truncreverse_subring;subring_complex_QinC] THEN have `!q:(1->num)->complex. q IN P DELETE p ==> ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [x_truncreverse_subring;subring_complex_QinC] THEN have `FINITE(P DELETE (p:(1->num)->complex))` [FINITE_DELETE] THEN specialize[`complex_ring`;`\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q`;`\q:(1->num)->complex. poly_deg complex_ring q`;`P DELETE (p:(1->num)->complex)`]poly_deg_product_le THEN have `poly_deg complex_ring (Gp:(1->num)->complex) <= nsum(P DELETE p) (\q:(1->num)->complex. poly_deg complex_ring q)` [] THEN have `ring_polynomial complex_ring (Gp:(1->num)->complex)` [transcendence_Gp_poly] THEN have `poly_deg complex_ring (poly_pow complex_ring Gp (e+1):(1->num)->complex) <= (e+1) * (poly_deg complex_ring Gp)` [poly_deg_pow_le] THEN have `poly_deg complex_ring (poly_pow complex_ring Gp (e+1):(1->num)->complex) <= (e+1) * nsum(P DELETE p) (\q:(1->num)->complex. poly_deg complex_ring q)` [LE_MULT2;LE_REFL;LE_TRANS] THEN have `t = (poly_deg complex_ring (p:(1->num)->complex)) + nsum(P DELETE p) (\q:(1->num)->complex. poly_deg complex_ring q)` [NSUM_DELETE] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [poly_scaled_pow_newton_rightside;in_complex_ring] THEN have `poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring Gp (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)) <= poly_deg complex_ring (poly_pow complex_ring Gp (e + 1)) + poly_deg complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [POLY_DEG_MUL_LE;poly_pow_poly] THEN have `~(poly_deg complex_ring (p:(1->num)->complex) = 0)` [distinct_minpolys_deg_nonzero] THEN specialize[ `poly_deg complex_ring (p:(1->num)->complex)`; `poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring Gp (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))`; `poly_deg complex_ring (poly_pow complex_ring (Gp:(1->num)->complex) (e + 1))`; `poly_deg complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)`; `t:num`; `nsum (P DELETE p) (\q:(1->num)->complex. poly_deg complex_ring q)` ]transcendence_H_deg_G_lemma THEN qed[] );; let transcendence_H_deg_HG = prove(` !P t p Gp e k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> p IN P ==> Gp = poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q) ==> e < k ==> poly_deg complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H (k - (e + 1)) ) ( poly_mul complex_ring ( poly_pow complex_ring Gp (e + 1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) ) <= e + (t-1)*(e+1) + t*(k-(e+1)) `, intro THEN have `ring_polynomial complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex)` [transcendence_H_poly;poly_pow_poly;poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (Gp:(1->num)->complex) (e + 1):(1->num)->complex)` [poly_pow_poly;transcendence_Gp_poly;poly_complex_if_poly_QinC] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [poly_scaled_pow_newton_rightside;in_complex_ring] THEN have `!p:(1->num)->complex. p IN P ==> ring_powerseries QinC_ring (I p)` [I_THM;ring_polynomial;distinct_minpolys] THEN have `poly_deg QinC_ring (H:(1->num)->complex) <= t` [deg_x_truncreverse_le;poly_product_series] THEN have `poly_deg complex_ring (H:(1->num)->complex) <= t` [poly_deg_subring;subring_complex_QinC] THEN have `ring_polynomial QinC_ring (H:(1->num)->complex)` [x_truncreverse_poly;poly_product_series] THEN have `ring_polynomial complex_ring (H:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (H:(1->num)->complex) (k-(e+1)))` [poly_pow_poly] THEN have `poly_deg complex_ring (poly_pow complex_ring (H:(1->num)->complex) (k-(e+1))) <= (k-(e+1))*(poly_deg complex_ring H)` [poly_deg_pow_le] THEN have `poly_deg complex_ring (poly_pow complex_ring (H:(1->num)->complex) (k-(e+1))) <= (k-(e+1))*t` [LE_MULT2;LE_TRANS;LE_REFL] THEN have `poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring Gp (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))) <= poly_deg complex_ring (poly_pow complex_ring H (k - (e + 1))) + poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring Gp (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))` [POLY_DEG_MUL_LE;RING_POLYNOMIAL_MUL] THEN qed[transcendence_H_deg_G;ARITH_RULE `d <= h+g ==> g <= e + (t-1)*(e+1) ==> h <= (k-(e+1))*t ==> d <= e+(t-1)*(e+1)+t*(k-(e+1))`] );; let transcendence_H_deg_HG_simpler = prove(` !P t p Gp e k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> p IN P ==> Gp = poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q) ==> e < k ==> poly_deg complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H (k - (e + 1)) ) ( poly_mul complex_ring ( poly_pow complex_ring Gp (e + 1) ) ( scaled_pow_newton_rightside complex_ring I (complex_root p) e ) ) ) <= t*k-1 `, intro THEN subgoal `~(t = 0)` THENL [ set_fact `(p:(1->num)->complex) IN P ==> ~(P = {})` THEN qed[distinct_minpolys_total_deg] ; pass ] THEN subgoal `e + (t-1)*(e+1) + t*(k-(e+1)) = t*k-1` THENL [ num_linear_fact `~(t = 0) ==> t = (t-1)+1` THEN have `t*(k-(e+1)) = (t-1)*(k-(e+1))+1*(k-(e+1))` [RIGHT_ADD_DISTRIB] THEN have `1*(k-(e+1)) = k-(e+1)` [MULT_CLAUSES] THEN have `t*(k-(e+1)) = (t-1)*(k-(e+1))+(k-(e+1))` [] THEN have `e + (t-1)*(e+1) + t*(k-(e+1)) = e + (t-1)*(e+1) + (t-1)*(k-(e+1)) + (k-(e+1))` [] THEN num_linear_fact `e < k ==> (e+1) + (k-(e+1)) = k` THEN have `(t-1)*(e+1) + (t-1)*(k-(e+1)) = (t-1)*k` [LEFT_ADD_DISTRIB] THEN num_linear_fact `e < k ==> (t-1)*(e+1) + (t-1)*(k-(e+1)) = (t-1)*k ==> e + (t-1)*(e+1) + (t-1)*(k-(e+1)) + (k-(e+1)) = ((t-1)*k + 1*k) - 1` THEN have `(t-1)*k + 1*k = t*k` [RIGHT_ADD_DISTRIB] THEN qed[] ; pass ] THEN qed[transcendence_H_deg_HG;LE_TRANS] );; let transcendence_H_deg_overall = prove(` !P B t H u k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> poly_deg complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) ) ) <= t*k-1 `, intro THEN def `G:((1->num)->complex)->((1->num)->complex)` `\p. poly_product complex_ring (P DELETE p) (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p)` THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e)))) = poly_sum complex_ring {e | e < k} (\e. poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))))` [transcendence_H_main] THEN have `FINITE {e:num | e < k}` [FINITE_NUMSEG_LT] THEN subgoal `!e p. e IN {e | e < k} ==> p IN P ==> ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))` THENL [ intro THEN have `ring_polynomial complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex)` [transcendence_H_poly;poly_pow_poly;poly_complex_if_poly_QinC] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex)` [poly_pow_poly;transcendence_Gp_poly;poly_complex_if_poly_QinC] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [poly_scaled_pow_newton_rightside;in_complex_ring] THEN qed[RING_POLYNOMIAL_MUL] ; pass ] THEN subgoal `!e p. e IN {e | e < k} ==> p IN P ==> ring_polynomial complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))` THENL [ intro THEN have `ring_polynomial complex_ring (poly_const complex_ring ((complex_of_int(B(p:(1->num)->complex)))):(1->num)->complex)` [RING_POLYNOMIAL_CONST;in_complex_ring] THEN qed[RING_POLYNOMIAL_MUL] ; pass ] THEN have `!e. e IN {e | e < k} ==> ring_polynomial complex_ring (poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))))` [poly_sum_poly] THEN subgoal `!e. e IN {e | e < k} ==> poly_deg complex_ring (poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))) <= t*k-1` THENL [ intro THEN set_fact `e IN {e:num | e < k} ==> e < k` THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))` [] THEN have `!p:(1->num)->complex. p IN P ==> poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))) <= t*k-1` [transcendence_H_deg_HG_simpler] THEN have `!p:(1->num)->complex. p IN P ==> poly_deg complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))) <= poly_deg complex_ring (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))` [deg_const_mul_le;in_complex_ring] THEN have `!p:(1->num)->complex. p IN P ==> poly_deg complex_ring (poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))) <= t*k-1` [LE_TRANS] THEN specialize[`complex_ring`;`\p:(1->num)->complex. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))`;`t*k-1`;`P:((1->num)->complex)->bool`]poly_deg_sum_le THEN qed[] ; pass ] THEN subgoal `poly_deg complex_ring (poly_sum complex_ring {e | e < k} (\e. poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))))) <= t*k-1` THENL [ specialize_assuming[`complex_ring`;`\e. poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring ((complex_of_int(B p)))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))`;`t*k-1`;`{e:num | e < k}`]poly_deg_sum_le THEN qed[] ; pass ] THEN qed[transcendence_H_main] );; let transcendence_vw_diff = prove(` !u n k vn wn. vn = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) ==> wn = (if k <= n then ring_sum complex_ring (0..n-k) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) else Cx(&0)) ==> vn - wn = ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u (n-e)) `, intro THEN case `k <= n:num` THENL [ num_linear_fact `k <= n ==> n-k <= n:num` THEN specialize[`complex_ring`;`n-k:num`;`n:num`;`\i. Cx(&(FACT(n-i) * binom(n,i))) * u i`]ring_sum_numseg_0_diff_reflect THEN num_linear_fact `k <= n ==> n-(n-k) = k:num` THEN subgoal `{e:num | e < k} = {e | e < n-(n-k)}` THENL [ rw[EXTENSION;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `ring_sum complex_ring {e:num | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e)) = ring_sum complex_ring {i | i < k} (\i. Cx (&(FACT (n - (n - i)) * binom (n,n - i))) * u (n - i))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN set_fact `a IN {e:num | e < k} ==> a < k` THEN num_linear_fact `a < k:num ==> k <= n ==> n-(n-a) = a` THEN num_linear_fact `a < k:num ==> k <= n ==> a <= n` THEN have `binom(n,n-a) = binom(n,a)` [BINOM_SYM] THEN have `FACT(n-(n-a)) = FACT(a)` [] THEN qed[] ; pass ] THEN have `ring_sub complex_ring (ring_sum complex_ring (0..n) (\i. Cx (&(FACT (n - i) * binom (n,i))) * u i)) (ring_sum complex_ring (0..n - k) (\i. Cx (&(FACT (n - i) * binom (n,i))) * u i)) = ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e))` [] THEN qed[ring_sub_complex] ; have `wn = Cx(&0)` [] THEN complex_field_fact `wn = Cx(&0) ==> vn - wn = vn` THEN specialize[`complex_ring`;`n:num`;`\i. Cx (&(FACT (n - i) * binom (n,i))) * u i`]ring_sum_numseg_le_reflect THEN have `vn - wn = ring_sum complex_ring (0..n) (\i. Cx (&(FACT (n - (n-i)) * binom (n,n-i))) * u (n-i))` [] THEN subgoal `ring_sum complex_ring (0..n) (\i. Cx (&(FACT (n - (n-i)) * binom (n,n-i))) * u (n-i)) = ring_sum complex_ring (0..n) (\i. Cx (&(FACT i * binom (n,i))) * u (n-i))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN have `a <= n:num` [IN_NUMSEG] THEN num_linear_fact `a <= n:num ==> n-(n-a) = a` THEN have `FACT (n-(n-a)) = FACT a` [] THEN have `binom(n,n-a) = binom(n,a)` [BINOM_SYM] THEN qed[] ; pass ] THEN have `vn - wn = ring_sum complex_ring (0..n) (\i. Cx (&(FACT i * binom (n,i))) * u (n-i))` [] THEN num_linear_fact `~(k <= n) ==> n <= k-1` THEN specialize[`complex_ring`;`n:num`;`k-1`;`\i. Cx (&(FACT i * binom (n,i))) * u (n-i)`]ring_sum_numseg_le_expand THEN subgoal `ring_sum complex_ring (0..k - 1) (\a. if a <= n then Cx (&(FACT a * binom (n,a))) * u (n - a) else ring_0 complex_ring) = ring_sum complex_ring (0..k - 1) (\a. Cx (&(FACT a * binom (n,a))) * u (n - a))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN proven_if `a <= n:num` [] THEN num_linear_fact `~(a <= n) ==> n < a:num` THEN simp[BINOM_LT;MULT_0;COMPLEX_MUL_LZERO;complex_ring_clauses] ; pass ] THEN have `vn - wn = ring_sum complex_ring (0..k - 1) (\a. Cx (&(FACT a * binom (n,a))) * u (n - a))` [] THEN num_linear_fact `~(k <= n) ==> ~(k = 0)` THEN have `{e:num | e < k} = (0..k-1)` [NUMSEG_LT] THEN qed[] ] );; let transcendence_vw_diff_series = prove(` !P B u v w. FINITE P ==> distinct_minpolys P ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n. w n = if k <= n then ring_sum complex_ring (0..n-k) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) else Cx(&0)) ==> poly_sub complex_ring ( series_from_coeffs v ) ( series_from_coeffs w ) = series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_poly_sub;coeff_series_from_coeffs] THEN rw[ring_sub_complex] THEN qed[transcendence_vw_diff] );; let transcendence_vw_diff_series_H = prove(` !P B t H u v w k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n. w n = if k <= n then ring_sum complex_ring (0..n-k) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) else Cx(&0)) ==> poly_sub complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs w ) ) = poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx(&(FACT(e) * binom(n,e))) * u(n-e) ) ) ) `, intro THEN specialize[`P:((1->num)->complex)->bool`;`B:((1->num)->complex)->int`;`u:num->complex`;`v:num->complex`;`w:num->complex`]transcendence_vw_diff_series THEN specialize_assuming[`complex_ring`;`poly_pow complex_ring H k:(1->num)->complex`;`series_from_coeffs v:(1->num)->complex`;`series_from_coeffs w:(1->num)->complex`]poly_sub_ldistrib THEN have `poly_mul complex_ring (poly_pow complex_ring H k) (poly_sub complex_ring (series_from_coeffs v) (series_from_coeffs w)) = poly_sub complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w)):(1->num)->complex` [series_complex] THEN qed[] );; let transcendence_uv_trivial = prove(` !P B u v. P = {} ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n. u n = Cx(&0) /\ v n = Cx(&0)) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN have `!d:num. u d = Cx(&0)` [RING_SUM_CLAUSES;complex_ring_clauses] THEN subgoal `!d:num. ring_sum complex_ring (0..d) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) = Cx(&0)` THENL [ intro THEN rw[GSYM complex_ring_clauses] THEN sufficesby RING_SUM_EQ_0 THEN intro THEN rw[BETA_THM] THEN rw[complex_ring_clauses] THEN qed[COMPLEX_MUL_RZERO] ; qed[] ] );; let transcendence_vw_match = prove(` !P B t H u v w. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n k. k*t <= n ==> (!d. w d = if k <= d then ring_sum complex_ring (0..d-k) (\i. Cx(&(FACT(d-i) * binom(d,i))) * u i) else Cx(&0)) ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs w ) ) ) `, intro THEN case `P = {}:((1->num)->complex)->bool` THENL [ have `!n:num. u n = Cx(&0)` [transcendence_uv_trivial] THEN have `!n:num. v n = Cx(&0)` [transcendence_uv_trivial] THEN subgoal `!d:num. (if k <= d then ring_sum complex_ring (0..d-k) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) else Cx(&0)) = Cx(&0)` THENL [ intro THEN rw[GSYM complex_ring_clauses] THEN rw[COND_ELIM_THM] THEN sufficesby (prove(`x:A = y ==> (if p then x else y) = y`,qed[])) THEN sufficesby RING_SUM_EQ_0 THEN intro THEN rw[BETA_THM] THEN rw[complex_ring_clauses] THEN qed[COMPLEX_MUL_RZERO] ; pass ] THEN have `!n:num. w n = Cx(&0)` [] THEN have `v = w:num->complex` [FUN_EQ_THM] THEN qed[] ; pass ] THEN case `t = 0` THENL [ have `P = {}:((1->num)->complex)->bool` [distinct_minpolys_total_deg] THEN qed[] ; pass ] THEN num_linear_fact `~(t = 0) ==> 1 <= t` THEN have `k <= k:num` [LE_REFL] THEN have `k*1 <= k*t:num` [LE_MULT2] THEN num_linear_fact `k*1 <= k*t ==> k <= k*t` THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex`; `w:num->complex`; `k:num` ]transcendence_vw_diff_series_H THEN specialize[`complex_ring`;`poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)`;`poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w)`;`n:num`]coeff_poly_sub THEN have `(coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))) - (coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w))) = coeff n (poly_sub complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w)))` [ring_sub_complex] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `k:num` ]transcendence_H_deg_overall THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `k:num` ]transcendence_Htimes_poly THEN case `coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e))))) = ring_0 complex_ring` THENL [ have `coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e))))) = Cx(&0)` [complex_ring_clauses] THEN have `coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) - coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w)) = Cx(&0)` [] THEN qed[COMPLEX_FIELD `a - b = Cx(&0) ==> a = b`] ; specialize[`complex_ring`;`poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. ring_sum complex_ring {e | e < k} (\e. Cx (&(FACT e * binom (n,e))) * u (n - e))))`;`t*k-1`;`n:num`]coeff_deg_le THEN num_linear_fact `k*t <= n ==> n <= t*k-1 ==> n = 0` THEN num_linear_fact `k*t <= n ==> n = 0 ==> k*t = 0` THEN have `k = 0` [MULT_EQ_0] THEN subgoal `!d. ring_sum complex_ring (0..d) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) = (if k <= d then ring_sum complex_ring (0..d - k) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) else Cx (&0))` THENL [ intro THEN num_linear_fact `k = 0 ==> k <= d` THEN sufficesby(prove(`k <= d:num /\ x = y:A ==> x = (if k <= d then y else z)`,qed[])) THEN intro THENL [ qed[] ; pass ] THEN rw[know(`k = 0`)] THEN rw[ARITH_RULE `d-0 = d`] ; pass ] THEN have `!d. v d = (if k <= d then ring_sum complex_ring (0..d - k) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) else Cx (&0))` [] THEN have `!d. w d = (if k <= d then ring_sum complex_ring (0..d - k) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) else Cx (&0))` [] THEN specialize[ `v:num->complex`; `w:num->complex`; `\d. (if k <= d then ring_sum complex_ring (0..d - k) (\i. Cx (&(FACT (d - i) * binom (d,i))) * u i) else Cx (&0))` ](prove(`!v w f. (!d:num. v d = f d:complex) ==> (!d:num. w d = f d:complex) ==> (!d:num. v d = w d)`,qed[])) THEN have `v = w:num->complex` [FUN_EQ_THM] THEN qed[] ] );; let transcendence_Hw_denom = prove(` !P D B t H u w. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n k. k*t <= n ==> (!d. w d = if k <= d then ring_sum complex_ring (0..d-k) (\i. Cx(&(FACT(d-i) * binom(d,i))) * u i) else Cx(&0)) ==> (D pow (n-k) / Cx(&(FACT k))) * coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs w ) ) IN ZinC ) `, intro THEN rw[coeff_poly_mul_oneindex] THEN have `(D pow (n-k) / Cx(&(FACT k))) IN ring_carrier complex_ring` [in_complex_ring] THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN have `!i. i IN 0..n ==> ring_mul complex_ring (coeff i (poly_pow complex_ring H k)) (coeff (n - i) (series_from_coeffs w)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`(\a. ring_mul complex_ring (coeff a (poly_pow complex_ring H k)) (coeff (n - a) (series_from_coeffs w)))`;`D pow (n-k) / Cx(&(FACT k))`;`0..n`](GSYM RING_SUM_LMUL) THEN have `(D pow (n - k) / Cx(&(FACT k))) * ring_sum complex_ring (0..n) (\a. ring_mul complex_ring (coeff a (poly_pow complex_ring H k)) (coeff (n - a) (series_from_coeffs w))) = ring_sum complex_ring (0..n) (\x. ring_mul complex_ring (D pow (n - k) / Cx(&(FACT k))) (ring_mul complex_ring (coeff x (poly_pow complex_ring H k)) (coeff (n - x) (series_from_coeffs w))))` [complex_ring_clauses] THEN subgoal `ring_sum complex_ring (0..n) (\x. ring_mul complex_ring (D pow (n - k) / Cx(&(FACT k))) (ring_mul complex_ring (coeff x (poly_pow complex_ring H k)) (coeff (n - x) (series_from_coeffs w)))) IN ZinC` THENL [ rw[GSYM ZinC_ring_clauses;GSYM subring_complex_empty;coeff_series_from_coeffs] THEN sufficesby ring_sum_in_subring THEN rw[complex_ring_clauses;subring_complex_empty] THEN intro THEN rw[BETA_THM] THEN have `s <= n:num` [IN_NUMSEG_0] THEN have `D pow s * coeff s (poly_pow QinC_ring H k) IN ZinC` [transcendence_H_pow_denom] THEN have `!p:(1->num)->complex. p IN P ==> ring_powerseries QinC_ring (I p)` [I_THM;ring_polynomial;distinct_minpolys] THEN have `ring_powerseries QinC_ring (poly_product QinC_ring P I:(1->num)->complex)` [poly_product_series] THEN have `ring_powerseries QinC_ring (H:(1->num)->complex)` [x_truncreverse_series] THEN have `poly_pow complex_ring H k = poly_pow QinC_ring H k:(1->num)->complex` [poly_pow_subring;subring_complex_QinC] THEN have `D pow s * coeff s (poly_pow complex_ring H k) IN ZinC` [] THEN specialize[`P:((1->num)->complex)->bool`;`D:complex`;`B:((1->num)->complex)->int`;`u:num->complex`]transcendence_w_denom THEN specialize_assuming[ `n-s:num`; `k:num`; `w(n-s:num):complex` ]( know(`!n k wn. wn = (if k <= n then ring_sum complex_ring (0..n - k) (\i. Cx (&(FACT (n - i) * binom (n,i))) * u i) else Cx (&0)) ==> (D pow (n - k) * wn) / Cx (&(FACT k)) IN ZinC`) ) THEN have `(D pow ((n-s)-k) * w(n-s)) / Cx(&(FACT k)) IN ZinC` [] THEN case `k <= n-s:num` THENL [ num_linear_fact `s <= n ==> k <= n-s:num ==> n-k = (n-s-k)+s` THEN have `(D:complex) pow (n-k) = D pow (n-s-k) * D pow s` [COMPLEX_POW_ADD] THEN complex_field_fact `D pow (n - k) = D pow (n - s - k) * D pow s ==> D pow (n - k) / Cx (&(FACT k)) * coeff s (poly_pow complex_ring H k) * w (n - s) = (D pow s * coeff s (poly_pow complex_ring H k)) * (D pow ((n-s)-k) * w(n-s)) / Cx(&(FACT k))` THEN qed[mul_in_ZinC;ZinC_ring_clauses] ; have `w(n-s:num) = Cx(&0)` [] THEN simp[COMPLEX_MUL_RZERO] THEN qed[ZinC_ring_clauses;ZinC_0] ] ; pass ] THEN qed[] );; let transcendence_Hv_denom = prove(` !P D B t H u v. FINITE P ==> distinct_minpolys P ==> D IN ZinC ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n k. k*t <= n ==> (D pow (n-k) / Cx(&(FACT k))) * coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) IN ZinC ) `, intro THEN def `w:num->complex` `\n. if k <= n then ring_sum complex_ring (0..n-k) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) else Cx(&0)` THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex`; `w:num->complex` ]transcendence_vw_match THEN have `coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) = coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w))` [] THEN specialize[ `P:((1->num)->complex)->bool`; `D:complex`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `w:num->complex` ]transcendence_Hw_denom THEN have `D pow (n - k) / Cx (&(FACT k)) * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs w)) IN ZinC` [] THEN qed[] );; let transcendence_u_tail = prove(` !P A B u. FINITE P ==> distinct_minpolys P ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> !n. norm(vsum(0..n) (\i. u i / Cx(&(FACT i)))) <= (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))) * exp A * A pow (n + 1) / &(FACT n) `, intro THEN subgoal `!p z. p IN P ==> complex_root p z ==> norm (cexp z - vsum (0..n) (\i. z pow i / Cx(&(FACT i)))) <= exp A * A pow (n+1) / &(FACT n)` THENL [ intro THEN have `norm (cexp z - vsum (0..n) (\i. z pow i / Cx(&(FACT i)))) <= exp(abs(Re z)) * norm z pow (n+1) / &(FACT n)` [TAYLOR_CEXP] THEN have `norm(z:complex) <= A` [] THEN have `&0 <= norm(z:complex)` [NORM_POS_LE] THEN have `&0 <= norm(z:complex) pow (n+1)` [REAL_POW_LE] THEN have `norm(z:complex) pow (n+1) <= A pow (n+1)` [REAL_POW_LE2] THEN real_linear_fact `&0 <= &(FACT n):real` THEN have `&0 <= inv(&(FACT n):real)` [REAL_LE_INV] THEN have `norm(z:complex) pow (n+1) * inv(&(FACT n)) <= A pow (n+1) * inv(&(FACT n))` [REAL_LE_RMUL] THEN have `norm(z:complex) pow (n+1) / &(FACT n) <= A pow (n+1) / &(FACT n)` [real_div] THEN have `&0 <= norm(z:complex) pow (n+1) * inv(&(FACT n))` [REAL_LE_MUL] THEN have `&0 <= norm(z:complex) pow (n+1) / &(FACT n)` [real_div] THEN have `abs(Re z) <= A` [COMPLEX_NORM_GE_RE_IM;REAL_LE_TRANS] THEN have `exp(abs(Re z)) <= exp A` [REAL_EXP_MONO_LE] THEN have `&0 <= exp(abs(Re z))` [REAL_EXP_POS_LE] THEN have `exp(abs(Re z)) * norm z pow (n+1) / &(FACT n) <= exp A * A pow (n+1) / &(FACT n)` [REAL_LE_MUL2] THEN qed[REAL_LE_TRANS] ; pass ] THEN subgoal `!p. p IN P ==> vsum (0..n) (\i. complex_root_powersums p i / Cx(&(FACT i))) = vsum (complex_root p) (\z. vsum (0..n) (\i. z pow i / Cx(&(FACT i))))` THENL [ intro THEN rw[complex_root_powersums] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `FINITE(0..n)` [FINITE_NUMSEG] THEN specialize[`\z i. z pow i / Cx(&(FACT i))`;`complex_root p`;`0..n`]VSUM_SWAP THEN have `vsum (complex_root p) (\z. vsum (0..n) (\i. z pow i / Cx (&(FACT i)))) = vsum (0..n) (\i. vsum (complex_root p) (\z. z pow i / Cx (&(FACT i))))` [] THEN simp[] THEN sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN simp[vsum_ring_sum_complex] THEN rw[complex_div;GSYM complex_ring_clauses] THEN sufficesby(GSYM RING_SUM_RMUL) THEN qed[in_complex_ring] ; pass ] THEN subgoal `!p. p IN P ==> norm(ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. complex_root_powersums p i / Cx(&(FACT i)))) <= &(poly_deg complex_ring p) * exp A * A pow (n+1) / &(FACT n)` THENL [ intro THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN simp[vsum_ring_sum_complex] THEN rw[GSYM ring_sub_complex] THEN simp[GSYM ring_sum_sub;in_complex_ring] THEN simp[GSYM vsum_ring_sum_complex] THEN specialize[`complex_root p`]VSUM_NORM_LE THEN have `!z. z IN complex_root p ==> norm(cexp z - vsum (0..n) (\i. z pow i / Cx (&(FACT i)))) <= exp A * A pow (n + 1) / &(FACT n)` [IN] THEN have `norm(vsum (complex_root p) (\z. cexp z - vsum (0..n) (\i. z pow i / Cx (&(FACT i))))) <= sum (complex_root p) (\z. exp A * A pow (n + 1) / &(FACT n))` [VSUM_NORM_LE] THEN specialize[`exp A * A pow (n + 1) / &(FACT n)`;`complex_root p`]SUM_CONST THEN have `CARD(complex_root p) = poly_deg complex_ring p` [distinct_minpolys_card_root] THEN rw[ring_sub_complex] THEN qed[] ; pass ] THEN subgoal `!p. p IN P ==> norm (complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int(B p) * complex_root_powersums p i) / Cx (&(FACT i)))) <= abs(real_of_int(B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n)` THENL [ intro THEN subgoal `vsum (0..n) (\i. (complex_of_int(B p) * complex_root_powersums p i) / Cx (&(FACT i))) = vsum (0..n) (\i. complex_of_int(B p) * (complex_root_powersums p i / Cx (&(FACT i))))` THENL [ sufficesby VSUM_EQ THEN rw[complex_div] THEN qed[COMPLEX_MUL_ASSOC] ; pass ] THEN subgoal `vsum (0..n) (\i. complex_of_int(B p) * (complex_root_powersums p i / Cx (&(FACT i)))) = complex_of_int(B p) * vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i))))` THENL [ specialize[`complex_of_int(B(p:(1->num)->complex))`;`\i. complex_root_powersums p i / Cx (&(FACT i))`]VSUM_COMPLEX_LMUL THEN qed[FINITE_NUMSEG] ; pass ] THEN have `vsum (0..n) (\i. (complex_of_int(B p) * complex_root_powersums p i) / Cx (&(FACT i))) = complex_of_int(B p) * vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i))))` [] THEN have `complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int(B p) * complex_root_powersums p i) / Cx (&(FACT i))) = complex_of_int(B p) * (ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i)))))` [COMPLEX_SUB_LDISTRIB] THEN have `norm(complex_of_int(B p) * (ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i)))))) = norm(complex_of_int(B p)) * norm(ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i)))))` [COMPLEX_NORM_MUL] THEN have `norm(complex_of_int(B(p:(1->num)->complex))) = abs(real_of_int(B p))` [complex_of_int;COMPLEX_NORM_CX] THEN have `norm(complex_of_int(B p) * (ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i)))))) = abs(real_of_int(B p)) * norm(ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i)))))` [] THEN have `abs(real_of_int(B p)) * norm(ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_root_powersums p i / Cx (&(FACT i))))) <= abs(real_of_int(B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n)` [REAL_ABS_POS;REAL_LE_LMUL] THEN qed[REAL_LE_TRANS] ; pass ] THEN subgoal `norm(vsum P (\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))))) <= sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n))` THENL [ specialize[`P:((1->num)->complex)->bool`;`\p. (complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))))`;`\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n)`]VSUM_NORM_LE THEN specialize[`\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp`;`\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i)))`;`P:((1->num)->complex)->bool`]VSUM_SUB THEN have `vsum P (\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i)))) = vsum P (\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp) - vsum P (\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))))` [] THEN have `vsum P (\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i)))) = -- vsum P (\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))))` [vsum_ring_sum_complex;COMPLEX_SUB_LZERO] THEN have `norm(vsum P (\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp - vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))))) = norm(vsum P (\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i)))))` [NORM_NEG] THEN qed[] ; pass ] THEN subgoal `vsum P (\p. vsum (0..n) (\i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i)))) = vsum(0..n) (\i. u i / Cx(&(FACT i)))` THENL [ have `FINITE (0..n)` [FINITE_NUMSEG] THEN specialize[`\p i. (complex_of_int (B p) * complex_root_powersums p i) / Cx (&(FACT i))`;`P:((1->num)->complex)->bool`;`0..n`]VSUM_SWAP THEN simp[] THEN sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN simp[vsum_ring_sum_complex] THEN rw[complex_div] THEN specialize_assuming[`complex_ring`;`\p. ring_mul complex_ring (complex_of_int (B p)) (complex_root_powersums p x)`;`inv(Cx(&(FACT x)))`;`P:((1->num)->complex)->bool`]RING_SUM_RMUL THEN simp[GSYM complex_ring_clauses] THEN qed[in_complex_ring] ; pass ] THEN have `norm(vsum(0..n) (\i. u i / Cx(&(FACT i)))) <= sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n))` [] THEN subgoal `sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p) * exp A * A pow (n + 1) / &(FACT n)) = sum P (\p:(1->num)->complex. (abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) / &(FACT n))` THENL [ rw[REAL_DIV] THEN sufficesby SUM_EQ THEN CONV_TAC REAL_FIELD ; pass ] THEN subgoal `sum P (\p:(1->num)->complex. (abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) / &(FACT n)) = (sum P (\p:(1->num)->complex. (abs (real_of_int (B p)) * &(poly_deg complex_ring p)))) * exp A * A pow (n + 1) / &(FACT n)` THENL [ specialize[`\p:(1->num)->complex. (abs (real_of_int (B p)) * &(poly_deg complex_ring p))`;`exp A * A pow (n + 1) / &(FACT n)`]SUM_RMUL THEN qed[] ; pass ] THEN qed[] );; let transcendence_v_bound = prove(` !P A B u. FINITE P ==> distinct_minpolys P ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> !n vn. vn = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) ==> norm vn <= (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))) * exp A * A pow (n + 1) `, intro THEN subgoal `ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i) = Cx(&(FACT n)) * vsum(0..n) (\i. u i / Cx(&(FACT i)))` THENL [ have `FINITE(0..n)` [FINITE_NUMSEG] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[GSYM VSUM_COMPLEX_LMUL] THEN sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN rw[complex_div] THEN simp[GSYM VSUM_COMPLEX_RMUL] THEN simp[GSYM VSUM_COMPLEX_LMUL] THEN sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN have `x <= n:num` [IN_NUMSEG] THEN simp[GSYM REAL_OF_NUM_MUL] THEN simp[REAL_OF_NUM_BINOM] THEN have `~(FACT(n-x) = 0)` [FACT_NZ] THEN have `~(&(FACT(n-x)) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT(n-x))) = Cx(&0))` [CX_INJ] THEN have `~(FACT(x) = 0)` [FACT_NZ] THEN have `~(&(FACT(x)) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT(x))) = Cx(&0))` [CX_INJ] THEN rw[CX_MUL;CX_DIV] THEN rw[GSYM complex_div] THEN complex_field_fact `~(Cx(&(FACT x)) = Cx(&0)) ==> ~(Cx(&(FACT(n-x))) = Cx(&0)) ==> (Cx (&(FACT (n - x))) * Cx (&(FACT n)) / (Cx (&(FACT (n - x))) * Cx (&(FACT x)))) * complex_of_int (B x') * complex_root_powersums x' x = Cx (&(FACT n)) * (complex_of_int (B x') * complex_root_powersums x' x) / Cx (&(FACT x))` THEN qed[] ; pass ] THEN have `vn = Cx(&(FACT n)) * vsum(0..n) (\i. u i / Cx(&(FACT i)))` [] THEN have `norm(vn:complex) = norm(Cx(&(FACT n))) * norm(vsum(0..n) (\i. u i / Cx(&(FACT i))))` [COMPLEX_NORM_MUL] THEN have `norm(vn:complex) = abs(&(FACT n)) * norm(vsum(0..n) (\i. u i / Cx(&(FACT i))))` [COMPLEX_NORM_CX] THEN have `abs(&(FACT n)) = &(FACT n):real` [REAL_ABS_REFL;FACT_LT;LT_IMP_LE;REAL_OF_NUM_LE] THEN have `norm(vn:complex) = &(FACT n) * norm(vsum(0..n) (\i. u i / Cx(&(FACT i))))` [] THEN have `norm(vsum(0..n) (\i. u i / Cx(&(FACT i)))) <= (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))) * exp A * A pow (n + 1) / &(FACT n)` [transcendence_u_tail] THEN have `&(FACT n) * norm(vsum(0..n) (\i. u i / Cx(&(FACT i)))) <= &(FACT n) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))) * exp A * A pow (n + 1) / &(FACT n)` [REAL_LE_LMUL;FACT_LT;LT_IMP_LE;REAL_OF_NUM_LE] THEN have `norm(vn:complex) <= &(FACT n) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))) * exp A * A pow (n + 1) / &(FACT n)` [] THEN have `~(FACT(n) = 0)` [FACT_NZ] THEN have `~(&(FACT(n)) = &0:real)` [REAL_OF_NUM_EQ] THEN real_field_fact `~(&(FACT n) = &0:real) ==> &(FACT n) * sum P (\p. abs (real_of_int (B(p:(1->num)->complex))) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) / &(FACT n) = sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` THEN qed[] );; let transcendence_p_bound = prove(` !P p A. distinct_minpolys P ==> p IN P ==> (!z. complex_root p z ==> norm z <= A) ==> (!d. norm(coeff d (x_truncreverse QinC_ring (poly_deg complex_ring p) p)) <= A pow d * &(binom(poly_deg complex_ring p,d))) `, intro THEN have `monic_vanishing_at complex_ring (complex_root p) I = p` [distinct_minpolys_monic_vanishing_at] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `CARD(complex_root p) = poly_deg complex_ring p` [distinct_minpolys_card_root] THEN have `!z. z IN complex_root p ==> I z IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`I:complex->complex`;`complex_root p`]x_truncreverse_monic_vanishing_at THEN have `x_truncreverse complex_ring (CARD(complex_root p)) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [] THEN have `x_truncreverse complex_ring (poly_deg complex_ring p) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [] THEN have `x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [x_truncreverse_subring;subring_complex_QinC] THEN simp[] THEN have `!z. z IN complex_root p ==> ring_powerseries complex_ring (one_minus_constx complex_ring (I z))` [series_complex] THEN have `!z. z IN complex_root p ==> norm(I z) <= A` [IN;I_THM] THEN have `!z d. z IN complex_root p ==> norm(coeff d (one_minus_constx complex_ring (I z))) <= A pow d * &(binom(1,d))` [coeff_root_bound_one_minus_constx] THEN specialize[`\z. one_minus_constx complex_ring (I z)`;`A:real`;`\z:complex. 1:num`;`complex_root p`]coeff_root_bound_product THEN have `norm (coeff d (poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z)))) <= A pow d * &(binom (CARD (complex_root p),d))` [CARD_EQ_NSUM] THEN qed[] );; let transcendence_H_bound = prove(` !P A t H. FINITE P ==> distinct_minpolys P ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!d. norm(coeff d H) <= A pow d * &(binom(t,d))) `, intro THEN specialize[`P:((1->num)->complex)->bool`;`H:(1->num)->complex`;`t:num`]transcendence_H_product_complex_ring THEN have `!p d. p IN P ==> norm(coeff d (x_truncreverse QinC_ring (poly_deg complex_ring p) p)) <= A pow d * &(binom(poly_deg complex_ring p,d))` [transcendence_p_bound] THEN specialize[`\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`A:real`;`\p:(1->num)->complex. poly_deg complex_ring p`;`P:((1->num)->complex)->bool`]coeff_root_bound_product THEN qed[] );; let transcendence_Hk_bound = prove(` !P A t H k. FINITE P ==> distinct_minpolys P ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!d. norm(coeff d (poly_pow complex_ring H k)) <= A pow d * &(binom(k*t,d))) `, intro THEN have `!d. norm(coeff d H:complex) <= (A:real) pow d * &(binom(t,d))` [transcendence_H_bound] THEN qed[coeff_root_bound_pow;series_complex;MULT_SYM] );; let transcendence_A_nonnegative = prove(` !P A. FINITE P ==> distinct_minpolys P ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> (P = {} \/ &0 <= A) `, intro THEN proven_if `P = {}:((1->num)->complex)->bool` [] THEN set_fact `~(P = {}) ==> ?p:(1->num)->complex. p IN P` THEN choose `p:(1->num)->complex` `(p:(1->num)->complex) IN P` [] THEN have `~(poly_deg complex_ring (p:(1->num)->complex) = 0)` [distinct_minpolys_deg_nonzero] THEN have `~(CARD(complex_root p) = 0)` [distinct_minpolys_card_root] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `~(complex_root p = {})` [CARD_EQ_0] THEN set_fact `~(complex_root p = {}) ==> ?z:complex. z IN complex_root p` THEN choose `z:complex` `z IN complex_root p` [] THEN have `complex_root p z` [IN] THEN have `norm(z:complex) <= A` [] THEN qed[NORM_POS_LE;REAL_LE_TRANS] );; let transcendence_Hv_bound = prove(` !P A B t H u v. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n k. norm ( coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) ) <= &2 pow (k*t) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) ) `, intro THEN rw[coeff_poly_mul_oneindex;complex_ring_clauses;coeff_series_from_coeffs] THEN specialize[`P:((1->num)->complex)->bool`;`A:real`;`t:num`]transcendence_Hk_bound THEN have `!d. norm(coeff d (poly_pow complex_ring H k)) <= A pow d * &(binom(k*t,d))` [] THEN specialize[`P:((1->num)->complex)->bool`;`A:real`;`B:((1->num)->complex)->int`;`u:num->complex`]transcendence_v_bound THEN have `!d. norm(v(n-d):complex) <= sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow ((n-d) + 1)` [] THEN have `!d. d IN (0..n) ==> norm(coeff d (poly_pow complex_ring H k) * v(n-d)) <= (A pow d * &(binom(k*t,d))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow ((n-d) + 1)` [COMPLEX_NORM_MUL;REAL_LE_MUL2;NORM_POS_LE] THEN have `FINITE (0..n)` [FINITE_NUMSEG] THEN specialize[`0..n`;`\d. coeff d (poly_pow complex_ring H k) * v(n-d)`]VSUM_NORM_LE THEN have `norm(vsum(0..n) (\d. coeff d (poly_pow complex_ring H k) * v(n-d))) <= sum(0..n) (\d. (A pow d * &(binom(k*t,d))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow ((n-d) + 1))` [] THEN subgoal `sum (0..n) (\d. (A pow d * &(binom (k * t,d))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n - d + 1)) = sum (0..n) (\d. &(binom (k * t,d)) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1))` THENL [ sufficesby SUM_EQ THEN intro THEN rw[BETA_THM] THEN have `x <= n:num` [IN_NUMSEG] THEN num_linear_fact `x <= n:num ==> n+1 = x + (n-x+1)` THEN have `(A:real) pow (n+1) = A pow x * A pow (n-x+1)` [REAL_POW_ADD] THEN real_linear_fact `(A pow x * &(binom (k * t,x))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n - x + 1) = &(binom (k * t,x)) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * (A pow x * A pow (n-x+1))` THEN qed[] ; pass ] THEN have `norm(vsum(0..n) (\d. coeff d (poly_pow complex_ring H k) * v(n-d))) <= sum(0..n) (\d. &(binom(k*t,d)) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1))` [] THEN specialize[`\d. &(binom(k*t,d))`;`sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)`]SUM_RMUL THEN have `sum(0..n) (\d. &(binom(k*t,d)) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)) = sum(0..n) (\d. &(binom(k*t,d))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` [] THEN subgoal `sum(0..n) (\d. &(binom(k*t,d))) <= (&2:real) pow (k*t)` THENL [ simp[GSYM REAL_OF_NUM_SUM] THEN specialize[`k * nsum P (\p:(1->num)->complex. poly_deg complex_ring p)`;`n:num`]binom_rowsum_partial THEN qed[REAL_OF_NUM_POW;REAL_OF_NUM_LE] ; pass ] THEN subgoal `&0:real <= sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` THENL [ have `!p:(1->num)->complex. p IN P ==> &0 <= abs(real_of_int(B p))` [REAL_ABS_POS] THEN have `!p:(1->num)->complex. p IN P ==> &0:real <= &(poly_deg complex_ring p)` [REAL_OF_NUM_LE;ARITH_RULE `0 <= x`] THEN have `!p:(1->num)->complex. p IN P ==> &0 <= abs (real_of_int (B p)) * &(poly_deg complex_ring p)` [REAL_LE_MUL] THEN have `&0 <= sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))` [SUM_POS_LE] THEN have `&0 <= exp A` [REAL_EXP_POS_LE] THEN case `&0:real <= A` THENL [ have `&0:real <= A pow (n+1)` [REAL_POW_LE] THEN qed[REAL_LE_MUL] ; specialize[`P:((1->num)->complex)->bool`;`A:real`]transcendence_A_nonnegative THEN have `P = {}:((1->num)->complex)->bool` [] THEN simp[SUM_CLAUSES] THEN CONV_TAC REAL_FIELD ] ; pass ] THEN have `sum (0..n) (\d. &(binom (k * t,d))) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) <= &2 pow (k * t) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` [REAL_LE_RMUL] THEN have `norm (ring_sum complex_ring (0..n) (\a. coeff a (poly_pow complex_ring H k) * v (n - a))) = norm (vsum (0..n) (\a. coeff a (poly_pow complex_ring H k) * v (n - a)))` [vsum_ring_sum_complex] THEN qed[REAL_LE_TRANS] );; let transcendence_Hv_zero = prove(` !P A D B t H u v. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> D IN ZinC ==> ~(D = Cx(&0)) ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n k. k*t <= n ==> norm(D pow (n-k) / Cx(&(FACT k))) * &2 pow (k * t) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) < &1 ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) ) `, intro THEN specialize[ `P:((1->num)->complex)->bool`; `D:complex`; `B:((1->num)->complex)->int` ]transcendence_Hv_denom THEN have `D pow (n-k) / Cx(&(FACT k)) * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) IN ZinC` [] THEN specialize[ `P:((1->num)->complex)->bool`; `A:real`; `B:((1->num)->complex)->int` ]transcendence_Hv_bound THEN have `norm(coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))) <= &2 pow (k * t) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` [] THEN subgoal `~(D pow (n-k) / Cx(&(FACT k)) = Cx(&0))` THENL [ have `~(D pow (n-k) = Cx(&0))` [COMPLEX_POW_EQ_0] THEN have `~(FACT k = 0)` [FACT_NZ] THEN have `~(&(FACT k) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT k)) = Cx(&0))` [CX_INJ] THEN qed[COMPLEX_DIV_EQ_0] ; pass ] THEN qed[zero_if_ZinC_scale_bound] );; let transcendence_Hv_zero_Ptrivial = prove(` !P H u v. P = {} ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> (!n k. coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) ) `, intro THEN have `!n:num. v n = Cx(&0)` [transcendence_uv_trivial] THEN have `!n:num. v n = ring_0 complex_ring` [complex_ring_clauses] THEN rw[coeff_poly_mul_oneindex;coeff_series_from_coeffs] THEN rw[GSYM complex_ring_clauses] THEN sufficesby RING_SUM_EQ_0 THEN intro THEN rw[BETA_THM] THEN qed[RING_MUL_RZERO;in_complex_ring] );; (* XXX: can improve bounds in proof *) (* in particular: *) (* first 10t is a proxy for 10t-1 *) (* and 10t+1 is a proxy for 10t+1/k *) (* and the first max(1,...) is a proxy for that^(1/k) *) (* and 10 is a proxy for something like 1+1/k *) let transcendence_Hv_zero_k0 = prove(` !P A D B t H u v k0. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> D IN ZinC ==> ~(D = Cx(&0)) ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> 1 <= k0 ==> norm(D) pow (10*t) * &2 pow t * max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) < &k0 ==> (!n k. k0 <= k ==> k*t <= n ==> n <= 10*k*t ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) ) `, intro THEN proven_if `P = {}:((1->num)->complex)->bool` [transcendence_Hv_zero_Ptrivial] THEN case `t = 0` THENL [ have `P = {}:((1->num)->complex)->bool` [distinct_minpolys_total_deg] THEN qed[] ; pass ] THEN subgoal `norm (D pow (n - k) / Cx (&(FACT k))) * &2 pow (k * t) * sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1) < &1` THENL [ have `&k0 <= &k:real` [REAL_OF_NUM_LE] THEN have `norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) < &k` [REAL_LTE_TRANS] THEN real_linear_fact `norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) = (norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)) * exp(&1)` THEN have `(norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)) * exp(&1) < &k` [REAL_LTE_TRANS] THEN have `&0 < exp(&1)` [REAL_EXP_POS_LT] THEN have `norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) < &k / exp(&1)` [REAL_LT_RDIV_EQ] THEN subgoal `&0 <= norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)` THENL [ have `&0 <= norm(D:complex)` [NORM_POS_LE] THEN have `&0 <= norm(D:complex) pow (10*t)` [REAL_POW_LE] THEN real_linear_fact `&0 <= &2:real` THEN have `&0 <= &2 pow t` [REAL_POW_LE] THEN real_linear_fact `&0 <= &1:real` THEN have `&0 <= max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A)` [REAL_LE_MAX] THEN have `&0 <= max (&1) A` [REAL_LE_MAX] THEN have `&0 <= (max (&1) A) pow (10*t+1)` [REAL_POW_LE] THEN qed[REAL_LE_MUL] ; pass ] THEN num_linear_fact `1 <= k0 ==> k0 <= k ==> ~(k = 0)` THEN have `(norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)) pow k < (&k / exp (&1)) pow k` [REAL_POW_LT2] THEN have `(&k / exp(&1)) pow k <= &(FACT k)` [factorial_lower_bound] THEN have `(norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)) pow k < &(FACT k)` [REAL_LTE_TRANS] THEN havetac `(norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1)) pow k = norm(D:complex) pow ((10 * t) * k) * &2 pow (t * k) * (max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A)) pow k * (max (&1) A) pow ((10*t+1)*k)` (simp[REAL_POW_MUL;REAL_POW_POW]) THEN have `norm(D:complex) pow ((10*t) * k) * &2 pow (k * t) * (max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A)) pow k * (max (&1) A) pow ((10*t+1)*k) < &(FACT k)` [MULT_SYM] THEN rw[COMPLEX_NORM_DIV;COMPLEX_NORM_CX;COMPLEX_NORM_POW] THEN subgoal `norm(D:complex) pow (n - k) <= norm D pow ((10*t) * k)` THENL [ have `~(norm(D:complex) < &1)` [zero_if_ZinC_norm_lt1] THEN real_linear_fact `~(norm(D:complex) < &1) ==> &1 <= norm D` THEN num_linear_fact `n <= 10*k*t ==> n-k <= (10*t)*k` THEN qed[REAL_POW_MONO] ; pass ] THEN subgoal `sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A <= max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow k` THENL [ have `sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A <= max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A)` [REAL_LE_MAX;REAL_LE_REFL] THEN have `&1 <= max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A)` [REAL_LE_MAX;REAL_LE_REFL] THEN num_linear_fact `1 <= k0 ==> k0 <= k ==> 1 <= k` THEN have `max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow 1 <= max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow k` [REAL_POW_MONO] THEN qed[REAL_POW_1;REAL_LE_TRANS] ; pass ] THEN subgoal `(A:real) pow (n+1) <= (max (&1) A) pow ((10*t+1)*k)` THENL [ have `&0 <= A:real` [transcendence_A_nonnegative] THEN have `A:real <= max (&1) A` [REAL_LE_MAX;REAL_LE_REFL] THEN have `(A:real) pow (n+1) <= (max (&1) A) pow (n+1)` [REAL_POW_LE2] THEN have `&1:real <= max (&1) A` [REAL_LE_MAX;REAL_LE_REFL] THEN num_linear_fact `1 <= k0 ==> k0 <= k ==> 1 <= k` THEN num_linear_fact `n <= 10*k*t ==> 1 <= k ==> n+1 <= (10*t+1)*k` THEN have `(max (&1:real) A) pow (n+1) <= (max (&1:real) A) pow ((10*t+1)*k)` [REAL_POW_MONO] THEN qed[REAL_LE_TRANS] ; pass ] THEN subgoal `norm(D:complex) pow (n - k) * &2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1) <= norm D pow ((10 * t) * k) * &2 pow (k * t) * max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow k * max (&1) A pow ((10 * t + 1) * k)` THENL [ have `&0 <= norm(D:complex)` [NORM_POS_LE] THEN have `&0 <= norm(D:complex) pow (n-k)` [REAL_POW_LE] THEN real_linear_fact `&0 <= &2:real` THEN have `&0 <= &2 pow (k*t)` [REAL_POW_LE] THEN real_linear_fact `&0 <= &1:real` THEN have `!p:(1->num)->complex. p IN P ==> &0 <= abs(real_of_int(B p))` [REAL_ABS_POS] THEN have `!p:(1->num)->complex. p IN P ==> &0:real <= &(poly_deg complex_ring p)` [REAL_OF_NUM_LE;ARITH_RULE `0 <= x`] THEN have `!p:(1->num)->complex. p IN P ==> &0 <= abs (real_of_int (B p)) * &(poly_deg complex_ring p)` [REAL_LE_MUL] THEN have `&0 <= sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p))` [SUM_POS_LE] THEN have `&0 <= exp A` [REAL_EXP_POS_LE] THEN specialize[`P:((1->num)->complex)->bool`;`A:real`]transcendence_A_nonnegative THEN have `&0 <= A:real` [] THEN have `&0 <= (A:real) pow (n+1)` [REAL_POW_LE] THEN have `(sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1) <= max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow k * max (&1) A pow ((10 * t + 1) * k)` [REAL_LE_MUL2;REAL_LE_MUL] THEN have `&2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1) <= &2 pow (k * t) * max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) pow k * max (&1) A pow ((10 * t + 1) * k)` [REAL_LE_MUL2;REAL_LE_MUL;REAL_LE_REFL] THEN have `&0 <= &2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1)` [REAL_LE_MUL] THEN qed[REAL_LE_MUL2] ; pass ] THEN have `norm(D:complex) pow (n - k) * &2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1) < &(FACT k)` [REAL_LET_TRANS] THEN have `&0 < &(FACT k)` [REAL_OF_NUM_LT;FACT_LT] THEN have `&0 < inv(&(FACT k))` [REAL_LT_INV] THEN have `(norm(D:complex) pow (n - k) * &2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1)) * inv(&(FACT k)) < &(FACT k) * inv(&(FACT k))` [REAL_LT_RMUL] THEN real_field_fact `&0 < &(FACT k) ==> (norm(D:complex) pow (n - k) * &2 pow (k * t) * (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * A pow (n + 1)) * inv(&(FACT k)) = norm D pow (n - k) / abs (&(FACT k)) * &2 pow (k * t) * sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A * A pow (n + 1)` THEN real_field_fact `&0 < &(FACT k) ==> &(FACT k) * inv(&(FACT k)) = &1` THEN qed[] ; pass ] THEN specialize[ `P:((1->num)->complex)->bool`; `A:real`; `D:complex`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex`; ]transcendence_Hv_zero THEN qed[] );; let transcendence_Hv_induction = prove(` !P A D B t H u v k0. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> D IN ZinC ==> ~(D = Cx(&0)) ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> 1 <= k0 ==> norm(D) pow (10*t) * &2 pow t * max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) < &k0 ==> (!j n k. k0 <= k ==> k*t <= n ==> n <= 10*k*t+j ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN INDUCT_TAC THENL [ rw[ARITH_RULE `10*k*t+0 = 10*k*t`] THEN specialize[ `P:((1->num)->complex)->bool`; `A:real`; `D:complex`; `B:((1->num)->complex)->int` ]transcendence_Hv_zero_k0 THEN qed[] ; pass ] THEN intro THEN proven_if `P = {}:((1->num)->complex)->bool` [transcendence_Hv_zero_Ptrivial] THEN case `t = 0` THENL [ have `P = {}:((1->num)->complex)->bool` [distinct_minpolys_total_deg] THEN qed[] ; pass ] THEN proven_if `n <= 10*k*t + j` [] THEN num_linear_fact `n <= 10 * k * t + SUC j ==> ~(n <= 10 * k * t + j) ==> n = 10*k*t+j+1` THEN subgoal `coeff n (poly_mul complex_ring (poly_pow complex_ring H (1+k)) (series_from_coeffs v)) = Cx (&0)` THENL [ num_linear_fact `k0 <= k ==> k0 <= 1+k` THEN num_linear_fact `1 <= k0 ==> k0 <= k ==> 1+k <= 10*k` THEN have `(1+k)*t <= (10*k)*t` [LE_MULT2;LE_REFL] THEN num_linear_fact `(1+k)*t <= (10*k)*t ==> (1+k)*t <= 10*k*t+j+1` THEN have `(1+k)*t <= n` [] THEN num_linear_fact `~(t = 0) ==> 10*k*t+j+1 <= 10*(1+k)*t+j` THEN have `n <= 10*(1+k)*t+j` [] THEN qed[] ; pass ] THEN subgoal `poly_mul complex_ring (poly_pow complex_ring H (1 + k)) (series_from_coeffs v) = poly_mul complex_ring H (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))` THENL [ simp[poly_pow_add;poly_pow_1;series_complex] THEN simp[POLY_MUL_ASSOC;series_complex] ; pass ] THEN have `coeff n (poly_mul complex_ring H (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))) = Cx(&0)` [] THEN specialize[`complex_ring`;`n:num`;`H:(1->num)->complex`;`poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)`]coeff_poly_mul_oneindex THEN have `ring_sum complex_ring (0..n) (\i. ring_mul complex_ring (coeff i H) (coeff (n-i) (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)))) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring (0..n) (\i. ring_mul complex_ring (coeff i H) (coeff (n-i) (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)))) = ring_sum complex_ring (0..n) (\i. if i = 0 then (coeff 0 H) * (coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))) else ring_0 complex_ring)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN proven_if `a = 0` [ARITH_RULE `n-0 = n`;complex_ring_clauses] THEN case `a <= t:num` THENL [ num_linear_fact `~(a = 0) ==> (10*k*t+j+1)-a <= 10*k*t+j` THEN have `n-a <= 10*k*t+j` [] THEN num_linear_fact `1 <= k0 ==> k0 <= k ==> k <= 10*k-1` THEN have `k*t <= (10*k)*t-1*t` [LE_MULT2;LE_REFL;RIGHT_SUB_DISTRIB] THEN num_linear_fact `a <= t ==> k*t <= (10*k)*t-1*t ==> k*t <= (10*k*t+j+1)-a` THEN have `k*t <= n-a:num` [] THEN have `coeff (n-a) (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) = ring_0 complex_ring` [complex_ring_clauses] THEN qed[RING_MUL_RZERO;in_complex_ring] ; pass ] THEN have `H = x_truncreverse complex_ring t (poly_product QinC_ring P I)` [x_truncreverse_subring;subring_complex_QinC] THEN have `poly_deg complex_ring (H:(1->num)->complex) <= t` [deg_x_truncreverse_le;series_complex] THEN have `ring_polynomial complex_ring (H:(1->num)->complex)` [transcendence_H_poly] THEN have `coeff a H = ring_0 complex_ring` [coeff_deg_le] THEN qed[RING_MUL_LZERO;in_complex_ring] ; pass ] THEN subgoal `ring_sum complex_ring (0..n) (\i. if i = 0 then coeff 0 H * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) else ring_0 complex_ring) = coeff 0 H * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))` THENL [ have `0 IN (0..n)` [IN_NUMSEG;ARITH_RULE `0 <= n`] THEN specialize[`complex_ring`;`0..n`;`0`;`coeff 0 H * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v))`]RING_SUM_DELTA THEN qed[in_complex_ring] ; pass ] THEN have `coeff 0 H * coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) = Cx(&0)` [] THEN specialize[ `P:((1->num)->complex)->bool`; `H:(1->num)->complex` ]transcendence_H_botcoeff1 THEN have `coeff 0 (H:(1->num)->complex) = Cx(&1)` [] THEN qed[COMPLEX_MUL_LID] );; let transcendence_Hv_noupperbound = prove(` !P A D B t H u v k0. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> D IN ZinC ==> ~(D = Cx(&0)) ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> 1 <= k0 ==> norm(D) pow (10*t) * &2 pow t * max (&1) (sum P (\p. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) < &k0 ==> (!n k. k0 <= k ==> k*t <= n ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) ) `, intro THEN num_linear_fact `n <= 10*k*t+n` THEN specialize[ `P:((1->num)->complex)->bool`; `A:real`; `D:complex`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex` ]transcendence_Hv_induction THEN qed[] );; let transcendence_Hv_kexists = prove(` !P A D B t H u v. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p z. p IN P ==> complex_root p z ==> norm z <= A) ==> D IN ZinC ==> ~(D = Cx(&0)) ==> (!p i. p IN P ==> D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> ?k. 1 <= k /\ !n. k*t <= n ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) `, intro THEN choose `k0:num` `max (&1) (norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1)) < &k0` [REAL_ARCH_LT] THEN have `&1 <= max (&1) (norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1))` [REAL_LE_MAX;REAL_LE_REFL] THEN have `&1 <= &k0:real` [REAL_LET_TRANS;REAL_LT_IMP_LE] THEN have `1 <= k0` [REAL_OF_NUM_LE] THEN have `norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) <= max (&1) (norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1))` [REAL_LE_MAX;REAL_LE_REFL] THEN have `norm(D:complex) pow (10*t) * &2 pow t * max (&1) (sum P (\p:(1->num)->complex. abs (real_of_int (B p)) * &(poly_deg complex_ring p)) * exp A) * (max (&1) A) pow (10*t+1) * exp(&1) < &k0:real` [REAL_LET_TRANS] THEN specialize[ `P:((1->num)->complex)->bool`; `A:real`; `D:complex`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex`; `k0:num` ]transcendence_Hv_noupperbound THEN witness `k0:num` THEN qed[LE_REFL] );; (* could use sup in proof but sum seems a bit easier *) let transcendence_Hv_kexists_v2 = prove(` !P B t H u v. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> ?k. 1 <= k /\ !n. k*t <= n ==> coeff n ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) = Cx(&0) `, intro THEN proven_if `P = {}:((1->num)->complex)->bool` [transcendence_Hv_zero_Ptrivial;ARITH_RULE `1 <= 1`] THEN def `A:real` `sum P (\p:(1->num)->complex. sum (complex_root p) norm)` THEN subgoal `!p z. p IN P ==> complex_root p z ==> norm z <= sum (complex_root p) norm` THENL [ intro THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `z IN complex_root p` [IN] THEN specialize[`norm:complex->real`;`complex_root p`;`z:complex`]SUM_DELETE THEN have `&0 <= sum(complex_root p DELETE z) norm` [SUM_POS_LE;NORM_POS_LE] THEN real_linear_fact `&0 <= sum (complex_root p DELETE z) norm ==> sum (complex_root p DELETE z) norm = sum (complex_root p) norm - norm z ==> norm z <= sum(complex_root p) norm` THEN qed[] ; pass ] THEN subgoal `!p. p IN P ==> sum(complex_root p) norm <= A` THENL [ intro THEN specialize[`\p:(1->num)->complex. sum(complex_root p) norm`;`P:((1->num)->complex)->bool`;`p:(1->num)->complex`]SUM_DELETE THEN have `&0 <= sum(P DELETE p) (\p. sum(complex_root p) norm)` [SUM_POS_LE;NORM_POS_LE] THEN real_linear_fact `&0 <= sum(P DELETE p) (\p. sum(complex_root p) norm) ==> sum(P DELETE p) (\p. sum(complex_root p) norm) = sum P (\p. sum(complex_root p) norm) - sum(complex_root p) norm ==> sum(complex_root p) norm <= sum P (\p. sum(complex_root p) norm)` THEN qed[] ; pass ] THEN have `!p z. p IN P ==> complex_root p z ==> norm z <= A` [REAL_LE_TRANS] THEN choose `D:complex` `D IN ZinC /\ ~(D = Cx(&0)) /\ (!p. p IN P ==> !i. D pow i * (coeff (poly_deg complex_ring p - i) p) IN ZinC)` [distinct_minpolys_denominator] THEN specialize_assuming[ `P:((1->num)->complex)->bool`; `A:real`; `D:complex`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex` ]transcendence_Hv_kexists THEN qed[] );; let transcendence_Hv_poly = prove(` !P B t H u v. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> ?k. 1 <= k /\ ring_polynomial complex_ring ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs v ) ) `, intro THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex`; ]transcendence_Hv_kexists_v2 THEN choose `k:num` `1 <= k /\ !n. k*t <= n ==> coeff n (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs v)) = Cx(&0)` [] THEN witness `k:num` THEN qed[poly_if_coeff;series_complex;complex_ring_clauses] );; let transcendence_uv_diffeq = prove(` !u v. (!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)) ==> poly_sub complex_ring ( poly_sub complex_ring ( series_from_coeffs v ) ( poly_mul complex_ring ( x_pow complex_ring 1 ) ( series_from_coeffs v ) ) ) ( poly_mul complex_ring ( x_pow complex_ring 2 ) ( x_derivative complex_ring ( series_from_coeffs v ) ) ) = series_from_coeffs u `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_series_from_coeffs;coeff_poly_sub;complex_coeff_x_pow_times;coeff_x_derivative;complex_ring_clauses;ring_sub_complex;ring_of_num_complex;complex_of_num] THEN simp[] THEN case `d = 0` THENL [ simp[ARITH_RULE `0 < 1`;ARITH_RULE `0 < 2`] THEN rw[RING_SUM_CLAUSES_NUMSEG;in_complex_ring;binom;ARITH_RULE `0 - 0 = 0`;FACT;ARITH_RULE `1 * 1 = 1`] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN num_linear_fact `~(d = 0) ==> ~(d < 1)` THEN case `d = 1` THENL [ simp[ARITH_RULE `1 < 2`;ARITH_RULE `1 - 1 = 0`] THEN rw[RING_SUM_CLAUSES_NUMSEG;ARITH_RULE `1 = SUC 0`;in_complex_ring;ARITH_RULE `0 <= SUC 0`] THEN rw[ARITH_RULE `SUC 0 = 1`;ARITH_RULE `1 - 0 = 1`;ARITH_RULE `1 - 1 = 0`;ARITH_RULE `0 - 0 = 0`;FACT;binom;ARITH_RULE `0 + 1 = 1`;ARITH_RULE `x * 1 = x`;complex_ring_clauses;fact_1] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN num_linear_fact `~(d = 0) ==> ~(d = 1) ==> ~(d < 2)` THEN num_linear_fact `~(d < 2) ==> d - 2 + 1 = d - 1` THEN num_linear_fact `d - 1 <= d` THEN specialize[`complex_ring`;`d-1`;`d:num`]ring_sum_numseg_le_expand THEN simp[] THEN simp[GSYM vsum_ring_sum_complex;FINITE_NUMSEG;complex_ring_clauses] THEN simp[GSYM VSUM_COMPLEX_LMUL;FINITE_NUMSEG] THEN simp[GSYM VSUM_SUB;FINITE_NUMSEG] THEN subgoal `vsum (0..d) (\x. Cx (&(FACT (d - x) * binom (d,x))) * u x - (if x <= d - 1 then Cx (&(FACT (d - 1 - x) * binom (d - 1,x))) * u x else Cx (&0)) - Cx (&(d - 1)) * (if x <= d - 1 then Cx (&(FACT (d - 1 - x) * binom (d - 1,x))) * u x else Cx (&0))) = vsum (0..d) (\i. if i = d then u d else Cx(&0))` THENL [ sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN case `x = d:num` THENL [ num_linear_fact `~(d < 2) ==> ~(d <= d - 1)` THEN simp[ARITH_RULE `d - d = 0`;BINOM_REFL;FACT;ARITH_RULE `1 * 1 = 1`] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN have `x <= d:num` [IN_NUMSEG] THEN num_linear_fact `x <= d ==> ~(x = d) ==> x <= d-1` THEN complex_field_fact `Cx (&(FACT (d - x) * binom (d,x))) * u x - Cx (&(FACT (d - 1 - x) * binom (d - 1,x))) * u x - Cx (&(d - 1)) * Cx (&(FACT (d - 1 - x) * binom (d - 1,x))) * u x = (Cx (&(FACT (d - x) * binom (d,x))) - Cx (&(FACT (d - 1 - x) * binom (d - 1,x))) - Cx (&(d - 1)) * Cx (&(FACT (d - 1 - x) * binom (d - 1,x)))) * u x` THEN simp[GSYM CX_MUL;GSYM CX_SUB;REAL_OF_NUM_MUL] THEN num_linear_fact `x <= d ==> ~(x = d) ==> x < d:num` THEN simp[fact_binom_lemma_37_real] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN simp[vsum_delta_complex;FINITE_NUMSEG;IN_NUMSEG_0;ARITH_RULE `d <= d:num`] );; let ord_product_one_minus_constx_I = prove(` !S y yinv. FINITE S ==> y * yinv = Cx(&1) ==> ( ring_polynomial complex_ring (poly_product complex_ring S (\z. one_minus_constx complex_ring (I z))) /\ ~(poly_product complex_ring S (\z. one_minus_constx complex_ring (I z)) = poly_0 complex_ring) /\ poly_ord (poly_product complex_ring S (\z. one_minus_constx complex_ring (I z))) yinv = if y IN S then 1 else 0 ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN have `~(poly_product complex_ring S (\z. one_minus_constx complex_ring (I z)) = poly_0 complex_ring)` [nonzero_poly_product;integral_domain_complex;nonzero_one_minus_constx;in_complex_ring;integral_domain] THEN case `(y:complex) IN S` THENL [ specialize_assuming[`complex_ring`;`S:complex->bool`;`y:complex`;`\z. one_minus_constx complex_ring (I z)`]poly_product_delete THEN have `poly_product complex_ring (S) (\z. one_minus_constx complex_ring (I z)) = poly_mul complex_ring (one_minus_constx complex_ring y) (poly_product complex_ring ((S) DELETE y) (\z. one_minus_constx complex_ring (I z)))` [series_complex;IN;I_THM] THEN have `FINITE (S DELETE (y:complex))` [FINITE_DELETE] THEN subgoal `~(poly_eval complex_ring (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv = Cx(&0))` THENL [ simp[eval_poly_product;in_complex_ring;one_minus_constx_poly;eval_one_minus_constx;I_THM;complex_ring_clauses;ring_sub_complex] THEN intro THEN have `ring_product complex_ring (S DELETE y) (\a. Cx (&1) - a * yinv) = ring_0 complex_ring` [complex_ring_clauses] THEN choose `a:complex` `a IN S DELETE y /\ Cx (&1) - a * yinv = ring_0 complex_ring` [INTEGRAL_DOMAIN_PRODUCT_EQ_0;integral_domain_complex] THEN have `Cx(&1) - a*yinv = Cx(&0)` [complex_ring_clauses] THEN complex_field_fact `y*yinv = Cx(&1) ==> Cx(&1) - a*yinv = Cx(&0) ==> a = y` THEN qed[IN_DELETE] ; pass ] THEN have `~(complex_root (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv)` [complex_root] THEN have `ring_polynomial complex_ring (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z)))` [poly_product_poly;one_minus_constx_poly;in_complex_ring] THEN specialize[`poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))`;`yinv:complex`]poly_ord_unique_0 THEN have `poly_ord (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv = 0` [] THEN have `poly_ord (one_minus_constx complex_ring y) yinv = 1` [poly_ord_one_minus_constx] THEN have `~(one_minus_constx complex_ring y = poly_0 complex_ring)` [nonzero_one_minus_constx;field;field_complex;in_complex_ring] THEN have `ring_polynomial complex_ring (one_minus_constx complex_ring y)` [one_minus_constx_poly;in_complex_ring] THEN have `ring_polynomial complex_ring (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z)))` [poly_product_poly;FINITE_DELETE;one_minus_constx_poly;in_complex_ring] THEN have `~(poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z)) = poly_0(complex_ring))` [poly_product_ring_product_x_poly;x_poly_use] THEN specialize[`one_minus_constx complex_ring y`;`poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))`;`yinv:complex`]poly_ord_mul THEN have `poly_ord (poly_product complex_ring (S) (\z. one_minus_constx complex_ring (I z))) yinv = poly_ord (poly_mul complex_ring (one_minus_constx complex_ring y) (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z)))) yinv` [] THEN have `poly_ord (poly_product complex_ring (S) (\z. one_minus_constx complex_ring (I z))) yinv = poly_ord (one_minus_constx complex_ring y) yinv + poly_ord (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv` [] THEN have `poly_ord (poly_product complex_ring (S) (\z. one_minus_constx complex_ring (I z))) yinv = 1 + poly_ord (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv` [] THEN qed[ARITH_RULE `1 + 0 = 1`] ; pass ] THEN subgoal `~(poly_eval complex_ring (poly_product complex_ring S (\z. one_minus_constx complex_ring (I z))) yinv = Cx(&0))` THENL [ simp[eval_poly_product;in_complex_ring;one_minus_constx_poly;eval_one_minus_constx;I_THM;complex_ring_clauses;ring_sub_complex] THEN intro THEN have `ring_product complex_ring (S) (\a. Cx (&1) - a * yinv) = ring_0 complex_ring` [complex_ring_clauses] THEN choose `a:complex` `a IN S /\ Cx (&1) - a * yinv = ring_0 complex_ring` [INTEGRAL_DOMAIN_PRODUCT_EQ_0;integral_domain_complex] THEN have `Cx(&1) - a*yinv = Cx(&0)` [complex_ring_clauses] THEN complex_field_fact `y*yinv = Cx(&1) ==> Cx(&1) - a*yinv = Cx(&0) ==> a = y` THEN qed[] ; pass ] THEN have `~(complex_root (poly_product complex_ring S (\z. one_minus_constx complex_ring (I z))) yinv)` [complex_root] THEN have `ring_polynomial complex_ring (poly_product complex_ring S (\z. one_minus_constx complex_ring (I z)))` [poly_product_poly;one_minus_constx_poly;I_THM;in_complex_ring] THEN qed[poly_ord_unique_0] );; let ord_sum_product_one_minus_constx_I = prove(` !S y yinv. FINITE S ==> y * yinv = Cx(&1) ==> y IN S ==> ( ring_polynomial complex_ring ( poly_sum complex_ring S (\t. poly_product complex_ring (S DELETE t) (\z. one_minus_constx complex_ring (I z) ) ) ) /\ ~( poly_sum complex_ring S (\t. poly_product complex_ring (S DELETE t) (\z. one_minus_constx complex_ring (I z) ) ) = poly_0 complex_ring ) /\ poly_ord ( poly_sum complex_ring S (\t. poly_product complex_ring (S DELETE t) (\z. one_minus_constx complex_ring (I z) ) ) ) yinv = 0 ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN have `!s. s IN S DELETE(y:complex) ==> s IN S` [IN_DELETE] THEN have `FINITE(S DELETE(y:complex))` [FINITE_DELETE] THEN have `!s. s IN S ==> ring_polynomial complex_ring (poly_product complex_ring (S DELETE s) (\z. one_minus_constx complex_ring (I z)))` [poly_product_poly;one_minus_constx_poly;in_complex_ring;FINITE_DELETE] THEN subgoal `!s. s IN S ==> ~(s = y) ==> poly_product complex_ring (S DELETE s) (\z. one_minus_constx complex_ring (I z)) = poly_0 complex_ring \/ 1 <= poly_ord (poly_product complex_ring (S DELETE s) (\z. one_minus_constx complex_ring (I z))) yinv` THENL [ intro THEN have `y IN S DELETE(s:complex)` [IN_DELETE] THEN specialize[ `S DELETE(s:complex)` ]ord_product_one_minus_constx_I THEN qed[ARITH_RULE `1 <= 1`;FINITE_DELETE] ; pass ] THEN have `~(poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z)) = poly_0 complex_ring)` [nonzero_poly_product;integral_domain_complex;nonzero_one_minus_constx;integral_domain;in_complex_ring] THEN specialize[ `S DELETE(y:complex)`; `y:complex`; `yinv:complex` ]ord_product_one_minus_constx_I THEN have `~(y IN S DELETE(y:complex))` [IN_DELETE] THEN have `poly_ord (poly_product complex_ring (S DELETE y) (\z. one_minus_constx complex_ring (I z))) yinv < 1` [ARITH_RULE `0 < 1`] THEN specialize[ `\t. poly_product complex_ring (S DELETE t) (\z. one_minus_constx complex_ring (I z))`; `yinv:complex`; `1`; `S:complex->bool`; `y:complex` ]poly_ord_sum_dominant THEN qed[] );; let transcendence_ord_revp = prove(` !P p y yinv. FINITE P ==> distinct_minpolys P ==> p IN P ==> y * yinv = Cx(&1) ==> ( ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p) /\ ~(x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_0 complex_ring) /\ poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring p) p) yinv = if complex_root p y then 1 else 0 ) `, intro THEN have `monic_vanishing_at complex_ring (complex_root p) I = p` [distinct_minpolys_monic_vanishing_at] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN have `CARD(complex_root p) = poly_deg complex_ring p` [distinct_minpolys_card_root] THEN have `!z. z IN complex_root p ==> I z IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[`complex_ring`;`I:complex->complex`]x_truncreverse_monic_vanishing_at THEN have `x_truncreverse complex_ring (CARD(complex_root p)) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [] THEN have `x_truncreverse complex_ring (poly_deg complex_ring p) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [] THEN have `x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_product complex_ring (complex_root p) (\z. one_minus_constx complex_ring (I z))` [x_truncreverse_subring;subring_complex_QinC] THEN qed[ord_product_one_minus_constx_I;IN] );; let transcendence_ord_H = prove(` !P t H. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!q y yinv. q IN P ==> complex_root q y ==> y * yinv = Cx(&1) ==> poly_ord H yinv = 1 ) `, intro THEN have `poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring q) q) yinv = 1` [transcendence_ord_revp] THEN have `!p. p IN P ==> ~(p = q) ==> poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring p) p) yinv = 0` [transcendence_ord_revp;distinct_minpolys_distinct_roots] THEN subgoal `poly_ord H yinv = nsum P (\s. poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring s) s) yinv)` THENL [ specialize_assuming[`\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`yinv:complex`;`P:((1->num)->complex)->bool`]poly_ord_product THEN have `!p. p IN P ==> ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [x_truncreverse_poly;series_complex;x_truncreverse_subring;subring_complex_QinC] THEN have `!p. p IN P ==> ~(x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_0 complex_ring)` [distinct_minpolys_reverse_nonzero] THEN have `H = poly_product complex_ring P (\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [transcendence_H_product_complex_ring] THEN specialize[`\p. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`yinv:complex`;`P:((1->num)->complex)->bool`]poly_ord_product THEN qed[] ; pass ] THEN subgoal `nsum P (\s. poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring s) s) yinv) = nsum P (\s. if s = q then 1 else 0)` THENL [ sufficesby NSUM_EQ THEN intro THEN rw[BETA_THM] THEN qed[] ; pass ] THEN specialize[`1`;`P:((1->num)->complex)->bool`;`q:(1->num)->complex`](GEN `b:num` NSUM_DELTA) THEN qed[] );; let transcendence_ord_Hk = prove(` !P t H k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!q y yinv. q IN P ==> complex_root q y ==> y * yinv = Cx(&1) ==> poly_ord (poly_pow complex_ring H k) yinv = k ) `, intro THEN have `ring_polynomial complex_ring (H:(1->num)->complex)` [transcendence_H_poly] THEN have `coeff 0 (H:(1->num)->complex) = Cx(&1)` [transcendence_H_botcoeff1] THEN have `~(H = poly_0 complex_ring:(1->num)->complex)` [coeff_poly_0;complex_ring_clauses;ring_1_0_complex] THEN specialize[`H:(1->num)->complex`;`k:num`;`yinv:complex`]poly_ord_pow THEN have `poly_ord H yinv = 1` [transcendence_ord_H] THEN qed[ARITH_RULE `k * 1 = k`] );; let transcendence_ord_Hu = prove(` !P B t H u. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> (!q y yinv. q IN P ==> ~(B q = &0) ==> complex_root q y ==> y * yinv = Cx(&1) ==> ( ring_polynomial complex_ring (poly_mul complex_ring H (series_from_coeffs u)) /\ ~(poly_mul complex_ring H (series_from_coeffs u) = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring H (series_from_coeffs u)) yinv = 0 ) ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN def `G:((1->num)->complex)->((1->num)->complex)` `\p. poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q)` THEN have `!p:(1->num)->complex. p IN P ==> G p = poly_product complex_ring (P DELETE p) (\q. x_truncreverse QinC_ring (poly_deg complex_ring q) q)` [] THEN def `e:num` `0` THEN def `k:num` `1` THEN have `e < k:num` [ARITH_RULE `0 < 1`] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `G:((1->num)->complex)->(1->num)->complex`; `u:num->complex`; `e:num`; `k:num` ]transcendence_newton_Hk_psum THEN subgoal `poly_mul complex_ring H (series_from_coeffs u) = poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs (\n. Cx (&(FACT e * binom (n,e))) * u (n - e)))` THENL [ have `H = poly_pow complex_ring H k:(1->num)->complex` [poly_pow_1;series_complex] THEN subgoal `series_from_coeffs u = series_from_coeffs (\n. Cx (&(FACT e * binom (n,e))) * u (n - e))` THENL [ sufficesby eq_coeff THEN rw[coeff_series_from_coeffs] THEN simp[FACT;binom;ARITH_RULE `d - 0 = d`;ARITH_RULE `1 * 1 = 1`;COMPLEX_MUL_LID] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_polynomial complex_ring (poly_pow complex_ring H (k - (e + 1))) /\ ~(poly_pow complex_ring H (k - (e + 1)) = poly_0 complex_ring) /\ poly_ord(poly_pow complex_ring H (k - (e + 1))) yinv = 0` THENL [ have_rw `k - (e+1) = 0` [ARITH_RULE `1 - (0+1) = 0`] THEN rw[poly_pow_0;poly_ord_1] ; pass ] THEN def `Z:((1->num)->complex)->((1->num)->complex)` `\p. poly_mul complex_ring (poly_const complex_ring (complex_of_int (B p))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))` THEN subgoal `ring_polynomial complex_ring (Z q) /\ ~(Z q = poly_0 complex_ring) /\ poly_ord(Z(q:(1->num)->complex)) yinv = 0` THENL [ subgoal `ring_polynomial complex_ring (poly_const complex_ring (complex_of_int (B q)):(1->num)->complex) /\ ~(poly_const complex_ring (complex_of_int (B(q:(1->num)->complex))) = poly_0 complex_ring:(1->num)->complex) /\ poly_ord(poly_const complex_ring (complex_of_int (B(q:(1->num)->complex)))) yinv = 0` THENL [ have `~(complex_of_int (B(q:(1->num)->complex)) = Cx(&0))` [complex_of_int;CX_INJ;int_eq;REAL_OF_INT_CLAUSES] THEN qed[poly_ord_const] ; pass ] THEN subgoal `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root q) e) /\ ~(scaled_pow_newton_rightside complex_ring I (complex_root q) e = poly_0 complex_ring) /\ poly_ord(scaled_pow_newton_rightside complex_ring I (complex_root q) e) yinv = 0` THENL [ rw[scaled_pow_newton_rightside] THEN simp[FACT;RING_OF_NUM_1;ARITH_RULE `0+1 = 1`;const_1_x_pow] THEN simp[POLY_MUL_LID;poly_pow_1;series_complex] THEN have `y IN complex_root q` [IN] THEN have `FINITE(complex_root q)` [distinct_minpolys_finite_root] THEN specialize[`complex_root q`;`y:complex`;`yinv:complex`]ord_sum_product_one_minus_constx_I THEN qed[] ; pass ] THEN subgoal `ring_polynomial complex_ring (poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1)) /\ ~(poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1) = poly_0 complex_ring) /\ poly_ord(poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1)) yinv = 0` THENL [ have `e + 1 = 1` [ARITH_RULE `0 + 1 = 1`] THEN rw[know `e + 1 = 1`] THEN simp[poly_pow_1;series_complex] THEN have `FINITE (P DELETE (q:(1->num)->complex))` [FINITE_DELETE] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> p IN P` [IN_DELETE] THEN have `!p:(1->num)->complex. p IN P ==> ring_polynomial QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> monic QinC_ring p` [distinct_minpolys] THEN have `!p:(1->num)->complex. p IN P ==> poly_deg QinC_ring p = poly_deg complex_ring p` [poly_deg_subring;subring_complex_QinC] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> ring_polynomial QinC_ring (x_truncreverse QinC_ring (poly_deg QinC_ring p) p)` [x_truncreverse_poly;ring_polynomial] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> ring_polynomial QinC_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring p) p)` [poly_complex_if_poly_QinC] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> ~(x_truncreverse QinC_ring (poly_deg QinC_ring p) p = poly_0 QinC_ring)` [nonzero_if_x_truncreverse_monic;integral_domain;integral_domain_QinC] THEN have `!p:(1->num)->complex. p IN P DELETE q ==> ~(x_truncreverse QinC_ring (poly_deg complex_ring p) p = poly_0 complex_ring)` [poly_0_QinC_eq_poly_0_complex] THEN specialize[`\p:(1->num)->complex. x_truncreverse QinC_ring (poly_deg complex_ring p) p`;`yinv:complex`;`P DELETE (q:(1->num)->complex)`]poly_ord_product THEN simp[] THEN subgoal `nsum (P DELETE q) (\s. poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring s) s) yinv) = nsum (P DELETE q) (\s. 0)` THENL [ sufficesby NSUM_EQ THEN intro THEN rw[BETA_THM] THEN have `(x:(1->num)->complex) IN P` [IN_DELETE] THEN have `~((x:(1->num)->complex) = q)` [IN_DELETE] THEN have `~complex_root x y` [distinct_minpolys_distinct_roots] THEN qed[transcendence_ord_revp] ; pass ] THEN qed[NSUM_0] ; pass ] THEN specialize_assuming_nosplit[`poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1):(1->num)->complex`;`scaled_pow_newton_rightside complex_ring I (complex_root q) e`;`yinv:complex`]poly_ord_mul_0 THEN specialize_assuming_nosplit[`poly_pow complex_ring H (k - (e + 1)):(1->num)->complex`;`poly_mul complex_ring (poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1):(1->num)->complex) (scaled_pow_newton_rightside complex_ring I (complex_root q) e)`;`yinv:complex`]poly_ord_mul_0 THEN specialize_assuming_nosplit[`poly_const complex_ring (complex_of_int (B(q:(1->num)->complex))):(1->num)->complex`;`poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex) (poly_mul complex_ring (poly_pow complex_ring (G(q:(1->num)->complex)) (e + 1):(1->num)->complex) (scaled_pow_newton_rightside complex_ring I (complex_root q) e))`;`yinv:complex`]poly_ord_mul_0 THEN qed[] ; pass ] THEN subgoal `!p:(1->num)->complex. p IN P ==> ~(p = q) ==> (ring_polynomial complex_ring (Z p) /\ (Z p = poly_0 complex_ring \/ 1 <= poly_ord (Z p:(1->num)->complex) yinv))` THENL [ GEN_TAC THEN REPEAT DISCH_TAC THEN have `ring_polynomial complex_ring (poly_const complex_ring (complex_of_int (B(p:(1->num)->complex))):(1->num)->complex)` [RING_POLYNOMIAL_CONST;in_complex_ring] THEN subgoal `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` THENL [ rw[scaled_pow_newton_rightside] THEN simp[FACT;RING_OF_NUM_1;ARITH_RULE `0+1 = 1`;const_1_x_pow] THEN simp[POLY_MUL_LID;poly_pow_1;series_complex] THEN have `FINITE(complex_root p)` [distinct_minpolys_finite_root] THEN subgoal `!t. ring_polynomial complex_ring (poly_product complex_ring (complex_root p DELETE t) (\s. one_minus_constx complex_ring (I s)))` THENL [ intro THEN qed[poly_product_poly;FINITE_DELETE;one_minus_constx_poly;in_complex_ring] ; pass ] THEN qed[poly_sum_poly] ; pass ] THEN subgoal `ring_polynomial complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1)) /\ 1 <= poly_ord(poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1)) yinv` THENL [ have `e + 1 = 1` [ARITH_RULE `0 + 1 = 1`] THEN rw[know `e + 1 = 1`] THEN simp[poly_pow_1;series_complex] THEN have `FINITE(P DELETE(p:(1->num)->complex))` [FINITE_DELETE] THEN subgoal `!s. s IN P DELETE p ==> ring_polynomial complex_ring (x_truncreverse QinC_ring (poly_deg complex_ring s) s) /\ ~(x_truncreverse QinC_ring (poly_deg complex_ring s) s = poly_0 complex_ring)` THENL [ intro THEN specialize_assuming[`P:((1->num)->complex)->bool`;`s:(1->num)->complex`;`y:complex`;`yinv:complex`]transcendence_ord_revp THEN qed[IN_DELETE] ; pass ] THEN intro THENL [ specialize[`complex_ring`;`\s. x_truncreverse QinC_ring (poly_deg complex_ring s) s`]poly_product_poly THEN qed[FINITE_DELETE] ; simp[poly_ord_product] THEN have `q IN (P DELETE (p:(1->num)->complex))` [IN_DELETE] THEN specialize[`P:((1->num)->complex)->bool`;`q:(1->num)->complex`;`y:complex`;`yinv:complex`]transcendence_ord_revp THEN have `poly_ord (x_truncreverse QinC_ring (poly_deg complex_ring q) q) yinv = 1` [] THEN qed[term_le_nsum] ] ; pass ] THEN have `ring_polynomial complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex)` [] THEN have `poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1) = poly_0 complex_ring \/ 1 <= poly_ord (poly_pow complex_ring (G p) (e + 1)) yinv` [] THEN have `ring_polynomial complex_ring (scaled_pow_newton_rightside complex_ring I (complex_root p) e)` [] THEN specialize[ `poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex`; `scaled_pow_newton_rightside complex_ring I (complex_root p) e`; `yinv:complex`; `1` ]poly_ord_mul_ge_ge0 THEN have `ring_polynomial complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex)` [] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))` [] THEN have `poly_mul complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e) = poly_0 complex_ring \/ 1 <= poly_ord (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)) yinv` [] THEN specialize[ `poly_pow complex_ring H (k - (e + 1)):(1->num)->complex`; `poly_mul complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)`; `yinv:complex`; `1` ]poly_ord_mul_ge0_ge THEN have `ring_polynomial complex_ring (poly_const complex_ring (complex_of_int (B(p:(1->num)->complex))):(1->num)->complex)` [] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))` [] THEN have `poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)) = poly_0 complex_ring \/ 1 <= poly_ord (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))) yinv` [] THEN specialize[ `poly_const complex_ring (complex_of_int (B(p:(1->num)->complex):int)):(1->num)->complex`; `poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1)):(1->num)->complex) (poly_mul complex_ring (poly_pow complex_ring (G(p:(1->num)->complex)) (e + 1):(1->num)->complex) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))`; `yinv:complex`; `1` ]poly_ord_mul_ge0_ge THEN qed[] ; pass ] THEN subgoal `ring_polynomial complex_ring (poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring (complex_of_int (B p))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))) /\ ~(poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring (complex_of_int (B p))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e)))) = poly_0 complex_ring) /\ poly_ord(poly_sum complex_ring P (\p. poly_mul complex_ring (poly_const complex_ring (complex_of_int (B p))) (poly_mul complex_ring (poly_pow complex_ring H (k - (e + 1))) (poly_mul complex_ring (poly_pow complex_ring (G p) (e + 1)) (scaled_pow_newton_rightside complex_ring I (complex_root p) e))))) yinv = 0` THENL [ have `!s:(1->num)->complex. s IN P ==> ring_polynomial complex_ring (Z s:(1->num)->complex)` [] THEN have `!s:(1->num)->complex. s IN P ==> ~(s = q) ==> Z s = poly_0 complex_ring \/ 1 <= poly_ord (Z s) yinv` [] THEN have `~(Z(q:(1->num)->complex) = poly_0 complex_ring:(1->num)->complex)` [] THEN have `poly_ord (Z(q:(1->num)->complex)) yinv < 1` [ARITH_RULE `0 < 1`] THEN specialize[`Z:((1->num)->complex)->((1->num)->complex)`;`yinv:complex`;`1`;`P:((1->num)->complex)->bool`;`q:(1->num)->complex`]poly_ord_sum_dominant THEN qed[] ; pass ] THEN qed[] );; let transcendence_ord_Hku = prove(` !P B t H u k. FINITE P ==> distinct_minpolys P ==> t = nsum P (\p. poly_deg complex_ring p) ==> H = x_truncreverse QinC_ring t (poly_product QinC_ring P I) ==> (!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))) ==> 1 <= k ==> (!q y yinv. q IN P ==> ~(B q = &0) ==> complex_root q y ==> y * yinv = Cx(&1) ==> ( ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) /\ ~(poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u) = poly_0 complex_ring) /\ poly_ord ( poly_mul complex_ring ( poly_pow complex_ring H k ) ( series_from_coeffs u ) ) yinv = k-1 ) ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN num_linear_fact `1 <= k ==> k = (k-1)+1` THEN have `poly_pow complex_ring H k = poly_mul complex_ring (poly_pow complex_ring H (k-1)) (poly_pow complex_ring H 1):(1->num)->complex` [poly_pow_add;poly_pow_1;series_complex] THEN have `poly_pow complex_ring H k = poly_mul complex_ring (poly_pow complex_ring H (k-1)) H:(1->num)->complex` [poly_pow_1;series_complex] THEN subgoal `poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u) = poly_mul complex_ring (poly_pow complex_ring H (k-1)) (poly_mul complex_ring H (series_from_coeffs u))` THENL [ specialize_assuming[`complex_ring`;`poly_pow complex_ring H (k-1):(1->num)->complex`;`H:(1->num)->complex`;`series_from_coeffs u:(1->num)->complex`]POLY_MUL_ASSOC THEN qed[series_complex] ; pass ] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex` ]transcendence_ord_Hu THEN have `poly_ord (poly_pow complex_ring H (k-1)) yinv = k-1` [transcendence_ord_Hk] THEN have `ring_polynomial complex_ring (H:(1->num)->complex)` [transcendence_H_poly] THEN have `coeff 0 (H:(1->num)->complex) = Cx(&1)` [transcendence_H_botcoeff1] THEN have `~(H = poly_0 complex_ring:(1->num)->complex)` [coeff_poly_0;complex_ring_clauses;ring_1_0_complex] THEN have `ring_polynomial complex_ring (poly_pow complex_ring H (k - 1):(1->num)->complex)` [poly_pow_poly] THEN have `ring_polynomial complex_ring (poly_mul complex_ring H (series_from_coeffs u))` [] THEN have `~(poly_pow complex_ring H (k - 1) = poly_0 complex_ring:(1->num)->complex)` [nonzero_poly_pow;integral_domain_complex;series_complex] THEN have `~(poly_mul complex_ring H (series_from_coeffs u) = poly_0 complex_ring)` [] THEN specialize[`poly_pow complex_ring H (k-1):(1->num)->complex`;`poly_mul complex_ring H (series_from_coeffs u)`;`yinv:complex`]poly_ord_mul THEN have `poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) yinv = poly_ord (poly_pow complex_ring H (k - 1)) yinv + poly_ord (poly_mul complex_ring H (series_from_coeffs u)) yinv` [] THEN have `poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) yinv = (k-1) + poly_ord (poly_mul complex_ring H (series_from_coeffs u)) yinv` [] THEN have `poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) yinv = (k-1) + 0` [] THEN qed[ARITH_RULE `(k-1)+0 = k-1`] );; (* if rational functions U = e/g and V = f/g satisfy U = V - xV - x^2 V' then U cannot have a pole of order 1 at any nonzero point *) let transcendence_pole_order_ne1 = prove(` !U V e f g y yinv. y * yinv = Cx(&1) ==> ring_polynomial complex_ring e ==> ring_polynomial complex_ring f ==> ring_polynomial complex_ring g ==> ~(g = poly_0 complex_ring) ==> ~(e = poly_0 complex_ring) ==> poly_ord e yinv + 1 = poly_ord g yinv ==> poly_mul complex_ring g U = e ==> poly_mul complex_ring g V = f ==> poly_sub complex_ring ( poly_sub complex_ring V ( poly_mul complex_ring (x_pow complex_ring 1) V ) ) ( poly_mul complex_ring (x_pow complex_ring 2) (x_derivative complex_ring V) ) = U ==> F `, intro THEN recall series_complex THEN have `x_derivative complex_ring f = poly_add complex_ring (poly_mul complex_ring (x_derivative complex_ring g) V) (poly_mul complex_ring g (x_derivative complex_ring V))` [x_derivative_mul] THEN subgoal `poly_mul complex_ring e g = poly_sub complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))))` THENL [ subgoal `ring_mul (x_series complex_ring) e g = ring_sub (x_series complex_ring) (ring_mul (x_series complex_ring) (ring_sub (x_series complex_ring) (ring_1 (x_series complex_ring)) (x_pow complex_ring 1)) (ring_mul (x_series complex_ring) f g)) (ring_mul (x_series complex_ring) (x_pow complex_ring 2) (ring_sub (x_series complex_ring) (ring_mul (x_series complex_ring) (x_derivative complex_ring f) g) (ring_mul (x_series complex_ring) f (x_derivative complex_ring g))))` THENL [ have `U IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `V IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `x_derivative complex_ring V IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `e IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `f IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `x_derivative complex_ring f IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `g IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `x_derivative complex_ring g IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `x_pow complex_ring 1 IN ring_carrier(x_series complex_ring)` [x_series_use] THEN have `x_pow complex_ring 2 IN ring_carrier(x_series complex_ring)` [x_series_use] THEN specialize_assuming[ `x_series complex_ring`; `U:(1->num)->complex`; `V:(1->num)->complex`; `x_derivative complex_ring V:(1->num)->complex`; `e:(1->num)->complex`; `f:(1->num)->complex`; `x_derivative complex_ring f:(1->num)->complex`; `g:(1->num)->complex`; `x_derivative complex_ring g:(1->num)->complex`; `x_pow complex_ring 1`; `x_pow complex_ring 2` ](GENL[ `r:R ring`; `U:R`; `V:R`; `V':R`; `e:R`; `f:R`; `f':R`; `g:R`; `g':R`; `x1:R`; `x2:R` ](RING_RULE ` ring_mul(r:R ring) g U = e ==> ring_mul r g V = f ==> ring_sub r (ring_sub r V (ring_mul r x1 V)) (ring_mul r x2 V') = U ==> f' = ring_add r (ring_mul r g' V) (ring_mul r g V') ==> ring_mul r e g = ring_sub r (ring_mul r (ring_sub r (ring_1 r) x1) (ring_mul r f g)) (ring_mul r x2 (ring_sub r (ring_mul r f' g) (ring_mul r f g'))) `)) THEN have `ring_mul (x_series complex_ring) g U = e` [x_series_use] THEN have `ring_mul (x_series complex_ring) g V = f` [x_series_use] THEN have `ring_sub (x_series complex_ring) (ring_sub (x_series complex_ring) V (ring_mul (x_series complex_ring) (x_pow complex_ring 1) V)) (ring_mul (x_series complex_ring) (x_pow complex_ring 2) (x_derivative complex_ring V)) = U` [x_series_use;x_series_sub_use] THEN have `x_derivative complex_ring f = ring_add (x_series complex_ring) (ring_mul (x_series complex_ring) (x_derivative complex_ring g) V) (ring_mul (x_series complex_ring) g (x_derivative complex_ring V))` [x_series_use] THEN qed[] ; pass ] THEN simp[x_series_use;x_series_sub_use] ; pass ] THEN have `ring_polynomial complex_ring (poly_mul complex_ring e g) /\ ~(poly_mul complex_ring e g = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring e g) yinv = poly_ord e yinv + poly_ord g yinv` [poly_ord_mul] THEN have `poly_ord (poly_mul complex_ring e g) yinv + 1 = 2 * poly_ord g yinv` [ARITH_RULE `E+1 = G ==> M = E+G ==> M+1 = 2*G`] THEN case `f = poly_0 complex_ring \/ poly_ord g yinv <= poly_ord f yinv` THENL [ have `ring_polynomial complex_ring (poly_mul complex_ring f g) /\ (poly_mul complex_ring f g = poly_0 complex_ring \/ 2 * poly_ord g yinv <= poly_ord (poly_mul complex_ring f g) yinv)` [poly_ord_mul_ge_ge;LE_REFL;ARITH_RULE `a + a = 2*a`] THEN specialize_assuming[ `poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)`; `poly_mul complex_ring f g:(1->num)->complex`; `yinv:complex`; `2 * poly_ord g yinv` ]poly_ord_mul_ge0_ge THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1))` [RING_POLYNOMIAL_SUB;x_pow_poly;RING_POLYNOMIAL_1] THEN have `poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g) = poly_0 complex_ring \/ 2 * poly_ord g yinv <= poly_ord (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) yinv` [] THEN have `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)) = poly_0 complex_ring \/ 2 * poly_ord g yinv <= poly_ord (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) yinv` [poly_ord_derivative_ge] THEN have `ring_polynomial complex_ring (x_pow complex_ring 2)` [x_pow_poly] THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))` [RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_MUL;x_derivative_polynomial] THEN specialize[ `x_pow complex_ring 2`; `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))`; `yinv:complex`; `2 * poly_ord g yinv` ]poly_ord_mul_ge0_ge THEN subgoal `2 * poly_ord g yinv <= poly_ord (poly_mul complex_ring e g) yinv` THENL [ have `ring_polynomial complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g))` [] THEN have `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)) = poly_0 complex_ring \/ 2 * poly_ord g yinv <= poly_ord (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) yinv` [] THEN have `poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g) = poly_0 complex_ring \/ 2 * poly_ord g yinv <= poly_ord (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) yinv` [] THEN specialize_assuming[ `poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)`; `poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))`; `yinv:complex`; `2 * poly_ord g yinv` ]poly_ord_sub_ge THEN qed[] ; pass ] THEN num_linear_fact `2 * poly_ord g yinv <= poly_ord (poly_mul complex_ring e g) yinv ==> poly_ord (poly_mul complex_ring e g) yinv + 1 = 2 * poly_ord g yinv ==> poly_ord e yinv + 1 = poly_ord g yinv ==> F` THEN qed[] ; pass ] THEN have `poly_ord f yinv < poly_ord g yinv` [ARITH_RULE `~(a <= b) ==> b < a:num`] THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) /\ ~(poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)) = poly_0 complex_ring) /\ poly_ord (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) yinv = (poly_ord f yinv + poly_ord g yinv) - 1` [poly_ord_derivative_pole] THEN have `ring_polynomial complex_ring (poly_mul complex_ring f g) /\ ~(poly_mul complex_ring f g = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring f g) yinv = poly_ord f yinv + poly_ord g yinv` [poly_ord_mul] THEN subgoal `ring_polynomial complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) /\ (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g) = poly_0 complex_ring \/ poly_ord f yinv + poly_ord g yinv <= poly_ord (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) yinv)` THENL [ specialize_assuming[ `poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)`; `poly_mul complex_ring f g:(1->num)->complex`; `yinv:complex`; `poly_ord f yinv + poly_ord g yinv` ]poly_ord_mul_ge0_ge THEN have `ring_polynomial complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1))` [RING_POLYNOMIAL_SUB;x_pow_poly;RING_POLYNOMIAL_1] THEN qed[LE_REFL] ; pass ] THEN subgoal `ring_polynomial complex_ring (x_pow complex_ring 2) /\ ~(x_pow complex_ring 2 = poly_0 complex_ring) /\ poly_ord (x_pow complex_ring 2) yinv = 0` THENL [ complex_field_fact `y * yinv = Cx (&1) ==> ~(yinv = Cx(&0))` THEN qed[poly_ord_x_pow] ; pass ] THEN subgoal `ring_polynomial complex_ring ( poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) /\ ~(poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) yinv = (poly_ord f yinv + poly_ord g yinv) - 1` THENL [ specialize_assuming[ `x_pow complex_ring 2`; `poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))`; `yinv:complex` ]poly_ord_mul THEN have `ring_polynomial complex_ring (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) /\ ~(poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) yinv = 0 + poly_ord (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) yinv` [] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) /\ ~(poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))) = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) yinv = 0 + (poly_ord f yinv + poly_ord g yinv) - 1` [] THEN qed[ARITH_RULE `0+d = d`] ; pass ] THEN subgoal `poly_ord (poly_mul complex_ring e g) yinv = (poly_ord f yinv + poly_ord g yinv) - 1` THENL [ num_linear_fact `poly_ord e yinv + 1 = poly_ord g yinv ==> (poly_ord f yinv + poly_ord g yinv) - 1 < poly_ord f yinv + poly_ord g yinv` THEN specialize_assuming[ `poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)`; `poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))`; `yinv:complex`; `poly_ord f yinv + poly_ord g yinv` ]poly_ord_sub_dominant THEN have `poly_ord (poly_sub complex_ring (poly_mul complex_ring (poly_sub complex_ring (poly_1 complex_ring) (x_pow complex_ring 1)) (poly_mul complex_ring f g)) (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g))))) yinv = poly_ord (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) yinv` [] THEN have `poly_ord (poly_mul complex_ring e g) yinv = poly_ord (poly_mul complex_ring (x_pow complex_ring 2) (poly_sub complex_ring (poly_mul complex_ring (x_derivative complex_ring f) g) (poly_mul complex_ring f (x_derivative complex_ring g)))) yinv` [] THEN qed[] ; pass ] THEN have `poly_ord e yinv + poly_ord g yinv = (poly_ord f yinv + poly_ord g yinv) - 1` [] THEN num_linear_fact `poly_ord e yinv + 1 = poly_ord g yinv ==> poly_ord e yinv + poly_ord g yinv = (poly_ord f yinv + poly_ord g yinv) - 1 ==> ~(poly_ord f yinv < poly_ord g yinv)` THEN qed[] );; let transcendence_mostly_0 = prove(` !P B q y. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> q IN P ==> ~(B q = &0) ==> complex_root q y ==> y = Cx(&0) `, intro THEN def `t:num` `nsum P (\p. poly_deg complex_ring (p:(1->num)->complex))` THEN def `H:(1->num)->complex` `x_truncreverse QinC_ring t (poly_product QinC_ring P I):(1->num)->complex` THEN def `u:num->complex` `\n. ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n)` THEN have `!n. u n = (ring_sum complex_ring P (\p. (complex_of_int(B p)) * complex_root_powersums p n))` [] THEN def `v:num->complex` `\n. ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)` THEN have `!n. v n = ring_sum complex_ring (0..n) (\i. Cx(&(FACT(n-i) * binom(n,i))) * u i)` [] THEN case `~(y = Cx(&0))` THENL [ def `yinv:complex` `inv y:complex` THEN have `y * yinv = Cx(&1)` [COMPLEX_MUL_RINV] THEN def `U:(1->num)->complex` `series_from_coeffs u:(1->num)->complex` THEN def `V:(1->num)->complex` `series_from_coeffs v:(1->num)->complex` THEN have `poly_sub complex_ring (poly_sub complex_ring (V) (poly_mul complex_ring (x_pow complex_ring 1) (V))) (poly_mul complex_ring (x_pow complex_ring 2) (x_derivative complex_ring (V))) = U` [transcendence_uv_diffeq] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `v:num->complex` ]transcendence_Hv_poly THEN choose `k:num` `1 <= k /\ ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H k:(1->num)->complex) V)` [] THEN have `1 <= k` [] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->int`; `t:num`; `H:(1->num)->complex`; `u:num->complex`; `k:num` ]transcendence_ord_Hku THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) /\ ~(poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u) = poly_0 complex_ring) /\ poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) (series_from_coeffs u)) yinv = k - 1` [] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) U:(1->num)->complex)` [] THEN have `ring_polynomial complex_ring (poly_mul complex_ring (poly_pow complex_ring H k) V:(1->num)->complex)` [] THEN have `ring_polynomial complex_ring (poly_pow complex_ring H k:(1->num)->complex)` [poly_pow_poly;transcendence_H_poly] THEN have `coeff 0 (H:(1->num)->complex) = Cx(&1)` [transcendence_H_botcoeff1] THEN have `~(H = poly_0 complex_ring:(1->num)->complex)` [coeff_poly_0;complex_ring_clauses;ring_1_0_complex] THEN have `~(poly_pow complex_ring H k = poly_0 complex_ring:(1->num)->complex)` [nonzero_poly_pow;integral_domain_complex;series_complex] THEN have `~(poly_mul complex_ring (poly_pow complex_ring H k) U = poly_0 complex_ring:(1->num)->complex)` [] THEN have `poly_ord (poly_pow complex_ring H k) yinv = k` [transcendence_ord_Hk] THEN have `poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) U) yinv = k-1` [] THEN num_linear_fact `1 <= k ==> (k-1)+1 = k` THEN have `poly_ord (poly_mul complex_ring (poly_pow complex_ring H k) U) yinv + 1 = poly_ord (poly_pow complex_ring H k) yinv` [] THEN specialize[ `U:(1->num)->complex`; `V:(1->num)->complex`; `poly_mul complex_ring (poly_pow complex_ring H k) U:(1->num)->complex`; `poly_mul complex_ring (poly_pow complex_ring H k) V:(1->num)->complex`; `poly_pow complex_ring H k:(1->num)->complex`; `y:complex`; `yinv:complex` ]transcendence_pole_order_ne1 THEN qed[] ; qed[] ] );; let transcendence_mostly_0_x = prove(` !P B q. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> q IN P ==> ~(B q = &0) ==> q = x_pow QinC_ring 1 `, intro THEN have `ring_polynomial QinC_ring (q:(1->num)->complex)` [distinct_minpolys] THEN have `ring_polynomial complex_ring (q:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(poly_deg complex_ring (q:(1->num)->complex) = 0)` [distinct_minpolys_deg_nonzero] THEN num_linear_fact `~(poly_deg complex_ring (q:(1->num)->complex) = 0) ==> 1 <= poly_deg complex_ring q` THEN choose `z:complex` `poly_eval complex_ring q z = Cx(&0)` [nonconstant_complex_root] THEN have `complex_root q z` [complex_root] THEN have `z = Cx(&0)` [transcendence_mostly_0] THEN qed[distinct_minpolys_zero_root] );; let transcendence_all_0 = prove(` !P B. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. complex_of_int(B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!q. q IN P ==> B q = &0) `, intro THEN case `q = x_pow QinC_ring 1` THENL [ subgoal `ring_sum complex_ring P (\p. complex_of_int (B p) * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring P (\p. if p = x_pow QinC_ring 1 then complex_of_int (B(x_pow QinC_ring 1)) * ring_sum complex_ring (complex_root(x_pow QinC_ring 1)) cexp else ring_0 complex_ring)` THENL [ sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM] THEN proven_if `a = x_pow QinC_ring 1` [] THEN have `B(a:(1->num)->complex) = &0:int` [transcendence_mostly_0_x] THEN simp[complex_of_int;GSYM REAL_OF_INT_CLAUSES] THEN qed[COMPLEX_MUL_LZERO;complex_ring_clauses] ; pass ] THEN have `ring_sum complex_ring P (\p. if p = x_pow QinC_ring 1 then complex_of_int (B(x_pow QinC_ring 1)) * ring_sum complex_ring (complex_root(x_pow QinC_ring 1)) cexp else ring_0 complex_ring) = Cx(&0)` [] THEN specialize[ `complex_ring`; `P:((1->num)->complex)->bool`; `x_pow QinC_ring 1`; `complex_of_int (B (x_pow QinC_ring 1)) * ring_sum complex_ring (complex_root (x_pow QinC_ring 1)) cexp` ]RING_SUM_DELTA THEN have `complex_of_int (B (x_pow QinC_ring 1)) * ring_sum complex_ring (complex_root (x_pow QinC_ring 1)) cexp = Cx(&0)` [in_complex_ring] THEN subgoal `ring_sum complex_ring (complex_root (x_pow QinC_ring 1)) cexp = Cx(&1)` THENL [ rw[complex_root_x_pow;ARITH_RULE `~(1 = 0)`;RING_SUM_SING;in_complex_ring;CEXP_0] ; pass ] THEN have `complex_of_int (B (x_pow QinC_ring 1)) * Cx(&1) = Cx(&0)` [] THEN have `complex_of_int (B (x_pow QinC_ring 1)) = Cx(&0)` [COMPLEX_MUL_RID] THEN have `real_of_int (B (x_pow QinC_ring 1)) = &0` [complex_of_int;CX_INJ] THEN qed[REAL_OF_INT_CLAUSES] ; pass ] THEN qed[transcendence_mostly_0_x] );; let transcendence_all_0_ZinC = prove(` !P B. FINITE P ==> distinct_minpolys P ==> (!p. p IN P ==> B p IN ZinC) ==> ring_sum complex_ring P (\p. (B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p. p IN P ==> B p = Cx(&0)) `, intro THEN def `intB:((1->num)->complex)->int` `\p:(1->num)->complex. @i. B p = complex_of_int i` THEN subgoal `!q:(1->num)->complex. q IN P ==> B q = complex_of_int(intB q)` THENL [ intro THEN simp[] THEN have `B(q:(1->num)->complex) IN ZinC` [] THEN qed[ZinC;IN_ELIM_THM] ; pass ] THEN subgoal `ring_sum complex_ring P (\p. complex_of_int (intB p) * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp)` THENL [ sufficesby RING_SUM_EQ THEN qed[] ; pass ] THEN have `ring_sum complex_ring P (\p. complex_of_int (intB p) * ring_sum complex_ring (complex_root p) cexp) = Cx (&0)` [] THEN specialize_assuming[ `P:((1->num)->complex)->bool`; `intB:((1->num)->complex)->int` ]transcendence_all_0 THEN have `intB(p:(1->num)->complex) = &0:int` [] THEN have `B(p:(1->num)->complex) = complex_of_int(&0)` [] THEN qed[complex_of_int;REAL_OF_INT_CLAUSES] );; let transcendence_all_0_QinC = prove(` !P B. FINITE P ==> distinct_minpolys P ==> (!p. p IN P ==> B p IN QinC) ==> ring_sum complex_ring P (\p. (B p) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> (!p. p IN P ==> B p = Cx(&0)) `, intro THEN choose `d:num` `(~(d = 0) /\ !p:(1->num)->complex. p IN P ==> Cx(&d) * B p IN ZinC)` [multi_QinC_to_ZinC] THEN def `C:((1->num)->complex)->complex` `\p:(1->num)->complex. Cx(&d) * B p` THEN have `!p:(1->num)->complex. p IN P ==> C p IN ZinC` [] THEN subgoal `ring_sum complex_ring P (\p. C p * ring_sum complex_ring (complex_root p) cexp) = Cx(&0)` THENL [ simp[GSYM COMPLEX_MUL_ASSOC] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[VSUM_COMPLEX_LMUL] THEN simp[vsum_ring_sum_complex] THEN simp[COMPLEX_MUL_RZERO] ; pass ] THEN have `C(p:(1->num)->complex) = Cx(&0)` [transcendence_all_0_ZinC] THEN have `Cx(&d) * B(p:(1->num)->complex) = Cx(&0)` [] THEN have `~(&d = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&d) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx(&d) = Cx(&0)) ==> Cx(&d) * B(p:(1->num)->complex) = Cx(&0) ==> B p = Cx(&0)` THEN qed[] );; let transcendence_all_0_QinC_weight1 = prove(` !P. FINITE P ==> distinct_minpolys P ==> ring_sum complex_ring P (\p. ring_sum complex_ring (complex_root p) cexp) = Cx(&0) ==> P = {} `, intro THEN def `B:((1->num)->complex)->complex` `\p:(1->num)->complex. Cx(&1)` THEN have `!p:(1->num)->complex. p IN P ==> B p IN QinC` [QinC_1] THEN subgoal `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring P (\p. ring_sum complex_ring (complex_root p) cexp)` THENL [ sufficesby RING_SUM_EQ THEN qed[COMPLEX_MUL_LID] ; pass ] THEN have `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = Cx(&0)` [] THEN have `!p:(1->num)->complex. p IN P ==> B p = Cx(&0)` [transcendence_all_0_QinC] THEN complex_field_fact `~(Cx(&1) = Cx(&0))` THEN have `!p:(1->num)->complex. ~(p IN P)` [] THEN qed[EMPTY;EXTENSION;IN] );; let transcendence_all_0_QinC_monic_vanishing_at = prove(` !S c. FINITE S ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring S c) ==> ring_sum complex_ring S (\s:X. cexp(c s)) = Cx(&0) ==> S = {} `, intro THEN have `monic complex_ring (monic_vanishing_at complex_ring S (c:X->complex))` [monic_vanishing_at_monic;in_complex_ring] THEN have `monic QinC_ring (monic_vanishing_at complex_ring S (c:X->complex))` [monic_subring;subring_complex_QinC] THEN recall field_QinC THEN specialize[ `monic_vanishing_at complex_ring S (c:X->complex)` ]monic_factorization_distinct_minpolys THEN choose2 `P:((1->num)->complex)->bool` `e:((1->num)->complex)->num` `FINITE P /\ distinct_minpolys P /\ (!p. p IN P ==> ~(e p = 0)) /\ poly_product QinC_ring P (\p. poly_pow QinC_ring p (e p)) = monic_vanishing_at complex_ring S (c:X->complex)` [] THEN have `!p:(1->num)->complex. Cx(&(e p)) IN QinC` [num_in_QinC] THEN subgoal `ring_sum complex_ring P (\p. Cx(&(e p)) * ring_sum complex_ring (complex_root p) cexp) = Cx(&0)` THENL [ have `!z. poly_ord (monic_vanishing_at complex_ring S c) z = CARD {s:X | s IN S /\ z = c s:complex}` [poly_ord_monic_vanishing_at] THEN have `!z. poly_ord (poly_product QinC_ring P (\p:(1->num)->complex. poly_pow QinC_ring p (e p))) z = nsum P (\p. if complex_root p z then e p else 0)` [ord_product_distinct_minpolys] THEN have `!z. CARD {s:X | s IN S /\ z = c s:complex} = nsum P (\p. if complex_root p z then e p else 0)` [] THEN set_fact `!z. {s:X | s IN S /\ c s = z:complex} = {s:X | s IN S /\ z = c s:complex}` THEN subgoal `!z. FINITE {s:X | s IN S /\ c s = z:complex}` THENL [ intro THEN specialize[ `S:X->bool`; `\s:X. c s = z:complex` ]FINITE_RESTRICT THEN qed[] ; pass ] THEN have `FINITE(IMAGE (c:X->complex) S)` [FINITE_IMAGE] THEN subgoal `ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. ring_sum complex_ring P (\p. Cx(&(if complex_root p z then e p else 0)) * cexp z)) = Cx(&0)` THENL [ subgoal `ring_sum complex_ring (IMAGE c S) (\z. Cx(&(CARD {x:X | x IN S /\ c x = z})) * cexp z) = Cx(&0)` THENL [ specialize[ `complex_ring`; `c:X->complex`; `\s:X. cexp (c s)`; `S:X->bool`; ]RING_SUM_IMAGE_GEN THEN have `ring_sum complex_ring (IMAGE c S) (\y. ring_sum complex_ring {x | x IN S /\ c x = y} (\s:X. cexp (c s))) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring (IMAGE c S) (\y. ring_sum complex_ring {x | x IN S /\ c x = y} (\s:X. cexp (c s))) = ring_sum complex_ring (IMAGE c S) (\y. ring_sum complex_ring {x | x IN S /\ c x = y} (\s:X. cexp y))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN sufficesby RING_SUM_EQ THEN rw[IN_ELIM_THM] THEN qed[] ; pass ] THEN have `ring_sum complex_ring (IMAGE c S) (\y. ring_sum complex_ring {x | x IN S /\ c x = y} (\s:X. cexp y)) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring (IMAGE c S) (\z. Cx(&(CARD {x | x IN S /\ c x = z})) * cexp z) = ring_sum complex_ring (IMAGE c S) (\y. ring_sum complex_ring {x | x IN S /\ c x = y} (\s:X. cexp y))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `FINITE {x:X | x IN S /\ c x = a:complex}` [] THEN have `cexp a IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `cexp a`; `{x:X | x IN S /\ c x = a:complex}` ]ring_sum_const THEN qed[ring_of_num_complex;complex_of_num;complex_ring_clauses] ; pass ] THEN simporqed ; pass ] THEN subgoal `ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(nsum P (\p. if complex_root p z then e p else 0))) * cexp z) = Cx(&0)` THENL [ subgoal `ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(nsum P (\p. if complex_root p z then e p else 0))) * cexp z) = ring_sum complex_ring (IMAGE c S) (\z. Cx(&(CARD {x:X | x IN S /\ c x = z})) * cexp z)` THENL [ sufficesby RING_SUM_EQ THEN simporqed ; pass ] THEN simporqed ; pass ] THEN subgoal `ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(nsum P (\p. if complex_root p z then e p else 0))) * cexp z) = ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. ring_sum complex_ring P (\p. Cx(&(if complex_root p z then e p else 0)) * cexp z))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN simp[GSYM vsum_ring_sum_complex] THEN simp[VSUM_COMPLEX_RMUL] THEN simp[vsum_ring_sum_complex] THEN rw[GSYM complex_of_num] THEN rw[GSYM ring_of_num_complex] THEN simp[ring_sum_num] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. ring_sum complex_ring P (\p. Cx(&(if complex_root p z then e p else 0)) * cexp z)) = ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(if complex_root p z then e p else 0)) * cexp z))` THENL [ specialize_assuming[ `complex_ring`; `\z p. Cx (&(if complex_root p z then e p else 0)) * cexp z`; `IMAGE (c:X->complex) S`; `P:((1->num)->complex)->bool` ]RING_SUM_SWAP THEN qed[in_complex_ring] ; pass ] THEN have `ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(if complex_root p z then e p else 0)) * cexp z)) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. Cx(&(if complex_root p z then e p else 0)) * cexp z)) = ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. if z IN complex_root p then Cx(&(e p)) * cexp z else ring_0 complex_ring))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN] THEN intro THEN rw[complex_ring_clauses] THEN qed[COMPLEX_MUL_LZERO] ; pass ] THEN have `ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. if z IN complex_root p then Cx(&(e p)) * cexp z else ring_0 complex_ring)) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring P (\p. ring_sum complex_ring (IMAGE (c:X->complex) S) (\z. if z IN complex_root p then Cx(&(e p)) * cexp z else ring_0 complex_ring)) = ring_sum complex_ring P (\p. ring_sum complex_ring {z | z IN IMAGE (c:X->complex) S /\ z IN complex_root p} (\z. Cx(&(e p)) * cexp z))` THENL [ sufficesby RING_SUM_EQ THEN simp[RING_SUM_RESTRICT_SET] ; pass ] THEN have `ring_sum complex_ring P (\p. ring_sum complex_ring {z | z IN IMAGE (c:X->complex) S /\ z IN complex_root p} (\z. Cx(&(e p)) * cexp z)) = Cx(&0)` [] THEN subgoal `!p. p IN P ==> {z | z IN IMAGE (c:X->complex) S /\ z IN complex_root p} = complex_root p` THENL [ intro THEN rw[EXTENSION;IN_ELIM_THM;IN_IMAGE] THEN intro THEN splitiff THENL [ qed[] ; intro THEN have `FINITE(P:((1->num)->complex)->bool)` [] THEN specialize[ `\p. if complex_root p x then e p else 0`; `P:((1->num)->complex)->bool`; `p:(1->num)->complex` ]term_le_nsum THEN have `e(p:(1->num)->complex) <= nsum P (\p. if complex_root p x then e p else 0)` [IN] THEN have `~(nsum P (\p. if complex_root p x then e p else 0) = 0)` [ARITH_RULE `~(e = 0) ==> e <= n ==> ~(n = 0)`] THEN have `~(CARD {s:X | s IN S /\ x = c s:complex} = 0)` [] THEN specialize[ `S:X->bool`; `\s:X. x = c s:complex` ]FINITE_RESTRICT THEN have `~({s:X | s IN S /\ x = c s:complex} = {})` [CARD_EQ_0] THEN ASM SET_TAC[] ] ; pass ] THEN subgoal `ring_sum complex_ring P (\p. ring_sum complex_ring {z | z IN IMAGE (c:X->complex) S /\ z IN complex_root p} (\z. Cx (&(e p)) * cexp z)) = ring_sum complex_ring P (\p. Cx(&(e p)) * ring_sum complex_ring (complex_root p) cexp)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN simp[] THEN have `FINITE(complex_root a)` [distinct_minpolys_finite_root] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[VSUM_COMPLEX_LMUL] ; pass ] THEN qed[] ; pass ] THEN specialize_assuming[ `P:((1->num)->complex)->bool`; `\p:(1->num)->complex. Cx(&(e p))` ]transcendence_all_0_QinC THEN have `!p:(1->num)->complex. p IN P ==> Cx (&(e p)) = Cx (&0)` [] THEN have `!p:(1->num)->complex. p IN P ==> &(e p) = &0:real` [CX_INJ] THEN have `!p:(1->num)->complex. p IN P ==> e p = 0` [REAL_OF_NUM_EQ] THEN subgoal `poly_product QinC_ring P (\p:(1->num)->complex. poly_pow QinC_ring p (e p)) = poly_1 QinC_ring` THENL [ subgoal `poly_product QinC_ring P (\p:(1->num)->complex. poly_pow QinC_ring p (e p)) = poly_product QinC_ring P (\p. poly_1 QinC_ring)` THENL [ sufficesby poly_product_eq THEN qed[poly_pow_0] ; pass ] THEN specialize[ `QinC_ring`; `P:((1->num)->complex)->bool` ]poly_product_1 THEN qed[] ; pass ] THEN have `monic_vanishing_at complex_ring S (c:X->complex) = poly_1 QinC_ring` [] THEN have `poly_deg complex_ring (monic_vanishing_at complex_ring S (c:X->complex)) = CARD S` [deg_monic_vanishing_at;ring_1_0_complex;in_complex_ring] THEN have `poly_deg complex_ring (poly_1 QinC_ring:(1->num)->complex) = 0` [poly_1_QinC_eq_poly_1_complex;POLY_DEG_1] THEN have `CARD(S:X->bool) = 0` [] THEN qed[CARD_EQ_0] );; let transcendence_weighted_QinC_monic_vanishing_at = prove(` !S:X->bool b:X->complex c:Y->complex f:X->(Y->bool). FINITE S ==> (!s. s IN S ==> b s IN QinC) ==> (!s. s IN S ==> FINITE(f s)) ==> (!s. s IN S ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring (f s) c)) ==> ring_sum complex_ring S (\s. b s * ring_sum complex_ring (f s) (cexp o c)) = Cx(&0) ==> (!a. ring_sum complex_ring S (\s. b s * Cx(&(numpreimages c (f s) a))) = Cx(&0)) `, intro THEN have `!s:X. s IN S ==> monic complex_ring (monic_vanishing_at complex_ring (f s) (c:Y->complex))` [monic_vanishing_at_monic;in_complex_ring] THEN have `!s:X. s IN S ==> monic QinC_ring (monic_vanishing_at complex_ring (f s) (c:Y->complex))` [monic_subring;subring_complex_QinC] THEN have `!s:X. s IN S ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring (f s) (c:Y->complex))` [monic_vanishing_at_poly] THEN recall field_QinC THEN specialize[ `S:X->bool`; `\s:X. monic_vanishing_at complex_ring (f s) (c:Y->complex)` ]multi_monic_factorization_distinct_minpolys THEN choose2 `P:((1->num)->complex)->bool` `e:X->((1->num)->complex)->num` `FINITE P /\ distinct_minpolys P /\ (!s:X. s IN S ==> poly_product QinC_ring P (\p. poly_pow QinC_ring p (e s p)) = monic_vanishing_at complex_ring (f s) (c:Y->complex))` [] THEN def `B:((1->num)->complex)->complex` `\p:(1->num)->complex. ring_sum complex_ring S (\s:X. b s * Cx(&(e s p)))` THEN subgoal `!p:(1->num)->complex. p IN P ==> B p IN QinC` THENL [ simp[] THEN rw[GSYM QinC_ring_clauses] THEN rw[GSYM subring_complex_QinC] THEN intro THEN sufficesby ring_sum_in_subring THEN rw[BETA_THM] THEN intro THEN sufficesby RING_MUL THEN rw[subring_complex_QinC] THEN rw[QinC_ring_clauses] THEN qed[num_in_QinC] ; pass ] THEN have `FINITE(P:((1->num)->complex)->bool)` [] THEN have `distinct_minpolys P` [] THEN subgoal `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = Cx (&0)` THENL [ simp[] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[GSYM VSUM_COMPLEX_RMUL] THEN specialize_assuming[ `\p s:X. (b s * Cx(&(e s p))) * ring_sum complex_ring (complex_root p) cexp`; `P:((1->num)->complex)->bool`; `S:X->bool` ]VSUM_SWAP THEN simp[] THEN rw[GSYM COMPLEX_MUL_ASSOC] THEN simp[VSUM_COMPLEX_LMUL] THEN subgoal `!s:X. s IN S ==> vsum P (\p. (Cx (&(e s p))) * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring (f s:Y->bool) (cexp o c)` THENL [ simp[vsum_ring_sum_complex] THEN intro THEN have `FINITE(f(s:X):Y->bool)` [] THEN have `poly_product QinC_ring P (\p. poly_pow QinC_ring p (e s p)) = monic_vanishing_at complex_ring (f(s:X):Y->bool) c` [] THEN specialize[ `P:((1->num)->complex)->bool`; `f(s:X):Y->bool`; `e(s:X):((1->num)->complex)->num`; `c:Y->complex`; `cexp` ]sum_root_decomposition_if_monic_vanishing_at_factorization THEN qed[] ; pass ] THEN simp[] THEN qed[vsum_ring_sum_complex] ; pass ] THEN subgoal `!p:(1->num)->complex. p IN P ==> B p = Cx(&0)` THENL [ specialize_assuming[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->complex` ]transcendence_all_0_QinC THEN qed[] ; pass ] THEN subgoal `!s:X. s IN S ==> numpreimages c (f s:Y->bool) a = nsum P (\p. if complex_root p a then e s p else 0)` THENL [ intro THEN specialize_assuming[ `P:((1->num)->complex)->bool`; `e(s:X):((1->num)->complex)->num`; `a:complex` ]ord_product_distinct_minpolys THEN have `FINITE(f(s:X):Y->bool)` [] THEN specialize[ `f(s:X):Y->bool`; `c:Y->complex`; `a:complex` ]poly_ord_monic_vanishing_at_numpreimages THEN qed[] ; pass ] THEN case `!p:(1->num)->complex. p IN P ==> ~complex_root p a` THENL [ subgoal `!s:X. s IN S ==> nsum P (\p. if complex_root p a then e s p else 0) = nsum P (\p. 0)` THENL [ intro THEN sufficesby NSUM_EQ THEN qed[] ; pass ] THEN have `!s:X. s IN S ==> numpreimages c (f s:Y->bool) (a:complex) = 0` [NSUM_0] THEN subgoal `ring_sum complex_ring S (\s:X. b s * Cx (&(numpreimages c (f s:Y->bool) (a:complex)))) = ring_sum complex_ring S (\s. Cx(&0))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `numpreimages c (f(a':X):Y->bool) (a:complex) = 0` [] THEN simp[] THEN qed[COMPLEX_MUL_RZERO] ; pass ] THEN simp[] THEN rw[GSYM complex_ring_clauses] THEN rw[RING_SUM_0] ; pass ] THEN simp[] THEN choose `q:(1->num)->complex` `q IN P /\ complex_root q a` [] THEN have_rw `Cx(&0) = B(q:(1->num)->complex)` [] THEN have_rw `B(q:(1->num)->complex) = ring_sum complex_ring S (\s:X. b s * Cx(&(e s q)))` [] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `nsum P (\p. if complex_root p a then e a' p else 0) = nsum P (\p. if p = q then e(a':X) q else 0)` THENL [ sufficesby NSUM_EQ THEN qed[distinct_minpolys_distinct_roots] ; pass ] THEN simp[NSUM_DELTA] );; (* ===== e is transcendental *) (* warmup, marginally easier *) let e_is_irrational = prove(` ~(cexp(Cx(&1)) IN QinC) `, intro THEN def `P:((1->num)->complex)->bool` `IMAGE (\n. x_minus_const QinC_ring (Cx(&n))) (0..1)` THEN have `FINITE(P:((1->num)->complex)->bool)` [FINITE_IMAGE;FINITE_NUMSEG] THEN subgoal `distinct_minpolys P` THENL [ rw[distinct_minpolys] THEN REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN choose `n:num` `n IN (0..1) /\ p = x_minus_const QinC_ring(Cx(&n))` [IN_IMAGE] THEN have `Cx(&n) IN QinC` [num_in_ZinC;ZinC_in_QinC] THEN qed[irred_x_minus_const;field_QinC;QinC_ring_clauses] ; pass ] THEN def `B:((1->num)->complex)->complex` `\p:(1->num)->complex. if p = x_minus_const QinC_ring (Cx(&0)) then -- cexp(Cx(&1)) else Cx(&1)` THEN subgoal `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = Cx (&0)` THENL [ subgoal `!m n. x_minus_const QinC_ring (Cx(&m)) = x_minus_const QinC_ring (Cx(&n)) ==> m = n` THENL [ intro THEN have `coeff 0 (x_minus_const QinC_ring (Cx(&m))) = coeff 0 (x_minus_const QinC_ring (Cx(&n)))` [] THEN have `coeff 0 (x_minus_const QinC_ring (Cx(&m))) = ring_neg QinC_ring (Cx(&n))` [coeff_x_minus_const;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses;ARITH_RULE `~(0 = 1)`] THEN have `ring_neg QinC_ring (Cx(&m)) = ring_neg QinC_ring (Cx(&n))` [coeff_x_minus_const;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses;ARITH_RULE `~(0 = 1)`] THEN have `Cx(&m) = Cx(&n)` [RING_NEG_EQ;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses] THEN have `&m = &n:real` [CX_INJ] THEN qed[REAL_OF_NUM_CLAUSES] ; pass ] THEN have `!m n. m IN (0..1) ==> n IN (0..1) ==> x_minus_const QinC_ring (Cx(&m)) = x_minus_const QinC_ring (Cx(&n)) ==> m = n` [] THEN specialize[ `complex_ring`; `\n. x_minus_const QinC_ring (Cx(&n))`; `\p:(1->num)->complex. B p * ring_sum complex_ring (complex_root p) cexp`; `(0..1)` ]RING_SUM_IMAGE THEN have_rw `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring (0..1) ((\p. B p * ring_sum complex_ring (complex_root p) cexp) o (\n. x_minus_const QinC_ring (Cx (&n))))` [] THEN rw[RING_SUM_CLAUSES_NUMSEG;ARITH_RULE `1 = SUC 0`;in_complex_ring;ARITH_RULE `0 <= SUC 0`;o_THM] THEN rw[ARITH_RULE `SUC 0 = 1`] THEN have_rw `complex_root (x_minus_const QinC_ring (Cx (&0))) = {Cx(&0)}` [complex_root_x_minus_const;x_minus_const_QinC_eq_x_minus_const_complex] THEN have_rw `complex_root (x_minus_const QinC_ring (Cx (&1))) = {Cx(&1)}` [complex_root_x_minus_const;x_minus_const_QinC_eq_x_minus_const_complex] THEN rw[RING_SUM_SING;in_complex_ring] THEN have_rw `B (x_minus_const QinC_ring (Cx (&0))) = --cexp(Cx(&1))` [] THEN have_rw `B (x_minus_const QinC_ring (Cx (&1))) = Cx(&1)` [ARITH_RULE `~(0 = 1)`] THEN rw[complex_ring_clauses;CEXP_0] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN subgoal `!p:(1->num)->complex. p IN P ==> B p IN QinC` THENL [ intro THEN have `--cexp(Cx(&1)) IN QinC` [neg_in_QinC] THEN have `Cx(&1) IN QinC` [QinC_1] THEN qed[] ; pass ] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->complex` ]transcendence_all_0_QinC THEN have `x_minus_const QinC_ring(Cx(&0)) IN P` [IN_NUMSEG_0;ARITH_RULE `0 <= 1`;IN_IMAGE] THEN have `--cexp(Cx(&1)) = Cx(&0)` [] THEN complex_field_fact `--cexp(Cx(&1)) = Cx(&0) ==> cexp(Cx(&1)) = Cx(&0)` THEN qed[CEXP_NZ] );; let e_is_transcendental = prove(` ~(algebraic_number(cexp(Cx(&1)))) `, rw[algebraic_number] THEN intro THEN def `d:num` `poly_deg ZinC_ring (p:(1->num)->complex)` THEN have `d = poly_deg complex_ring (p:(1->num)->complex)` [poly_deg_subring;subring_complex_empty] THEN def `P:((1->num)->complex)->bool` `IMAGE (\n. x_minus_const QinC_ring (Cx(&n))) (0..d)` THEN have `FINITE(P:((1->num)->complex)->bool)` [FINITE_IMAGE;FINITE_NUMSEG] THEN subgoal `distinct_minpolys P` THENL [ rw[distinct_minpolys] THEN REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN choose `n:num` `p' = x_minus_const QinC_ring(Cx(&n))` [IN_IMAGE] THEN have `Cx(&n) IN QinC` [num_in_ZinC;ZinC_in_QinC] THEN qed[irred_x_minus_const;field_QinC;QinC_ring_clauses] ; pass ] THEN subgoal `!m n. x_minus_const QinC_ring (Cx(&m)) = x_minus_const QinC_ring (Cx(&n)) ==> m = n` THENL [ intro THEN have `coeff 0 (x_minus_const QinC_ring (Cx(&m))) = coeff 0 (x_minus_const QinC_ring (Cx(&n)))` [] THEN have `coeff 0 (x_minus_const QinC_ring (Cx(&m))) = ring_neg QinC_ring (Cx(&n))` [coeff_x_minus_const;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses;ARITH_RULE `~(0 = 1)`] THEN have `ring_neg QinC_ring (Cx(&m)) = ring_neg QinC_ring (Cx(&n))` [coeff_x_minus_const;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses;ARITH_RULE `~(0 = 1)`] THEN have `Cx(&m) = Cx(&n)` [RING_NEG_EQ;num_in_ZinC;ZinC_in_QinC;QinC_ring_clauses] THEN have `&m = &n:real` [CX_INJ] THEN qed[REAL_OF_NUM_CLAUSES] ; pass ] THEN def `B:((1->num)->complex)->complex` `\q:(1->num)->complex. coeff (@n. q = x_minus_const QinC_ring (Cx(&n))) (p:(1->num)->complex)` THEN subgoal `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = Cx (&0)` THENL [ have `!m n. m IN (0..d) ==> n IN (0..d) ==> x_minus_const QinC_ring (Cx(&m)) = x_minus_const QinC_ring (Cx(&n)) ==> m = n` [] THEN specialize[ `complex_ring`; `\n. x_minus_const QinC_ring (Cx(&n))`; `\p:(1->num)->complex. B p * ring_sum complex_ring (complex_root p) cexp`; `(0..d)` ]RING_SUM_IMAGE THEN have_rw `ring_sum complex_ring P (\p. B p * ring_sum complex_ring (complex_root p) cexp) = ring_sum complex_ring (0..d) ((\p. B p * ring_sum complex_ring (complex_root p) cexp) o (\n. x_minus_const QinC_ring (Cx (&n))))` [] THEN subgoal `poly_eval complex_ring p (cexp(Cx(&1))) = ring_sum complex_ring (0..d) ((\p. B p * ring_sum complex_ring (complex_root p) cexp) o (\n. x_minus_const QinC_ring (Cx (&n))))` THENL [ have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_ZinC] THEN have `p IN ring_carrier(x_poly complex_ring)` [x_poly_use] THEN have `p IN ring_carrier(poly_ring complex_ring (:1))` [x_poly] THEN specialize_assuming[`complex_ring`;`cexp(Cx(&1))`;`p:(1->num)->complex`]POLY_EVAL_EXPAND THEN have_rw `poly_eval complex_ring p (cexp (Cx (&1))) = ring_sum complex_ring (0..d) (\i. ring_mul complex_ring (p (\v. i)) (ring_pow complex_ring (cexp (Cx (&1))) i))` [in_complex_ring] THEN sufficesby RING_SUM_EQ THEN intro THEN rw[BETA_THM;o_THM] THEN rw[complex_ring_clauses;ring_pow_complex] THEN rw[complex_root_x_minus_const;x_minus_const_QinC_eq_x_minus_const_complex] THEN rw[GSYM CEXP_N;COMPLEX_MUL_RID] THEN have `B (x_minus_const QinC_ring (Cx (&a))) = coeff a p:complex` [] THEN have `B (x_minus_const QinC_ring (Cx (&a))) = p (\v:1. a):complex` [coeff;x_monomial] THEN have `B (x_minus_const complex_ring (Cx (&a))) = p (\v:1. a):complex` [x_minus_const_QinC_eq_x_minus_const_complex] THEN qed[RING_SUM_SING;in_complex_ring] ; pass ] THEN qed[] ; pass ] THEN subgoal `!q:(1->num)->complex. q IN P ==> B q IN QinC` THENL [ intro THEN have `ring_polynomial QinC_ring (p:(1->num)->complex)` [poly_QinC_if_poly_ZinC] THEN qed [coeff_poly_in_ring;QinC_ring_clauses] ; pass ] THEN specialize[ `P:((1->num)->complex)->bool`; `B:((1->num)->complex)->complex` ]transcendence_all_0_QinC THEN have `~(coeff d p = Cx(&0))` [topcoeff_nonzero;ZinC_ring_clauses] THEN have `x_minus_const QinC_ring (Cx(&d)) IN P` [IN_IMAGE;IN_NUMSEG_0;ARITH_RULE `d <= d:num`] THEN have `B(x_minus_const QinC_ring (Cx(&d))) = Cx(&0)` [] THEN have `B(x_minus_const QinC_ring (Cx(&d))) = coeff d p:complex` [] THEN qed[] );; (* ===== algebraic numbers closed under mul, add *) let algebraic_number_mul = prove(` !y z. algebraic_number y ==> algebraic_number z ==> algebraic_number (y*z) `, intro THEN choose `p:(1->num)->complex` `(ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ monic QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p y = Cx(&0))` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN choose `q:(1->num)->complex` `(ring_polynomial QinC_ring q /\ ~(q = poly_0 QinC_ring) /\ monic QinC_ring q /\ ring_irreducible(x_poly QinC_ring) q /\ poly_eval complex_ring q z = Cx(&0))` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN have `!e. complex_root_powersums p e IN QinC` [complex_root_powersums_QinC_monic_irreducible_QinC] THEN have `!e. complex_root_powersums q e IN QinC` [complex_root_powersums_QinC_monic_irreducible_QinC] THEN def `pq:(1->num)->complex` `monic_vanishing_at complex_ring (complex_root p CROSS complex_root q) (\(y,z). y*z)` THEN have `FINITE (complex_root p)` [distinct_minpolys_finite_root_simple] THEN have `FINITE (complex_root q)` [distinct_minpolys_finite_root_simple] THEN have `FINITE (complex_root p CROSS complex_root q)` [FINITE_CROSS] THEN have `y IN complex_root p` [complex_root;IN] THEN have `z IN complex_root q` [complex_root;IN] THEN have `(y,z) IN complex_root p CROSS complex_root q` [IN_CROSS] THEN subgoal `ring_polynomial QinC_ring (pq:(1->num)->complex)` THENL [ have `ring_hasQ (subring_generated complex_ring QinC)` [ring_hasQ_QinC;subring_complex_QinC] THEN subgoal `!d. ring_sum complex_ring (complex_root p CROSS complex_root q) (\s. ring_pow complex_ring ((\(y,z). y * z) s) d) IN ring_carrier (subring_generated complex_ring QinC)` THENL [ intro THEN rw[subring_complex_QinC] THEN rw[GSYM complex_ring_clauses] THEN rw[LAMBDA_PAIR] THEN simp[RING_MUL_POW;in_complex_ring] THEN rw[GSYM LAMBDA_PAIR] THEN have `!x. x IN complex_root p ==> ring_pow complex_ring x d IN ring_carrier complex_ring` [in_complex_ring] THEN have `!y. y IN complex_root q ==> ring_pow complex_ring y d IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `complex_root p`; `complex_root q`; `\x. ring_pow complex_ring x d`; `\y. ring_pow complex_ring y d` ]ring_sum_cross_mul THEN simp[] THEN rw[ring_pow_complex;complex_ring_clauses;GSYM complex_root_powersums;QinC_ring_clauses] THEN qed[mul_in_QinC] ; pass ] THEN specialize_assuming[ `complex_ring`; `QinC`; `complex_root p CROSS complex_root q`; `\(y,z). y * z:complex` ]poly_subring_if_powersums_subring THEN qed[in_complex_ring;subring_complex_QinC] ; pass ] THEN have `algebraic_number ((\(y,z). y*z) (y,z))` [algebraic_number_if_monic_vanishing_at_QinC] THEN havetac `((\(y,z). y*z) (y,z)) = y*z:complex` (rw[]) THEN qed[] );; (* or can prove this more directly on poly coeffs *) let algebraic_number_neg = prove(` !z. algebraic_number z ==> algebraic_number (-- z) `, intro THEN complex_field_fact `--z = (-- Cx(&1)) * z` THEN have `Cx(&1) IN ZinC` [ZinC_1] THEN have `-- Cx(&1) IN ZinC` [neg_in_ZinC] THEN have `algebraic_number (-- Cx(&1))` [algebraic_number_ZinC] THEN qed[algebraic_number_mul] );; let algebraic_number_add = prove(` !y z. algebraic_number y ==> algebraic_number z ==> algebraic_number (y+z) `, intro THEN choose `p:(1->num)->complex` `(ring_polynomial QinC_ring p /\ ~(p = poly_0 QinC_ring) /\ monic QinC_ring p /\ ring_irreducible(x_poly QinC_ring) p /\ poly_eval complex_ring p y = Cx(&0))` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN choose `q:(1->num)->complex` `(ring_polynomial QinC_ring q /\ ~(q = poly_0 QinC_ring) /\ monic QinC_ring q /\ ring_irreducible(x_poly QinC_ring) q /\ poly_eval complex_ring q z = Cx(&0))` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN have `!e. complex_root_powersums p e IN QinC` [complex_root_powersums_QinC_monic_irreducible_QinC] THEN have `!e. complex_root_powersums q e IN QinC` [complex_root_powersums_QinC_monic_irreducible_QinC] THEN def `pq:(1->num)->complex` `monic_vanishing_at complex_ring (complex_root p CROSS complex_root q) (\(y,z). y+z)` THEN have `FINITE (complex_root p)` [distinct_minpolys_finite_root_simple] THEN have `FINITE (complex_root q)` [distinct_minpolys_finite_root_simple] THEN have `FINITE (complex_root p CROSS complex_root q)` [FINITE_CROSS] THEN have `y IN complex_root p` [complex_root;IN] THEN have `z IN complex_root q` [complex_root;IN] THEN have `(y,z) IN complex_root p CROSS complex_root q` [IN_CROSS] THEN subgoal `ring_polynomial QinC_ring (pq:(1->num)->complex)` THENL [ have `ring_hasQ (subring_generated complex_ring QinC)` [ring_hasQ_QinC;subring_complex_QinC] THEN subgoal `!d. ring_sum complex_ring (complex_root p CROSS complex_root q) (\s. ring_pow complex_ring ((\(y,z). y + z) s) d) IN ring_carrier (subring_generated complex_ring QinC)` THENL [ intro THEN rw[GSYM complex_ring_clauses] THEN rw[LAMBDA_PAIR] THEN simp[RING_MUL_POW;in_complex_ring] THEN rw[GSYM LAMBDA_PAIR] THEN simp[RING_BINOMIAL_THEOREM;in_complex_ring] THEN specialize_assuming[ `complex_ring`; `\(x,y) k. ring_mul complex_ring (ring_of_num complex_ring (binom (d,k))) (ring_mul complex_ring (ring_pow complex_ring x k) (ring_pow complex_ring y (d - k)))`; `complex_root p CROSS complex_root q`; `0..d` ]RING_SUM_SWAP THEN subgoal `ring_sum complex_ring (complex_root p CROSS complex_root q) (\(x,y). ring_sum complex_ring (0..d) (\k. ring_mul complex_ring (ring_of_num complex_ring (binom (d,k))) (ring_mul complex_ring (ring_pow complex_ring x k) (ring_pow complex_ring y (d - k))))) = ring_sum complex_ring (complex_root p CROSS complex_root q) (\i. ring_sum complex_ring (0..d) ((\(x,y) k. ring_mul complex_ring (ring_of_num complex_ring (binom (d,k))) (ring_mul complex_ring (ring_pow complex_ring x k) (ring_pow complex_ring y (d - k)))) i))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;LAMBDA_PAIR] ; pass ] THEN have `ring_sum complex_ring (complex_root p CROSS complex_root q) (\i. ring_sum complex_ring (0..d) ((\(x,y) k. ring_mul complex_ring (ring_of_num complex_ring (binom (d,k))) (ring_mul complex_ring (ring_pow complex_ring x k) (ring_pow complex_ring y (d - k)))) i)) = ring_sum complex_ring (0..d) (\j. ring_sum complex_ring (complex_root p CROSS complex_root q) (\i. (\(x,y) k. ring_mul complex_ring (ring_of_num complex_ring (binom (d,k))) (ring_mul complex_ring (ring_pow complex_ring x k) (ring_pow complex_ring y (d - k)))) i j))` [in_complex_ring;FINITE_NUMSEG] THEN simp[] THEN sufficesby ring_sum_in_subring THEN intro THEN rw[BETA_THM;LAMBDA_PAIR] THEN rw[GSYM LAMBDA_PAIR] THEN subgoal `ring_sum complex_ring (complex_root p CROSS complex_root q) (\(x,y). ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) (ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s)))) = ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) (ring_sum complex_ring (complex_root p CROSS complex_root q) (\(x,y). ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))))` THENL [ specialize_assuming[ `complex_ring`; `\(x,y). ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))`; `ring_of_num complex_ring (binom(d,s))`; `complex_root p CROSS complex_root q` ]RING_SUM_LMUL THEN have `ring_sum complex_ring (complex_root p CROSS complex_root q) (\x. ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) ((\(x,y). ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))) x)) = ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) (ring_sum complex_ring (complex_root p CROSS complex_root q) (\(x,y). ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))))` [in_complex_ring] THEN subgoal `ring_sum complex_ring (complex_root p CROSS complex_root q) (\x. ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) ((\(x,y). ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))) x)) = ring_sum complex_ring (complex_root p CROSS complex_root q) (\(x,y). ring_mul complex_ring (ring_of_num complex_ring (binom (d,s))) (ring_mul complex_ring (ring_pow complex_ring x s) (ring_pow complex_ring y (d - s))))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;LAMBDA_PAIR] ; pass ] THEN qed[] ; pass ] THEN simp[] THEN have `!x. x IN complex_root p ==> ring_pow complex_ring x s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!y. y IN complex_root q ==> ring_pow complex_ring y (d - s) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize_assuming[ `complex_ring`; `complex_root p`; `complex_root q`; `\x. ring_pow complex_ring x s`; `\y. ring_pow complex_ring y (d - s)` ]ring_sum_cross_mul THEN simp[] THEN rw[ring_of_num_complex;ring_pow_complex;complex_ring_clauses;GSYM complex_root_powersums;QinC_ring_clauses;subring_complex_QinC] THEN qed[mul_in_QinC;num_in_QinC;complex_of_num] ; pass ] THEN specialize_assuming[ `complex_ring`; `QinC`; `complex_root p CROSS complex_root q`; `\(y,z). y + z:complex` ]poly_subring_if_powersums_subring THEN qed[in_complex_ring;subring_complex_QinC] ; pass ] THEN have `algebraic_number ((\(y,z). y+z) (y,z))` [algebraic_number_if_monic_vanishing_at_QinC] THEN havetac `((\(y,z). y+z) (y,z)) = y+z:complex` (rw[]) THEN qed[] );; let algebraic_number_subring = prove(` algebraic_number subring_of complex_ring `, rw[subring_of] THEN rw[complex_ring_clauses;IN;SUBSET] THEN qed[IN_UNIV;IN;algebraic_number_ZinC;ZinC_0;ZinC_1;algebraic_number_neg;algebraic_number_add;algebraic_number_mul] );; let algebraic_number_sum = prove(` !f S. FINITE S ==> (!s:X. s IN S ==> algebraic_number (f s)) ==> algebraic_number (ring_sum complex_ring S f) `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;complex_ring_clauses] THEN qed[algebraic_number_ZinC;num_in_ZinC] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_SUM_CLAUSES;in_complex_ring] THEN rw[complex_ring_clauses] THEN qed[algebraic_number_add] ] );; (* ===== expformal c: the power series exp(cx) *) let expformal = new_definition ` expformal (c:complex) = series_from_coeffs (\n. c pow n / Cx(&(FACT n))) `;; let mul_expformal = prove(` !u v. poly_mul complex_ring (expformal u) (expformal v) = expformal (u+v) `, intro THEN sufficesby eq_coeff THEN intro THEN rw[coeff_poly_mul_oneindex;expformal;coeff_series_from_coeffs] THEN rw[GSYM ring_pow_complex;GSYM complex_ring_clauses] THEN simp[RING_BINOMIAL_THEOREM;in_complex_ring] THEN simp[GSYM vsum_ring_sum_complex;FINITE_NUMSEG] THEN rw[complex_div] THEN simp[GSYM VSUM_COMPLEX_RMUL;FINITE_NUMSEG] THEN rw[GSYM complex_div] THEN sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN have `x <= d:num` [IN_NUMSEG] THEN num_linear_fact `x <= d:num ==> (d-x)+x = d` THEN have `FACT(d-x) * FACT x * binom(d,x) = FACT d` [BINOM_FACT] THEN have `&(FACT(d-x)) * &(FACT x) * &(binom(d,x)) = &(FACT d):real` [REAL_OF_NUM_MUL] THEN have `Cx(&(FACT(d-x))) * Cx(&(FACT x)) * Cx(&(binom(d,x))) = Cx(&(FACT d))` [CX_MUL] THEN rw[complex_ring_clauses;ring_pow_complex] THEN rw[ring_of_num_complex;complex_of_num] THEN have `~(Cx(&(FACT d)) = Cx(&0))` [CX_INJ;REAL_OF_NUM_EQ;FACT_NZ] THEN complex_field_fact ` ~(Cx(&(FACT d)) = Cx(&0)) ==> Cx(&(FACT(d-x))) * Cx(&(FACT x)) * Cx(&(binom(d,x))) = Cx(&(FACT d)) ==> u pow x / Cx (&(FACT x)) * v pow (d - x) / Cx (&(FACT (d - x))) = (Cx (&(binom (d,x))) * u pow x * v pow (d - x)) / Cx (&(FACT d))` THEN qed[] );; let expformal_0 = prove(` expformal(Cx(&0)) = poly_1 complex_ring `, sufficesby eq_coeff THEN rw[expformal;coeff_series_from_coeffs;coeff_poly_1;complex_ring_clauses] THEN rw[COMPLEX_POW_ZERO] THEN intro THEN case `d = 0` THENL [ simp[FACT] THEN CONV_TAC COMPLEX_FIELD ; simp[] THEN CONV_TAC COMPLEX_FIELD ] );; let pow_expformal = prove(` !u n. poly_pow complex_ring (expformal u) n = expformal (Cx(&n)*u) `, GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;COMPLEX_MUL_LZERO;expformal_0] ; rw[ARITH_RULE `SUC n = n + 1`] THEN simp[poly_pow_add;series_complex;poly_pow_1] THEN simp[GSYM REAL_OF_NUM_ADD;CX_ADD;COMPLEX_ADD_RDISTRIB;COMPLEX_MUL_LID;mul_expformal] ] );; let product_expformal = prove(` !c S. FINITE S ==> poly_product complex_ring S (\s:X. expformal (c s)) = expformal (ring_sum complex_ring S c) `, GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[poly_product_empty;RING_SUM_CLAUSES] THEN rw[expformal_0;complex_ring_clauses] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[poly_product_insert;RING_SUM_CLAUSES;in_complex_ring;series_complex] THEN qed[mul_expformal;complex_ring_clauses] ] );; let product_expformal_I = prove(` !S. FINITE S ==> poly_product complex_ring S expformal = expformal (ring_sum complex_ring S I) `, simp[GSYM product_expformal] THEN simp[I_THM] THEN intro THEN sufficesby poly_product_eq THEN qed[] );; let expformal_powersums_QinC = prove(` !S c:X->complex. FINITE S ==> (!n. ring_sum complex_ring S (\s:X. (c s) pow n) IN QinC) ==> ring_powerseries QinC_ring (poly_sum complex_ring S (\s:X. expformal (c s))) `, GEN_TAC THEN GEN_TAC THEN DISCH_TAC THEN rw[coeff_series_in_ring] THEN simp[coeff_poly_sum;series_complex;expformal;coeff_series_from_coeffs;QinC_ring_clauses] THEN simp[GSYM vsum_ring_sum_complex] THEN rw[complex_div] THEN simp[VSUM_COMPLEX_RMUL] THEN rw[GSYM complex_div] THEN intro THEN have `~(FACT d = 0)` [FACT_NZ] THEN have `~(&(FACT d) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT d)) = Cx(&0))` [CX_INJ] THEN have `Cx(&(FACT d)) IN QinC` [num_in_QinC] THEN qed[div_in_QinC] );; let pow_expformal_powersums_QinC = prove(` !S m. FINITE S ==> (!n. ring_sum complex_ring S (\z. z pow n) IN QinC) ==> ring_powerseries QinC_ring ( poly_sum complex_ring S (\z. poly_pow complex_ring (expformal z) m ) ) `, intro THEN subgoal `!n. ring_sum complex_ring S (\z. (Cx(&m) * z) pow n) IN QinC` THENL [ rw[COMPLEX_POW_MUL] THEN simp[GSYM vsum_ring_sum_complex] THEN simp[VSUM_COMPLEX_LMUL] THEN simp[vsum_ring_sum_complex] THEN intro THEN have `Cx(&m) IN QinC` [num_in_QinC] THEN have `Cx(&m) pow n IN QinC` [ring_pow_QinC;RING_POW;QinC_ring_clauses] THEN qed[mul_in_QinC] ; pass ] THEN rw[pow_expformal] THEN specialize[ `S:complex->bool`; `\z:complex. Cx(&m) * z` ]expformal_powersums_QinC THEN qed[] );; let carrier_QinC_series_subring_generated_refl = prove(` ring_carrier(subring_generated (x_series complex_ring) (ring_carrier(x_series QinC_ring)) ) = ring_carrier(x_series QinC_ring) `, sufficesby CARRIER_SUBRING_GENERATED_SUBRING THEN rw[subring_of] THEN rw[GSYM x_series_use] THEN intro THENL [ rw[SUBSET;GSYM x_series_use] THEN qed[series_complex_if_series_QinC] ; qed[poly_0_QinC_eq_poly_0_complex;RING_POWERSERIES_0] ; qed[poly_1_QinC_eq_poly_1_complex;RING_POWERSERIES_1] ; qed[poly_neg_QinC_eq_poly_neg_complex;RING_POWERSERIES_NEG] ; qed[poly_add_QinC_eq_poly_add_complex;RING_POWERSERIES_ADD] ; qed[poly_mul_QinC_eq_poly_mul_complex;RING_POWERSERIES_MUL] ] );; let symfun_expformal_powersums_QinC = prove(` !S m. FINITE S ==> (!n. ring_sum complex_ring S (\z. z pow n) IN QinC) ==> poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. poly_product complex_ring U expformal) IN ring_carrier(x_series QinC_ring) `, intro THEN rw[poly_product] THEN rw[poly_sum] THEN recall ring_hasQ_subring_series_complex THEN have `!s. s IN S ==> expformal s IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN subgoal `!d. ring_sum (x_series complex_ring) S (\s. ring_pow (x_series complex_ring) (expformal s) d) IN ring_carrier (subring_generated (x_series complex_ring) (ring_carrier (x_series QinC_ring)))` THENL [ intro THEN have `ring_powerseries QinC_ring (poly_sum complex_ring S (\z. poly_pow complex_ring (expformal z) d))` [pow_expformal_powersums_QinC] THEN have `ring_sum(x_series complex_ring) S (\z. poly_pow complex_ring (expformal z) d) IN ring_carrier(x_series QinC_ring)` [x_series_use;poly_sum] THEN have `ring_carrier(x_series QinC_ring) SUBSET ring_carrier(x_series complex_ring)` [x_series_use;series_complex_if_series_QinC;SUBSET] THEN have `ring_carrier(x_series QinC_ring) SUBSET ring_carrier (subring_generated (x_series complex_ring) (ring_carrier(x_series QinC_ring)))` [SUBSET_CARRIER_SUBRING_GENERATED;SUBSET_REFL] THEN rw[GSYM x_series_use_pow] THEN qed[SUBSET] ; pass ] THEN specialize[ `x_series complex_ring`; `ring_carrier(x_series QinC_ring)`; `S:complex->bool`; `expformal`; `m:num` ]symfun_subring_if_powersums_subring THEN qed[carrier_QinC_series_subring_generated_refl] );; let symfun_expformal_powersums_QinC_v2 = prove(` !S m. FINITE S ==> (!n. ring_sum complex_ring S (\z. z pow n) IN QinC) ==> poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. expformal (ring_sum complex_ring U I)) IN ring_carrier(x_series QinC_ring) `, intro THEN subgoal `poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. poly_product complex_ring U expformal) = poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. expformal (ring_sum complex_ring U I))` THENL [ sufficesby poly_sum_eq THEN rw[IN_ELIM_THM] THEN intro THEN qed[FINITE_SUBSET;product_expformal_I] ; pass ] THEN qed[symfun_expformal_powersums_QinC] );; let symfun_expformal_powersums_QinC_v3 = prove(` !S m d. FINITE S ==> (!n. ring_sum complex_ring S (\z. z pow n) IN QinC) ==> ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d) IN QinC `, intro THEN have `poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. expformal (ring_sum complex_ring U I)) IN ring_carrier(x_series QinC_ring)` [symfun_expformal_powersums_QinC_v2] THEN have `coeff d (poly_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. expformal (ring_sum complex_ring U I))) IN QinC` [x_series_use;coeff_series_in_ring;QinC_ring_clauses] THEN specialize_assuming[ `complex_ring`; `\U. expformal (ring_sum complex_ring U I)`; `d:num`; `{U:complex->bool | U SUBSET S /\ CARD U = m}` ]coeff_poly_sum THEN have `FINITE {U:complex->bool | U SUBSET S /\ CARD U = m}` [finite_subsets_card] THEN have `ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. coeff d (expformal (ring_sum complex_ring U I))) IN QinC` [series_complex] THEN subgoal `ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. coeff d (expformal (ring_sum complex_ring U I))) = ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d / Cx (&(FACT d)))` THENL [ sufficesby RING_SUM_EQ THEN qed[coeff_series_from_coeffs;expformal] ; pass ] THEN have `ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d / Cx (&(FACT d))) IN QinC` [] THEN have `Cx(&(FACT d)) IN QinC` [num_in_QinC] THEN have `Cx(&(FACT d)) * ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d / Cx (&(FACT d))) IN QinC` [mul_in_QinC] THEN have `Cx(&(FACT d)) * vsum {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d / Cx (&(FACT d))) IN QinC` [vsum_ring_sum_complex] THEN subgoal `Cx(&(FACT d)) * vsum {U | U SUBSET S /\ CARD U = m} (\U. (ring_sum complex_ring U I) pow d / Cx (&(FACT d))) = vsum {U | U SUBSET S /\ CARD U = m} (\U. Cx(&(FACT d)) * ((ring_sum complex_ring U I) pow d / Cx (&(FACT d))))` THENL [ simp[VSUM_COMPLEX_LMUL] ; pass ] THEN have `vsum {U | U SUBSET S /\ CARD U = m} (\U. Cx(&(FACT d)) * ((ring_sum complex_ring U I) pow d / Cx (&(FACT d)))) IN QinC` [] THEN subgoal `vsum {U | U SUBSET S /\ CARD U = m} (\U. Cx (&(FACT d)) * ring_sum complex_ring U I pow d / Cx (&(FACT d))) = vsum {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I pow d)` THENL [ sufficesby VSUM_EQ THEN intro THEN rw[BETA_THM] THEN have `~(FACT d = 0)` [FACT_NZ] THEN have `~(&(FACT d) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT d)) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx(&(FACT d)) = Cx(&0)) ==> Cx (&(FACT d)) * ring_sum complex_ring x I pow d / Cx (&(FACT d)) = ring_sum complex_ring x I pow d` THEN qed[] ; pass ] THEN qed[vsum_ring_sum_complex] );; let resolvent_if_powersums_QinC = prove(` !S m. FINITE S ==> (!n. ring_sum complex_ring S (\z. z pow n) IN QinC) ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I) ) `, intro THEN have `ring_hasQ (subring_generated complex_ring QinC)` [ring_hasQ_QinC;subring_complex_QinC] THEN specialize[`S:complex->bool`;`m:num`]finite_subsets_card THEN have `!s. s IN {U | U SUBSET S /\ CARD U = m} ==> ring_sum complex_ring s I IN ring_carrier complex_ring` [in_complex_ring] THEN subgoal `!d. ring_sum complex_ring {U | U SUBSET S /\ CARD U = m} (\s. ring_pow complex_ring (ring_sum complex_ring s I) d) IN ring_carrier (subring_generated complex_ring QinC)` THENL [ rw[ring_pow_complex] THEN qed[symfun_expformal_powersums_QinC_v3;QinC_ring_clauses;subring_complex_QinC] ; pass ] THEN specialize[ `complex_ring`; `QinC`; `{U:complex->bool | U SUBSET S /\ CARD U = m}`; `\U. ring_sum complex_ring U I` ]poly_subring_if_powersums_subring THEN qed[subring_complex_QinC] );; (* can alternatively prove this via symmetric-function argument *) let resolvent_if_poly_QinC = prove(` !S m. FINITE S ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring S I ) ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I) ) `, intro THEN subgoal `!n. ring_sum complex_ring S (\z. z pow n) IN QinC` THENL [ intro THEN specialize_assuming[ `complex_ring`; `QinC`; `S:complex->bool`; `I:complex->complex`; `n:num` ]powersums_subring_if_poly_subring THEN have `!n. coeff n (monic_vanishing_at complex_ring S I) IN ring_carrier(subring_generated complex_ring QinC)` [subring_complex_QinC;coeff_poly_in_ring] THEN have `ring_sum complex_ring S (\s. ring_pow complex_ring (I s) n) IN ring_carrier (subring_generated complex_ring QinC)` [in_complex_ring] THEN subgoal `ring_sum complex_ring S (\s. ring_pow complex_ring (I s) n) = ring_sum complex_ring S (\z. z pow n)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;I_THM;ring_pow_complex] ; pass ] THEN qed[subring_complex_QinC;QinC_ring_clauses] ; pass ] THEN qed[resolvent_if_powersums_QinC] );; let prod_resolvent_if_poly_QinC_lemma = prove(` !S. FINITE S ==> monic_vanishing_at complex_ring {U | U SUBSET S} (\U. ring_sum complex_ring U I) = poly_product complex_ring (0..CARD S) (\m. monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I) ) `, intro THEN rw[poly_product;monic_vanishing_at] THEN have `IMAGE CARD {U:complex->bool | U SUBSET S} = (0..CARD S)` [image_card_powerset] THEN have `FINITE {A | A SUBSET S:complex->bool}` [FINITE_POWERSET] THEN specialize[ `x_series complex_ring`; `CARD:(complex->bool)->num`; `\s. x_minus_const complex_ring (ring_sum complex_ring s I)`; `{U:complex->bool | U SUBSET S}` ]RING_PRODUCT_IMAGE_GEN THEN simp[] THEN sufficesby RING_PRODUCT_EQ THEN intro THEN rw[BETA_THM] THEN set_fact `{x:complex->bool | x IN {U | U SUBSET S} /\ CARD x = a} = {U | U SUBSET S /\ CARD U = a}` THEN qed[] );; let prod_resolvent_if_poly_QinC = prove(` !S. FINITE S ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring S I ) ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring {U | U SUBSET S} (\U. ring_sum complex_ring U I) ) `, intro THEN simp[prod_resolvent_if_poly_QinC_lemma] THEN have `FINITE (0..CARD(S:complex->bool))` [FINITE_NUMSEG] THEN have `!m. m IN 0..CARD S ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I))` [resolvent_if_poly_QinC] THEN have `poly_product complex_ring (0..CARD S) (\m. monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I)) = poly_product QinC_ring (0..CARD S) (\m. monic_vanishing_at complex_ring {U | U SUBSET S /\ CARD U = m} (\U. ring_sum complex_ring U I))` [subring_complex_QinC;poly_product_subring;ring_polynomial] THEN simp[] THEN qed[poly_product_poly] );; (* ===== pi is transcendental *) let pi_is_transcendental = prove(` ~(algebraic_number(Cx pi)) `, intro THEN recall algebraic_number_ii THEN have `algebraic_number (ii * Cx pi)` [algebraic_number_mul] THEN choose `f:(1->num)->complex` `ring_polynomial QinC_ring f /\ ~(f = poly_0 QinC_ring) /\ monic QinC_ring f /\ ring_irreducible(x_poly QinC_ring) f /\ poly_eval complex_ring f (ii * Cx pi) = Cx(&0)` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN recall CEXP_II_PI THEN complex_field_fact `cexp (ii * Cx pi) = --Cx (&1) ==> Cx(&1) + cexp(ii * Cx pi) = Cx(&0)` THEN have `complex_root f (ii * Cx pi)` [complex_root] THEN have `FINITE (complex_root f)` [distinct_minpolys_finite_root_simple] THEN subgoal `ring_product complex_ring (complex_root f) (\z. Cx(&1) + cexp z) = Cx(&0)` THENL [ specialize_assuming[ `complex_ring`; `complex_root f`; `ii * Cx pi`; `\z. Cx(&1) + cexp z` ]ring_product_delete THEN have `ring_product complex_ring (complex_root f) (\z. Cx (&1) + cexp z) = ring_mul complex_ring (Cx (&1) + cexp (ii * Cx pi)) (ring_product complex_ring (complex_root f DELETE (ii * Cx pi)) (\z. Cx (&1) + cexp z))` [IN;in_complex_ring] THEN simp[complex_ring_clauses;COMPLEX_MUL_LZERO] ; pass ] THEN subgoal `ring_product complex_ring (complex_root f) (\z. Cx(&1) + cexp z) = ring_sum complex_ring {S | S SUBSET complex_root f} (\S. ring_product complex_ring S cexp)` THENL [ rw[GSYM complex_ring_clauses] THEN specialize_assuming[ `complex_ring`; `cexp`; `complex_root f` ]ring_product_1_plus_expand THEN qed[in_complex_ring] ; pass ] THEN have `ring_sum complex_ring {S | S SUBSET complex_root f} (\S. ring_product complex_ring S cexp) = Cx(&0)` [] THEN subgoal `ring_sum complex_ring {S | S SUBSET complex_root f} (\S. ring_product complex_ring S cexp) = ring_sum complex_ring {S | S SUBSET complex_root f} (\S. cexp (ring_sum complex_ring S I))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN have `FINITE(a:complex->bool)` [FINITE_SUBSET] THEN simp[GSYM vsum_ring_sum_complex;GSYM cproduct_ring_product_complex;CEXP_VSUM;I_THM] ; pass ] THEN have `ring_sum complex_ring {S | S SUBSET complex_root f} (\S. cexp (ring_sum complex_ring S I)) = Cx(&0)` [] THEN have `monic_vanishing_at complex_ring (complex_root f) I = f` [distinct_minpolys_monic_vanishing_at_simple] THEN have `ring_polynomial QinC_ring (monic_vanishing_at complex_ring (complex_root f) I)` [] THEN specialize[`complex_root f`]prod_resolvent_if_poly_QinC THEN have `ring_polynomial complex_ring (monic_vanishing_at complex_ring (complex_root f) I)` [monic_vanishing_at_poly;in_complex_ring] THEN have `~(monic_vanishing_at complex_ring (complex_root f) I = poly_0 complex_ring)` [monic_vanishing_at_monic;monic_poly_0;ring_1_0_complex;in_complex_ring] THEN have `FINITE(complex_root f)` [complex_root_le_deg] THEN specialize[`complex_root f`]FINITE_POWERSET THEN specialize[ `{U | U SUBSET complex_root f}`; `\U. ring_sum complex_ring U I` ]transcendence_all_0_QinC_monic_vanishing_at THEN ASM SET_TAC[] );; (* ===== more on monomials *) let monomial_pow = new_definition ` monomial_pow (m:V->num) (e:num) = (\v. e * m v) `;; let monomial_product = new_definition ` monomial_product (S:X->bool) (m:X->V->num) = (\v. nsum S (\x. m x v)) `;; let finite_monomial_vars_permutation = prove(` !m:V->num f:V->V. f permutes (:V) ==> ( FINITE(monomial_vars (m o f)) <=> FINITE(monomial_vars m) ) `, rw[monomial_vars] THEN rw[permutes;IN_UNIV] THEN intro THEN subgoal `{i | ~(m i = 0)} = IMAGE (f:V->V) {i | ~((m o f) i = 0)}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM;o_DEF] THEN qed[] ; pass ] THEN choose `g:V->V` `(!y:V. f(g y) = y) /\ (!x:V. g(f x) = x)` [BIJECTIVE_LEFT_RIGHT_INVERSE] THEN subgoal `{i | ~((m o f) i = 0)} = IMAGE (g:V->V) {i | ~(m i = 0)}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM;o_DEF] THEN qed[] ; pass ] THEN qed[FINITE_IMAGE] );; let finite_monomial_vars_swap = prove(` !m:num->num i. FINITE(monomial_vars (m o swap(i,i+1))) <=> FINITE(monomial_vars m) `, qed[PERMUTES_SWAP;IN_UNIV;finite_monomial_vars_permutation] );; let monomial_eq_swap = prove(` !m:V->num i j. m o swap(i,j) = m <=> m i = m j `, rw[swap;o_DEF;FUN_EQ_THM] THEN qed[] );; let monomial_le_swap = prove(` !m:num->num i. FINITE(monomial_vars m) ==> ( monomial_le (<=) (m o swap(i,i+1)) m <=> m i <= m(i+1) ) `, intro THEN simp[MONOMIAL_LE_TOSET;TOSET_num;FLD_num] THEN simp[monomial;SUBSET_UNIV;finite_monomial_vars_swap] THEN rw[monomial_eq_swap] THEN rw[properly_le;swap;o_DEF] THEN splitiff THENL [ intro THENL [ ASM_ARITH_TAC ; case `i' < i+1` THENL [ num_linear_fact `~(i+1 = i)` THEN have `m i = m(i+1):num` [] THEN qed[LE_REFL] ; pass ] THEN case `i' = i+1` THENL [ num_linear_fact `~(i+1 = i)` THEN have `m(i:num):num < m(i+1)` [] THEN qed[LT_IMP_LE] ; pass ] THEN num_linear_fact `~(i' < i+1) ==> ~(i' = i+1) ==> ~(i' = i)` THEN qed[LT_REFL] ] ; intro THEN proven_if `m(i):num = m(i+1)` [] THEN num_linear_fact `~(m(i):num = m(i+1)) ==> m i <= m(i+1) ==> m i < m(i+1)` THEN simp[] THEN witness `i+1` THEN num_linear_fact `~(i + 1 = i)` THEN simp[] THEN intro THEN num_linear_fact `i+1 < j ==> ~(j = i)` THEN num_linear_fact `i+1 < j ==> ~(j = i+1)` THEN qed[] ] );; let monomial_induction = prove(` !P:(num->num)->bool. (!n. (!m. monomial_lt (<=) m n ==> P m) ==> P n) ==> (!n. P n) `, once_rw[GSYM MONOMIAL_LT_PROPERLY] THEN simp[properly_le] THEN intro THEN have `WF ((<):num->num->bool)` [WF_num] THEN have `WF (monomial_lt ((<):num->num->bool))` [WF_MONOMIAL_LT] THEN qed[WF_IND] );; let monomial_le_mul2_eq = prove(` !m n M N:num->num. monomial_le (<=) m M ==> monomial_le (<=) n N ==> monomial_mul m n = monomial_mul M N ==> ( m = M /\ n = N ) `, REPEAT GEN_TAC THEN REPEAT DISCH_TAC THEN case `~(m = M:num->num)` THENL [ have `monomial_lt (<=) m (M:num->num)` [monomial_lt] THEN have `monomial_lt (<=) (monomial_mul m n) (monomial_mul M N:num->num)` [MONOMIAL_LTE_MUL2;POSET_num] THEN qed[monomial_lt] ; pass ] THEN case `~(n = N:num->num)` THENL [ have `monomial_lt (<=) n (N:num->num)` [monomial_lt] THEN have `monomial_lt (<=) (monomial_mul m n) (monomial_mul M N:num->num)` [MONOMIAL_LET_MUL2;POSET_num] THEN qed[monomial_lt] ; pass ] THEN qed[] );; let monomial_pow_0 = prove(` !m:V->num. monomial_pow m 0 = monomial_1 `, rw[monomial_pow;monomial_1;FUN_EQ_THM] THEN ARITH_TAC );; let monomial_pow_1 = prove(` !m:V->num. monomial_pow m 1 = m `, rw[monomial_pow;FUN_EQ_THM] THEN ARITH_TAC );; let monomial_pow_add = prove(` !m:V->num d e. monomial_pow m (d+e) = monomial_mul (monomial_pow m d) (monomial_pow m e) `, rw[monomial_pow;monomial_mul] THEN rw[RIGHT_ADD_DISTRIB] );; let monomial_1_le = prove(` !m. monomial (:num) m ==> monomial_le (<=) monomial_1 m `, simp[MONOMIAL_LE_TOSET;TOSET_num;FLD_num;MONOMIAL_1] THEN rw[properly_le;monomial_1;monomial;SUBSET_UNIV;monomial_vars] THEN intro THEN case `{i:num | ~(m i = 0)} = {}` THENL [ DISJ1_TAC THEN rw[FUN_EQ_THM] THEN ASM SET_TAC[] ; DISJ2_TAC THEN pass ] THEN specialize[`{i:num | ~(m i = 0)}`]max_finite THEN choose `i:num` `i IN {i:num | ~(m i = 0)} /\ (!j. i < j ==> ~(j IN {i | ~(m i = 0)}))` [] THEN witness `i:num` THEN intro THENL [ have `i IN {i:num | ~(m i = 0)}` [] THEN set_fact `i IN {i:num | ~(m i = 0)} ==> ~(m i = 0)` THEN qed[ARITH_RULE `~(x = 0) ==> 0 < x`] ; have `~(j IN {i:num | ~(m i = 0)})` [] THEN set_fact `~(j IN {i:num | ~(m i = 0)}) ==> m j = 0` THEN qed[] ] );; let monomial_product_empty = prove(` monomial_product ({}:X->bool) (m:X->V->num) = monomial_1 `, rw[monomial_product;monomial_1] THEN rw[NSUM_CLAUSES] );; let monomial_product_insert = prove(` !(S:X->bool) (m:X->V->num) t. FINITE S ==> ~(t IN S) ==> monomial_product (t INSERT S) m = monomial_mul (m t) (monomial_product S m) `, rw[monomial_product;monomial_mul] THEN simp[NSUM_CLAUSES] );; (* ===== support *) let support_lt = new_definition ` support_lt (r:R ring) (p:(num->num)->R) (M:num->num) <=> (!m. ~(p m = ring_0 r) ==> monomial_lt (<=) m M) `;; let support_le = new_definition ` support_le (r:R ring) (p:(num->num)->R) (M:num->num) <=> (!m. ~(p m = ring_0 r) ==> monomial_le (<=) m M) `;; let support_le1 = new_definition ` support_le1 (r:R ring) (p:(num->num)->R) (M:num->num) <=> ( support_le r p M /\ p M = ring_1 r ) `;; (* POLY_TOP_MONOMIAL_LE is deg first *) let poly_first_monomial = prove(` !(r:R ring) p:(num->num)->R. ring_polynomial r p ==> ~(p = poly_0 r) ==> (?t. ~(p t = ring_0 r) /\ support_le r p t) `, intro THEN rw[support_le] THEN def `M:(num->num)->bool` `{m:num->num | ~(p m = ring_0(r:R ring))}` THEN have `FINITE (M:(num->num)->bool)` [ring_polynomial] THEN subgoal `~(M = {}:(num->num)->bool)` THENL [ rw[EMPTY;EXTENSION;IN_ELIM_THM] THEN intro THEN subgoal `p = poly_0 r:(num->num)->R` THENL [ rw[poly_0;poly_const;FUN_EQ_THM;COND_ID] THEN ASM SET_TAC[] ; pass ] THEN qed[] ; pass ] THEN have `toset (monomial_le ((<=):num->num->bool))` [TOSET_MONOMIAL_LE;FLD_num;TOSET_num] THEN subgoal `M SUBSET fld (monomial_le ((<=):num->num->bool))` THENL [ simp[FLD_MONOMIAL_LE;monomial;SUBSET;IN_ELIM_THM;IN_UNIV] THEN qed[ring_polynomial;ring_powerseries;INFINITE] ; pass ] THEN specialize[ `monomial_le ((<=):num->num->bool)`; `M:(num->num)->bool` ]TOSET_MAX THEN choose `t:num->num` `t IN M /\ (!m:num->num. m IN M ==> monomial_le (<=) m t)` [] THEN witness `t:num->num` THEN ASM SET_TAC[] );; let support_le_0 = prove(` !(r:R ring). support_le r (poly_0 r) monomial_1 `, rw[support_le;poly_0;poly_const;COND_ID] );; let support_le_exists = prove(` !(r:R ring) p:(num->num)->R. ring_polynomial r p ==> ?t. support_le r p t `, qed[support_le_0;poly_first_monomial] );; let support_le_mul = prove(` !(r:R ring) p q M N. ring_polynomial r p ==> ring_polynomial r q ==> support_le r p M ==> support_le r q N ==> support_le r (poly_mul r p q) (monomial_mul M N) `, rw[support_le;poly_mul] THEN intro THEN case `!a. a IN {m1,m2 | monomial_mul m1 m2 = m:num->num} ==> (\(m1,m2). ring_mul r (p m1) (q m2)) a IN ring_carrier r ==> (\(m1,m2). ring_mul r (p m1) (q m2)) a = ring_0(r:R ring)` THENL [ specialize[ `r:R ring`; `(\(m1,m2). ring_mul(r:R ring) (p(m1:num->num)) (q(m2:num->num)))`; `{m1,m2 | monomial_mul m1 m2 = m:num->num}` ]RING_SUM_EQ_0 THEN qed[] ; pass ] THEN choose `a:(num->num)#(num->num)` `a IN {m1,m2 | monomial_mul m1 m2 = m:num->num} /\ (\(m1,m2). ring_mul r (p m1) (q m2)) a IN ring_carrier r /\ ~((\(m1,m2). ring_mul r (p m1) (q m2)) a = ring_0(r:R ring))` [] THEN subgoal `(FST a,SND a) IN {m1,m2 | monomial_mul m1 m2 = m:num->num} <=> monomial_mul (FST a) (SND a) = m` THENL [ rw[IN_ELIM_PAIR_THM] ; pass ] THEN have `a IN {m1,m2 | monomial_mul m1 m2 = m:num->num}` [] THEN have `(FST a,SND a) IN {m1,m2 | monomial_mul m1 m2 = m:num->num}` [PAIR] THEN have `monomial_mul (FST a) (SND a) = m:num->num` [] THEN subgoal `(\(m1,m2). ring_mul r (p(m1:num->num)) (q(m2:num->num))) a = ring_mul(r:R ring) (p (FST a)) (q (SND a))` THENL [ rw[LAMBDA_PAIR] ; pass ] THEN have `~(ring_mul(r:R ring) (p (FST a:num->num)) (q (SND a:num->num)) = ring_0 r)` [] THEN have `p (FST(a:(num->num)#(num->num))) IN ring_carrier(r:R ring)` [ring_polynomial;ring_powerseries] THEN have `q (SND(a:(num->num)#(num->num))) IN ring_carrier(r:R ring)` [ring_polynomial;ring_powerseries] THEN have `~(p (FST(a:(num->num)#(num->num))) = ring_0(r:R ring))` [RING_MUL_LZERO] THEN have `~(q (SND(a:(num->num)#(num->num))) = ring_0(r:R ring))` [RING_MUL_RZERO] THEN have `monomial_le (<=) (FST(a:(num->num)#(num->num))) M` [] THEN have `monomial_le (<=) (SND(a:(num->num)#(num->num))) N` [] THEN qed[MONOMIAL_LE_MUL2] );; let support_le_mul_top = prove(` !(r:R ring) p q M N. ring_polynomial r p ==> ring_polynomial r q ==> support_le r p M ==> support_le r q N ==> (poly_mul r p q) (monomial_mul M N) = ring_mul r (p M) (q N) `, rw[support_le;poly_mul] THEN intro THEN subgoal `ring_sum(r:R ring) {m1,m2 | monomial_mul m1 m2 = monomial_mul M N:num->num} (\(m1,m2). ring_mul r (p m1) (q m2)) = ring_sum r {m1,m2 | monomial_mul m1 m2 = monomial_mul M N} (\a. if a = M,N then ring_mul r (p (FST a)) (q (SND a)) else ring_0 r)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;FORALL_PAIR_THM;PAIR_EQ;IN_ELIM_PAIR_THM] THEN intro THEN proven_if `p1 = M:num->num /\ p2 = N:num->num` [] THEN case `p(p1:num->num) = ring_0(r:R ring)` THENL [ have `q(p2:num->num) IN ring_carrier(r:R ring)` [ring_polynomial;ring_powerseries] THEN qed[RING_MUL_LZERO] ; pass ] THEN case `q(p2:num->num) = ring_0(r:R ring)` THENL [ have `p(p1:num->num) IN ring_carrier(r:R ring)` [ring_polynomial;ring_powerseries] THEN qed[RING_MUL_RZERO] ; pass ] THEN have `monomial_le (<=) p1 (M:num->num)` [] THEN have `monomial_le (<=) p2 (N:num->num)` [] THEN qed[monomial_le_mul2_eq] ; pass ] THEN simp[RING_SUM_DELTA;IN_ELIM_PAIR_THM] THEN qed[RING_MUL;ring_polynomial;ring_powerseries] );; let support_le1_mul = prove(` !(r:R ring) p q M N. ring_polynomial r p ==> ring_polynomial r q ==> support_le1 r p M ==> support_le1 r q N ==> support_le1 r (poly_mul r p q) (monomial_mul M N) `, rw[support_le1] THEN qed[support_le_mul;support_le_mul_top;RING_1;RING_MUL_LID] );; let support_le_poly_1 = prove(` !(r:R ring) m. monomial (:num) m ==> support_le r (poly_1 r) m `, rw[support_le;poly_1;poly_const] THEN qed[monomial_1_le] );; let support_le1_poly_1 = prove(` !(r:R ring). support_le1 r (poly_1 r) monomial_1 `, simp[support_le1;support_le_poly_1;MONOMIAL_1] THEN rw[poly_1;poly_const] );; let support_le1_pow = prove(` !(r:R ring) p M e. ring_polynomial r p ==> support_le1 r p M ==> support_le1 r (poly_pow r p e) (monomial_pow M e) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;monomial_pow_0;support_le1_poly_1] ; rw[ARITH_RULE `SUC e = e+1`] THEN intro THEN have `ring_powerseries r (p:(num->num)->R)` [ring_polynomial] THEN simp[poly_pow_add;monomial_pow_add] THEN simp[poly_pow_1;monomial_pow_1] THEN qed[support_le1_mul;poly_pow_poly] ] );; let support_le1_product = prove(` !(r:R ring) p M S. FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (p s)) ==> (!s:X. s IN S ==> support_le1 r (p s) (M s)) ==> support_le1 r (ring_product(poly_ring r (:num)) S p) (monomial_product S M) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;monomial_product_empty] THEN rw[POLY_RING;support_le1_poly_1] ; subgoal `!s:X. s IN x INSERT S ==> p s IN ring_carrier(poly_ring(r:R ring) (:num))` THENL [ rw[POLY_RING;SUBSET_UNIV;IN_ELIM_THM] THEN qed[] ; pass ] THEN set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES;monomial_product_insert] THEN rw[POLY_RING] THEN subgoal `ring_polynomial r (ring_product(poly_ring(r:R ring) (:num)) (S:X->bool) p)` THENL [ have `ring_product(poly_ring(r:R ring) (:num)) (S:X->bool) p IN ring_carrier(poly_ring r (:num))` [RING_PRODUCT] THEN specialize[`r:R ring`;`(:num)`](CONJUNCT1 POLY_RING) THEN have `ring_product(poly_ring(r:R ring) (:num)) (S:X->bool) p IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:num)}` [] THEN set_fact `ring_product(poly_ring(r:R ring) (:num)) (S:X->bool) p IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:num)} ==> ring_polynomial r (ring_product(poly_ring(r:R ring) (:num)) (S:X->bool) p)` THEN qed[] ; pass ] THEN qed[support_le1_mul] ] );; let support_le_cancel = prove(` !(r:R ring) p q M. ring_polynomial r p ==> ring_polynomial r q ==> support_le r p M ==> support_le1 r q M ==> support_lt r (poly_sub r p (poly_mul r (poly_const r (p M)) q)) M `, rw[support_le;support_le1;support_lt;poly_sub] THEN intro_genonly THEN DISCH_TAC THEN DISCH_TAC THEN subgoal `poly_mul(r:R ring) (poly_const r (p(M:num->num))) q = (\d:num->num. ring_mul r (p M) (q d))` THENL [ have `ring_powerseries r (q:(num->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (p:(num->num)->R)` [ring_polynomial] THEN have `p(M:num->num) IN ring_carrier(r:R ring)` [ring_powerseries] THEN specialize[ `r:R ring`; `p(M:num->num):R`; `q:(num->num)->R` ]poly_const_times THEN qed[] ; pass ] THEN simp[] THEN intro THEN case `m = M:num->num` THENL [ have `q(m:num->num) = ring_1(r:R ring)` [] THEN have `ring_mul(r:R ring) (p(M:num->num)) (q(m:num->num)) = p M` [RING_MUL_RID;ring_polynomial;ring_powerseries] THEN have `ring_sub r (p m) (ring_mul(r:R ring) (p(M:num->num)) (q(m:num->num))) = ring_0 r` [RING_SUB_REFL;ring_polynomial;ring_powerseries] THEN qed[] ; pass ] THEN proven_if `monomial_le (<=) m (M:num->num)` [monomial_lt] THEN have `p(m:num->num) = ring_0(r:R ring)` [] THEN have `q(m:num->num) = ring_0(r:R ring)` [] THEN have `p(M:num->num) IN ring_carrier(r:R ring)` [ring_polynomial;ring_powerseries] THEN have `ring_mul(r:R ring) (p(M:num->num)) (q(m:num->num)) = ring_0 r` [RING_MUL_RZERO] THEN have `ring_sub r (ring_0 r) (ring_0(r:R ring)) = ring_0 r` [RING_SUB_REFL;RING_0] THEN have `ring_sub r (p(m:num->num)) (ring_0(r:R ring)) = ring_0 r` [] THEN have `ring_sub r (p m) (ring_mul(r:R ring) (p(M:num->num)) (q(m:num->num))) = ring_0 r` [] THEN qed[] );; (* ===== symmetric functions *) let elementary_sympoly_range = new_definition ` elementary_sympoly_range (r:R ring) n d = \m. if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r `;; let elementary_sympoly_range_subring = prove(` !(r:R ring) G n d. elementary_sympoly_range(subring_generated r G) n d = elementary_sympoly_range r n d `, rw[elementary_sympoly_range] THEN rw[SUBRING_GENERATED] );; let powerseries_elementary_sympoly_range = prove(` !(r:R ring) n d. ring_powerseries r (elementary_sympoly_range r n d) `, rw[ring_powerseries;elementary_sympoly_range;monomial_vars;INFINITE] THEN intro THENL [ qed[RING_1;RING_0] ; subgoal `~(?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0))` THENL [ rw[NOT_EXISTS_THM] THEN intro THEN subgoal `{i:num | ~(m i = 0)} = U` THENL [ simp[] THEN SET_TAC[ARITH_RULE `~(1 = 0)`] ; pass ] THEN qed[FINITE_SUBSET;finite_range] ; pass ] THEN qed[] ] );; let polynomial_elementary_sympoly_range = prove(` !(r:R ring) n d. ring_polynomial r (elementary_sympoly_range r n d) `, rw[ring_polynomial] THEN rw[powerseries_elementary_sympoly_range] THEN rw[elementary_sympoly_range] THEN intro THEN case `ring_1(r:R ring) = ring_0 r` THENL [ simp[COND_ID] THEN qed[FINITE_EMPTY;EMPTY_GSPEC] ; pass ] THEN subgoal `{m | ~((if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r) = ring_0(r:R ring))} = IMAGE (\U i. if i IN U then 1 else 0) {U | U SUBSET range n /\ CARD U = d}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[] THEN have `FINITE(range n)` [finite_range] THEN specialize[`range n`;`d:num`]finite_subsets_card THEN qed[FINITE_IMAGE] );; let elementary_sympoly_range_in_poly_ring = prove(` !(r:R ring) n d. elementary_sympoly_range r n d IN ring_carrier(poly_ring r (range n)) `, rw[POLY_RING;IN_ELIM_THM] THEN simp[polynomial_elementary_sympoly_range] THEN rw[poly_vars;UNIONS_SUBSET;IN_ELIM_THM;monomial_vars] THEN rw[SUBSET] THEN rw[elementary_sympoly_range] THEN intro THEN choose `U:num->bool` `U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0)` [] THEN set_fact `s = {i:num | ~(m i = 0)} ==> x IN s ==> ~(m x = 0)` THEN have `x:num IN U` [] THEN qed[SUBSET] );; let eval_elementary_sympoly_range = prove(` !(r:R ring) n d c. (!i. i < n ==> c i IN ring_carrier r) ==> poly_evaluate r (elementary_sympoly_range r n d) c = ring_sum r {U | U SUBSET range n /\ CARD U = d} (\U. ring_product r U c) `, intro THEN rw[poly_evaluate;poly_extend;elementary_sympoly_range;I_THM] THEN case `ring_1(r:R ring) = ring_0 r` THENL [ have `trivial_ring(r:R ring)` [TRIVIAL_RING_10] THEN have `ring_carrier(r:R ring) = {ring_0 r}` [trivial_ring] THEN have `ring_sum(r:R ring) {U | U SUBSET range n /\ CARD U = d} (\U. ring_product r U c) = ring_0 r` [RING_SUM;IN_SING] THEN qed[RING_SUM;IN_SING] ; pass ] THEN subgoal `{m | ~((if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r) = ring_0(r:R ring))} = IMAGE (\U i. if i IN U then 1 else 0) {U | U SUBSET range n /\ CARD U = d}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM] THEN intro THEN qed[] ; pass ] THEN simp[] THEN subgoal `ring_sum(r:R ring) (IMAGE (\U i. if i IN U then 1 else 0) {U | U SUBSET range n /\ CARD U = d}) (\m. ring_mul r (if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r) (ring_product r (monomial_vars m) (\i. ring_pow r (c i) (m i)))) = ring_sum r {U | U SUBSET range n /\ CARD U = d} ((\m. ring_mul r (if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r) (ring_product r (monomial_vars m) (\i. ring_pow r (c i) (m i)))) o (\U i. if i IN U then 1 else 0))` THENL [ subgoal `!X Y. X IN {U | U SUBSET range n /\ CARD U = d} ==> Y IN {U | U SUBSET range n /\ CARD U = d} ==> (\U i. if i IN U then 1 else 0) X = (\U i. if i IN U then 1 else 0) Y ==> X = Y` THENL [ rw[IN_ELIM_THM;EXTENSION;FUN_EQ_THM;IN] THEN num_linear_fact `~(1 = 0)` THEN qed[] ; pass ] THEN specialize_assuming[ `r:R ring`; `\U i:num. if i IN U then 1 else 0`; `(\m. ring_mul(r:R ring) (if ?U. U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0) then ring_1 r else ring_0 r) (ring_product r (monomial_vars m) (\i. ring_pow r (c i) (m i))))`; `{U | U SUBSET range n /\ CARD U = d}` ]RING_SUM_IMAGE THEN qed[] ; pass ] THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;o_THM;IN_ELIM_THM;monomial_vars] THEN intro THEN have `?U. U SUBSET range n /\ CARD U = d /\ (\i. if i IN a then 1 else 0) = (\i. if i IN U then 1 else 0)` [] THEN subgoal `{i:num | ~((if i IN a then 1 else 0) = 0)} = a` THENL [ SET_TAC[ARITH_RULE `~(1 = 0)`] ; pass ] THEN simp[RING_MUL_LID;RING_PRODUCT] THEN sufficesby RING_PRODUCT_EQ THEN qed[RING_POW_1;SUBSET;range_lt] );; let eval_elementary_sympoly_range_coeff = prove(` !(r:R ring) n d c. (!i. i < n ==> c i IN ring_carrier r) ==> d <= n ==> poly_evaluate r (elementary_sympoly_range r n d) c = ring_mul r ( ring_pow r (ring_neg r (ring_1 r)) d ) ( coeff (n - d) (monic_vanishing_at r (range n) c) ) `, intro THEN simp[eval_elementary_sympoly_range] THEN have `FINITE(range n)` [finite_range] THEN have `CARD(range n) = n` [card_range] THEN have `!i. i IN range n ==> c i IN ring_carrier(r:R ring)` [range_lt] THEN have `d <= CARD(range n)` [] THEN subgoal `coeff (CARD (range n) - d) (monic_vanishing_at(r:R ring) (range n) c) = ring_mul r (ring_pow r (ring_neg r (ring_1 r)) d) (ring_sum r {U | U SUBSET (range n) /\ CARD U = d} (\U. ring_product r U c))` THENL [ specialize[ `r:R ring`; `c:num->R`; `range n` ]coeff_monic_vanishing_at THEN qed[] ; pass ] THEN have `coeff (n - d) (monic_vanishing_at(r:R ring) (range n) c) = coeff (CARD (range n) - d) (monic_vanishing_at r (range n) c)` [] THEN simp[] THEN specialize[`r:R ring`;`n:num`]ring_pow_neg_1_mul_refl THEN simp[RING_MUL_ASSOC;RING_SUM;RING_POW;RING_NEG;RING_1;RING_MUL_LID] );; let elementary_sympoly_range_o_permutes = prove(` !(r:R ring) n d f m. f permutes range n ==> elementary_sympoly_range r n d (m o f) = elementary_sympoly_range r n d m `, rw[elementary_sympoly_range] THEN intro THEN case `?U. U SUBSET range n /\ CARD U = d /\ (m:num->num) = (\i. if i IN U then 1 else 0)` THENL [ choose `U:num->bool` `U SUBSET range n /\ CARD U = d /\ (m:num->num) = (\i. if i IN U then 1 else 0)` [] THEN def `V:num->bool` `(U:num->bool) o (f:num->num)` THEN subgoal `?V. V SUBSET range n /\ CARD V = d /\ (m:num->num) o f = (\i. if i IN V then 1 else 0)` THENL [ witness `V:num->bool` THEN intro THENL [ qed[o_permutes_subset] ; qed[card_o_permutes;FINITE_SUBSET;finite_range] ; rw[FUN_EQ_THM] THEN simp[IN;o_THM] ] ; pass ] THEN qed[] ; pass ] THEN case `?U. U SUBSET range n /\ CARD U = d /\ (m:num->num) o f = (\i. if i IN U then 1 else 0)` THENL [ choose `U:num->bool` `U SUBSET range n /\ CARD U = d /\ (m:num->num) o f = (\i. if i IN U then 1 else 0)` [] THEN def `V:num->bool` `IMAGE (f:num->num) (U:num->bool)` THEN subgoal `?V. V SUBSET range n /\ CARD V = d /\ (m:num->num) = (\i. if i IN V then 1 else 0)` THENL [ witness `V:num->bool` THEN intro THENL [ qed[image_permutes_subset] ; qed[card_image_permutes;FINITE_SUBSET;finite_range] ; rw[FUN_EQ_THM] THEN intro THEN def `y:num` `inverse f (x:num):num` THEN have `f(y:num) = x:num` [PERMUTES_INVERSES] THEN have `m(x:num):num = (m o f)(y:num)` [o_THM] THEN case `y:num IN U` THENL [ have `(m o (f:num->num))(y:num) = 1` [] THEN have `x:num IN V` [IN_IMAGE] THEN qed[] ; pass ] THEN have `(m o (f:num->num))(y:num) = 0` [] THEN subgoal `~(x:num IN V)` THENL [ intro THEN choose `z:num` `z IN U /\ f(z:num) = x:num` [IN_IMAGE] THEN have `y = z:num` [PERMUTES_INJECTIVE] THEN qed[] ; pass ] THEN qed[] ] ; pass ] THEN qed[] ; pass ] THEN qed[] );; let elementary_sympoly_range_le = prove(` !(r:R ring) n d m. ~(elementary_sympoly_range r n d m = ring_0 r) ==> monomial_le (<=) m (\i. if n-d <= i /\ i < n then 1 else 0) `, rw[elementary_sympoly_range] THEN intro THEN choose `U:num->bool` `U SUBSET range n /\ CARD U = d /\ m = (\i. if i IN U then 1 else 0)` [] THEN have `FINITE(range n)` [finite_range] THEN have `CARD(U:num->bool) <= CARD(range n)` [CARD_SUBSET] THEN have `d <= n:num` [card_range] THEN simp[MONOMIAL_LE_TOSET;TOSET_num;FLD_num] THEN simp[monomial;monomial_vars;SUBSET_UNIV] THEN rw[properly_le] THEN intro THENL [ subgoal `{i:num | ~((if i IN U then 1 else 0) = 0)} = U` THENL [ SET_TAC[ARITH_RULE `~(1 = 0)`] ; pass ] THEN qed[FINITE_SUBSET;finite_range] ; set_fact `{i:num | ~((if n-d <= i /\ i < n then 1 else 0) = 0)} SUBSET {i:num | i < n}` THEN qed[FINITE_SUBSET;FINITE_NUMSEG_LT] ; pass ] THEN specialize[`n-d:num`;`n:num`]finite_le_lt THEN have `FINITE({i:num | n-d <= i /\ i < n} DIFF U)` [FINITE_DIFF] THEN case `({i:num | n-d <= i /\ i < n} DIFF U) = {}` THENL [ set_fact `({i:num | n-d <= i /\ i < n} DIFF U) = {} ==> {i:num | n-d <= i /\ i < n} SUBSET U` THEN specialize[`n-d:num`;`n:num`]card_le_lt THEN num_linear_fact `d <= n:num ==> n-(n-d) = d` THEN have `CARD {i:num | n-d <= i /\ i < n} = d` [] THEN have `FINITE(U:num->bool)` [FINITE_SUBSET] THEN have `{i:num | n-d <= i /\ i < n} = U` [CARD_SUBSET_EQ] THEN DISJ1_TAC THEN rw[FUN_EQ_THM] THEN ASM SET_TAC[] ; pass ] THEN specialize[`{i:num | n-d <= i /\ i < n} DIFF U`]max_finite THEN choose `i:num` `i IN {i:num | n-d <= i /\ i < n} DIFF U /\ (!j. i < j ==> ~(j IN {i:num | n-d <= i /\ i < n} DIFF U))` [] THEN DISJ2_TAC THEN witness `i:num` THEN have `~(i:num IN U)` [IN_DIFF] THEN have `i IN {i:num | n-d <= i /\ i < n}` [IN_DIFF] THEN set_fact `i IN {i:num | n-d <= i /\ i < n} ==> n-d <= i /\ i < n` THEN simp[ARITH_RULE `0 < 1`] THEN intro THEN have `~(j IN {i:num | n-d <= i /\ i < n} DIFF U)` [] THEN have `n-d <= i:num` [] THEN num_linear_fact `n-d <= i ==> i < j ==> n-d <= j:num` THEN case `j < n:num` THENL [ set_fact `n - d <= j ==> j < n ==> j IN {i:num | n-d <= i /\ i < n}` THEN qed[IN_DIFF] ; pass ] THEN qed[SUBSET;range_lt] );; let support_le_elementary_sympoly_range = prove(` !(r:R ring) n d. support_le r (elementary_sympoly_range r n d) (\i. if n-d <= i /\ i < n then 1 else 0) `, rw[support_le] THEN qed[elementary_sympoly_range_le] );; let support_le1_elementary_sympoly_range = prove(` !(r:R ring) n d. d <= n ==> support_le1 r (elementary_sympoly_range r n d) (\i. if n-d <= i /\ i < n then 1 else 0) `, rw[support_le1;support_le_elementary_sympoly_range] THEN rw[elementary_sympoly_range] THEN intro THEN subgoal `?U. U SUBSET range n /\ CARD U = d /\ (\i. if n - d <= i /\ i < n then 1 else 0) = (\i. if i IN U then 1 else 0)` THENL [ witness `{i:num | n-d <= i /\ i < n}` THEN rw[IN_ELIM_THM;card_le_lt;range_set_lt] THEN num_linear_fact `d <= n:num ==> n-(n-d) = d` THEN ASM SET_TAC[] ; pass ] THEN simp[] );; let support_le1_pow_elementary_sympoly_range = prove(` !(r:R ring) n d e. d <= n ==> support_le1 r (poly_pow r (elementary_sympoly_range r n d) e) (\i. if n-d <= i /\ i < n then e else 0) `, intro THEN specialize[`r:R ring`;`n:num`;`d:num`]support_le1_elementary_sympoly_range THEN have `ring_polynomial(r:R ring) (elementary_sympoly_range r n d)` [polynomial_elementary_sympoly_range] THEN specialize[ `r:R ring`; `elementary_sympoly_range(r:R ring) n d`; `\i:num. if n - d <= i /\ i < n then 1 else 0`; `e:num` ]support_le1_pow THEN subgoal `monomial_pow (\i:num. if n - d <= i /\ i < n then 1 else 0) e = (\i. if n - d <= i /\ i < n then e else 0)` THENL [ rw[monomial_pow;FUN_EQ_THM] THEN ARITH_TAC ; pass ] THEN qed[] );; let support_le1_product_pow_elementary_sympoly_range = prove(` !(r:R ring) n e. support_le1 r (ring_product(poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) ) (\i. nsum (1..n) (\d. if n-d <= i /\ i < n then e d else 0)) `, intro THEN have `FINITE (1..n)` [FINITE_NUMSEG] THEN have `!d. d IN 1..n ==> ring_polynomial(r:R ring) (poly_pow r (elementary_sympoly_range r n d) (e d))` [polynomial_elementary_sympoly_range;poly_pow_poly] THEN subgoal `!d. d IN 1..n ==> support_le1(r:R ring) (poly_pow r (elementary_sympoly_range r n d) (e d)) (\i. if n-d <= i /\ i < n then e d else 0)` THENL [ intro THEN have `d <= n:num` [IN_NUMSEG] THEN specialize[ `r:R ring` ]support_le1_pow_elementary_sympoly_range THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `\d. poly_pow r (elementary_sympoly_range(r:R ring) n d) (e d)`; `\d i:num. if n-d <= i /\ i < n then e d else 0`; `1..n` ]support_le1_product THEN subgoal `monomial_product (1..n) (\d i:num. if n - d <= i /\ i < n then e d else 0) = (\i. nsum (1..n) (\d. if n - d <= i /\ i < n then e d else 0))` THENL [ rw[monomial_product] ; pass ] THEN qed[] );; let support_le1_product_pow_elementary_sympoly_range_v2 = prove(` !(r:R ring) n e. support_le1 r (ring_product(poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) ) (\i. if i < n then nsum (n-i..n) e else 0) `, intro THEN subgoal `(\i. if i < n then nsum (n-i..n) e else 0) = (\i. nsum (1..n) (\d. if n-d <= i /\ i < n then e d else 0))` THENL [ rw[FUN_EQ_THM] THEN intro THEN case `x < n:num` THENL [ have `FINITE(1..n)` [FINITE_NUMSEG] THEN specialize[`\d. n-d <= x:num`;`1..n`;`e:num->num`](GSYM NSUM_RESTRICT_SET) THEN subgoal `(n-x..n) = {d | d IN 1..n /\ n-d <= x}` THENL [ rw[EXTENSION;IN_NUMSEG;IN_ELIM_THM] THEN ASM_ARITH_TAC ; pass ] THEN simp[] ; pass ] THEN simp[NSUM_0] ; pass ] THEN qed[support_le1_product_pow_elementary_sympoly_range] );; let symmetric_subring_if_poly_subring_lemma = prove(` !(r:R ring) G n c:num->R M p. (!i. i < n ==> c i IN ring_carrier r) ==> (!d. coeff d (monic_vanishing_at r (range n) c) IN ring_carrier(subring_generated r G)) ==> ring_polynomial(subring_generated r G) p ==> poly_vars r p SUBSET range n ==> (!f m. f permutes range n ==> p(m o f) = p(m)) ==> support_le r p M ==> poly_evaluate r p c IN ring_carrier(subring_generated r G) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby monomial_induction THEN intro THEN have `ring_polynomial r (p:(num->num)->R)` [ring_polynomial_if_subring] THEN case `p = poly_0(r:R ring):(num->num)->R` THENL [ simp[POLY_EVALUATE_0] THEN qed[SUBRING_GENERATED;RING_0] ; pass ] THEN case `n = 0` THENL [ have `range n = {}` [range_0] THEN have `poly_vars(r:R ring) (p:(num->num)->R) = {}` [SUBSET_EMPTY] THEN have `p:(num->num)->R = poly_const r (p monomial_1)` [poly_vars_empty] THEN have `p(monomial_1:num->num):R IN ring_carrier r` [ring_polynomial;ring_powerseries] THEN have `poly_evaluate r (p:(num->num)->R) c = p monomial_1` [POLY_EVALUATE_CONST] THEN qed[ring_polynomial;ring_powerseries] ; pass ] THEN choose `t:num->num` `~(p t = ring_0(r:R ring)) /\ support_le r p t` [poly_first_monomial] THEN case `monomial_lt ((<=):num->num->bool) t M` THENL [ have `support_le(r:R ring) p t` [] THEN specialize[`t:num->num`](know(`!m:num->num. monomial_lt (<=) m M ==> (!p. (!i. i < n ==> c i IN ring_carrier(r:R ring)) ==> (!d. coeff d (monic_vanishing_at r (range n) c) IN ring_carrier (subring_generated r G)) ==> ring_polynomial(subring_generated r G) p ==> poly_vars r p SUBSET range n ==> (!f m. f permutes range n ==> p (m o f) = p m) ==> support_le r p m ==> poly_evaluate r p c IN ring_carrier (subring_generated r G))`)) THEN qed[] ; pass ] THEN subgoal `!i:num. i+1 < n ==> t i <= t (i+1):num` THENL [ intro THEN subgoal `swap(i,i+1) permutes range n` THENL [ num_linear_fact `i+1 < n ==> i < n` THEN qed[swap_permutes_range] ; pass ] THEN have `p(t o (swap(i,i+1))) = p(t:num->num):R` [] THEN have `~(p((t:num->num) o (swap(i,i+1))) = ring_0(r:R ring))` [] THEN have `monomial_le (<=) (t o (swap(i,i+1))) (t:num->num)` [support_le] THEN have `FINITE(monomial_vars(t:num->num))` [ring_polynomial;ring_powerseries;INFINITE] THEN qed[monomial_le_swap] ; pass ] THEN def `e:num->num` `\i. if i = n then t 0 else t(n-i) - t(n-i-1):num` THEN subgoal `t = (\i:num. if i < n then nsum(n-i..n) e else 0)` THENL [ rw[FUN_EQ_THM] THEN INDUCT_TAC THENL [ num_linear_fact `~(n = 0) ==> 0 < n` THEN num_linear_fact `n-0 = n` THEN simp[NSUM_SING_NUMSEG] ; rw[ARITH_RULE `SUC x = x+1`] THEN case `x+1 < n` THENL [ num_linear_fact `x+1 < n ==> n-(x+1)+1 = n-x` THEN num_linear_fact `n-(x+1) <= n` THEN specialize[ `e:num->num`; `n-(x+1)`; `n:num` ]NSUM_CLAUSES_LEFT THEN havetac `nsum (n - (x + 1) + 1..n) e = nsum(n-x..n) e` (simp[]) THEN num_linear_fact `x + 1 < n ==> x < n` THEN havetac `nsum (n - (x + 1) + 1..n) e = t x` (simp[]) THEN subgoal `t(x+1) = e(n-(x+1)) + t x:num` THENL [ num_linear_fact `x+1 < n ==> ~(n-(x+1) = n)` THEN have `e(n-(x+1)) = t(n-(n-(x+1))) - t(n-(n-(x+1))-1):num` [] THEN num_linear_fact `x+1 < n ==> n-(n-(x+1)) = x+1` THEN num_linear_fact `x+1 < n ==> n-(n-(x+1))-1 = x` THEN havetac `e(n-(x+1)) = t(x+1) - t(x):num` (simp[]) THEN have `t x <= t(x+1):num` [] THEN num_linear_fact `t x <= t(x+1):num ==> e (n - (x + 1)) = t (x + 1) - t x ==> t (x + 1) = e (n - (x + 1)) + t x` THEN qed[] ; pass ] THEN simp[] ; pass ] THEN have `~(x+1 IN range n)` [range_lt] THEN have `~(x+1 IN poly_vars(r:R ring) (p:(num->num)->R))` [SUBSET] THEN have `~(x+1 IN UNIONS {monomial_vars m | ~((p:(num->num)->R) m = ring_0 r)})` [poly_vars] THEN have `!v. v IN {monomial_vars m | ~((p:(num->num)->R) m = ring_0 r)} ==> ~(x+1 IN v)` [IN_UNIONS] THEN proven_if `t(x+1) = 0` [] THEN subgoal `x+1 IN monomial_vars t` THENL [ rw[monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `monomial_vars t IN {monomial_vars m | ~((p:(num->num)->R) m = ring_0 r)}` THENL [ rw[IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[] ] ; pass ] THEN specialize[ `r:R ring`; `n:num`; `e:num->num` ]support_le1_product_pow_elementary_sympoly_range_v2 THEN have `support_le1(r:R ring) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) t` [] THEN subgoal `ring_polynomial(r:R ring) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))` THENL [ have `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN ring_carrier(poly_ring r (:num))` [RING_PRODUCT] THEN specialize[`r:R ring`;`(:num)`](CONJUNCT1 POLY_RING) THEN have `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:num)}` [] THEN set_fact `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET (:num)} ==> ring_polynomial(r:R ring) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))` THEN qed[] ; pass ] THEN have `support_le(r:R ring) p t` [] THEN specialize[ `r:R ring`; `p:(num->num)->R`; `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))`; `t:num->num` ]support_le_cancel THEN def `s:(num->num)->R` `poly_sub(r:R ring) p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))` THEN have `ring_polynomial r (s:(num->num)->R)` [RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_CONST;ring_polynomial;ring_powerseries] THEN subgoal `poly_evaluate(r:R ring) (s:(num->num)->R) c IN ring_carrier(subring_generated r G)` THENL [ case `s = poly_0 r:(num->num)->R` THENL [ have `poly_evaluate(r:R ring) (s:(num->num)->R) c = ring_0 r` [POLY_EVALUATE_0] THEN have `ring_0(r:R ring) = ring_0(subring_generated r G)` [SUBRING_GENERATED] THEN have `poly_evaluate(r:R ring) (s:(num->num)->R) c = ring_0(subring_generated r G)` [] THEN qed[RING_0] ; pass ] THEN choose `S:num->num` `~(s S = ring_0(r:R ring)) /\ support_le r s S` [poly_first_monomial] THEN have `monomial_lt (<=) (S:num->num) t` [support_lt] THEN have `monomial_le (<=) (t:num->num) M` [support_le] THEN have `monomial_lt (<=) (S:num->num) M` [MONOMIAL_LTE_TRANS;POSET_num] THEN subgoal `ring_polynomial(subring_generated r G) (s:(num->num)->R)` THENL [ intro THEN subgoal `s(m:num->num) = ring_sub(r:R ring) (p m) (ring_mul r (p t) ((ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) m))` THENL [ once_rw[know(`s = poly_sub(r:R ring) p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))`)] THEN rw[poly_sub] THEN specialize_assuming[ `r:R ring`; `p(t:num->num):R`; `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))` ]poly_const_times THEN qed[ring_polynomial;ring_powerseries] ; pass ] THEN have `!d. elementary_sympoly_range r n d = elementary_sympoly_range(subring_generated(r:R ring) G) n d` [elementary_sympoly_range_subring] THEN have `!d. ring_polynomial(subring_generated r G) (elementary_sympoly_range(subring_generated(r:R ring) G) n d)` [polynomial_elementary_sympoly_range] THEN have `!d. poly_pow(r:R ring) (elementary_sympoly_range(subring_generated r G) n d) (e d) = poly_pow (subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d)` [poly_pow_subring;ring_polynomial] THEN have `!d. ring_polynomial(subring_generated(r:R ring) G) (poly_pow (subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))` [poly_pow_poly] THEN have `FINITE(1..n)` [FINITE_NUMSEG] THEN specialize_assuming[ `r:R ring`; `G:R->bool`; `\d. poly_pow (subring_generated(r:R ring) G) (elementary_sympoly_range(subring_generated r G) n d) (e d)`; `1..n` ]poly_product_subring_multi THEN subgoal `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) = ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow (subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))` THENL [ sufficesby RING_PRODUCT_EQ THEN qed[] ; pass ] THEN have `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow (subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d)) = ring_product (poly_ring(subring_generated r G) (:num)) (1..n) (\d. poly_pow (subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))` [ring_polynomial] THEN have `s = poly_sub(r:R ring) p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring(subring_generated r G) (:num)) (1..n) (\d. poly_pow(subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))))` [] THEN have `poly_const r (p t) = poly_const(subring_generated r G) (p(t:num->num)):(num->num)->R` [poly_const_subring] THEN have `s = poly_sub(r:R ring) p (poly_mul r (poly_const(subring_generated r G) (p t)) (ring_product (poly_ring(subring_generated r G) (:num)) (1..n) (\d. poly_pow(subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))))` [] THEN have `ring_polynomial (subring_generated(r:R ring) G) (poly_const (subring_generated r G) (p(t:num->num)):(num->num)->R)` [RING_POLYNOMIAL_CONST;ring_polynomial;ring_powerseries] THEN have `ring_powerseries (subring_generated(r:R ring) G) (poly_const (subring_generated r G) (p(t:num->num)):(num->num)->R)` [ring_polynomial] THEN have `ring_powerseries (subring_generated(r:R ring) G) (ring_product (poly_ring (subring_generated r G) (:num)) (1..n) (\d. poly_pow (subring_generated r G) (elementary_sympoly_range (subring_generated r G) n d) (e d)))` [RING_PRODUCT;GSYM poly_use;ring_polynomial] THEN specialize[ `r:R ring`; `G:R->bool`; `poly_const(subring_generated(r:R ring) G) (p(t:num->num)):(num->num)->R`; `ring_product (poly_ring(subring_generated(r:R ring) G) (:num)) (1..n) (\d. poly_pow(subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))` ]poly_mul_subring THEN have `s = poly_sub(r:R ring) p (poly_mul(subring_generated r G) (poly_const(subring_generated r G) (p t)) (ring_product (poly_ring(subring_generated r G) (:num)) (1..n) (\d. poly_pow(subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))))` [] THEN have `s = poly_sub(subring_generated(r:R ring) G) p (poly_mul(subring_generated r G) (poly_const(subring_generated r G) (p t)) (ring_product (poly_ring(subring_generated r G) (:num)) (1..n) (\d. poly_pow(subring_generated r G) (elementary_sympoly_range(subring_generated r G) n d) (e d))))` [poly_sub_subring] THEN qed[RING_POLYNOMIAL_SUB;RING_POLYNOMIAL_MUL;RING_POLYNOMIAL_CONST;poly_use;RING_PRODUCT] ; pass ] THEN subgoal `!m:num->num. s m IN ring_carrier(subring_generated(r:R ring) G)` THENL [ qed[ring_polynomial;ring_powerseries] ; pass ] THEN subgoal `poly_vars r (s:(num->num)->R) SUBSET range n` THENL [ have `!d. elementary_sympoly_range(r:R ring) n d IN ring_carrier(poly_ring r (range n))` [elementary_sympoly_range_in_poly_ring] THEN have `!d. d IN 1..n ==> poly_pow r (elementary_sympoly_range(r:R ring) n d) (e d) IN ring_carrier(poly_ring r (range n))` [poly_pow_in_poly_ring] THEN specialize_assuming[ `r:R ring`; `\d. poly_pow(r:R ring) (elementary_sympoly_range r n d) (e d)`; `range n`; `1..n` ]poly_product_in_poly_ring THEN have `ring_product(poly_ring (r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN ring_carrier(poly_ring r (range n))` [FINITE_NUMSEG] THEN have `poly_const r (p(t:num->num)):(num->num)->R IN ring_carrier(poly_ring r (range n))` [POLY_CONST;ring_polynomial;ring_powerseries] THEN have `poly_mul r (poly_const r (p(t:num->num))) (ring_product(poly_ring (r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) IN ring_carrier(poly_ring r (range n))` [poly_mul_in_poly_ring] THEN subgoal `p:(num->num)->R IN ring_carrier(poly_ring r (range n))` THENL [ rw[POLY_RING;IN_ELIM_THM] THEN qed[] ; pass ] THEN have `s:(num->num)->R IN ring_carrier(poly_ring r (range n))` [poly_sub_in_poly_ring] THEN specialize[`r:R ring`;`range n`](CONJUNCT1 POLY_RING) THEN have `s:(num->num)->R IN {q | ring_polynomial r q /\ poly_vars r q SUBSET range n}` [] THEN ASM SET_TAC[] ; pass ] THEN subgoal `!f m:num->num. f permutes range n ==> s(m o f) = s m:R` THENL [ intro THEN rw[know `s = poly_sub(r:R ring) p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))`] THEN rw[poly_sub_o_permutes] THEN have_rw `!m. p(m o f) = p(m:num->num):R` [] THEN subgoal `(\m:num->num. poly_mul(r:R ring) (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) (m o f)) = poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))` THENL [ rw[FUN_EQ_THM] THEN intro THEN specialize[ `r:R ring`; `poly_const(r:R ring) (p(t:num->num)):(num->num)->R` ]poly_mul_o_permutes THEN have_rw `poly_mul(r:R ring) (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) (x o f) = poly_mul r (\m:num->num. (poly_const r (p t)) (m o f)) (\m. (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) (m o f)) x` [] THEN have_rw `!m:num->num. poly_const r (p(t:num->num):R) (m o (f:num->num)) = poly_const r (p t) m` [poly_const_o_permutes] THEN subgoal `(\m:num->num. ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) (m o f)) = ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))` THENL [ rw[FUN_EQ_THM] THEN intro THEN have `FINITE(1..n)` [FINITE_NUMSEG] THEN specialize[ `r:R ring`; `\d. poly_pow(r:R ring) (elementary_sympoly_range r n d) (e d)`; `range n`; `f:num->num`; `1..n` ]poly_product_o_permutes THEN have `!s. s IN 1..n ==> ring_polynomial(r:R ring) (poly_pow r (elementary_sympoly_range r n s) (e s))` [poly_pow_poly;polynomial_elementary_sympoly_range] THEN specialize[ `x':num->num` ](know `!m. f permutes range n ==> (!s. s IN 1..n ==> ring_polynomial(r:R ring) (poly_pow r (elementary_sympoly_range r n s) (e s))) ==> ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) (m o f) = ring_product (poly_ring r (:num)) (1..n) (\s m. poly_pow r (elementary_sympoly_range r n s) (e s) (m o f)) m`) THEN rw[know `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) (x' o f) = ring_product (poly_ring r (:num)) (1..n) (\s m. poly_pow r (elementary_sympoly_range r n s) (e s) (m o f)) x'`] THEN subgoal `(\s m:num->num. poly_pow(r:R ring) (elementary_sympoly_range r n s) (e s) (m o f)) = (\d. poly_pow r (elementary_sympoly_range r n d) (e d))` THENL [ once_rw[FUN_EQ_THM] THEN rw[fun_eq_thm_v] THEN intro THEN have `ring_powerseries(r:R ring) (elementary_sympoly_range r n x'')` [powerseries_elementary_sympoly_range] THEN specialize[ `r:R ring`; `elementary_sympoly_range(r:R ring) n x''`; `range n`; `f:num->num`; `e(x'':num):num`; `v:num->num` ]poly_pow_o_permutes THEN rw[know `poly_pow(r:R ring) (elementary_sympoly_range r n x'') (e x'') (v o (f:num->num)) = poly_pow r (\m. elementary_sympoly_range r n x'' (m o f)) (e x'') v`] THEN subgoal `(\m. elementary_sympoly_range(r:R ring) n x'' (m o f)) = elementary_sympoly_range r n x''` THENL [ rw[FUN_EQ_THM] THEN qed[elementary_sympoly_range_o_permutes] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN rw[know `(\m:num->num. ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) (m o f)) = ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))`] THEN subgoal `(\m. poly_const r (p t) m) = poly_const(r:R ring) (p(t:num->num)):(num->num)->R` THENL [ rw[FUN_EQ_THM] ; pass ] THEN qed[] ; pass ] THEN rw[know `(\m:num->num. poly_mul(r:R ring) (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) (m o f)) = poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))`] THEN havetac `(p:(num->num)->R) = (\m. p m)` (rw[FUN_EQ_THM]) THEN qed[] ; pass ] THEN specialize[`S:num->num`](know(`!m:num->num. monomial_lt (<=) m M ==> (!p. (!i. i < n ==> c i IN ring_carrier(r:R ring)) ==> (!d. coeff d (monic_vanishing_at r (range n) c) IN ring_carrier (subring_generated r G)) ==> ring_polynomial(subring_generated r G) p ==> poly_vars r p SUBSET range n ==> (!f m. f permutes range n ==> p (m o f) = p m) ==> support_le r p m ==> poly_evaluate r p c IN ring_carrier (subring_generated r G))`)) THEN have `support_le(r:R ring) (s:(num->num)->R) S` [] THEN specialize[`s:(num->num)->R`](know(`!p. (!i. i < n ==> c i IN ring_carrier(r:R ring)) ==> (!d. coeff d (monic_vanishing_at r (range n) c) IN ring_carrier (subring_generated r G)) ==> ring_polynomial(subring_generated r G) p ==> poly_vars r p SUBSET range n ==> (!f m. f permutes range n ==> p (m o f) = p m) ==> support_le r p S ==> poly_evaluate r p c IN ring_carrier (subring_generated r G)`)) THEN qed[] ; pass ] THEN subgoal `poly_evaluate(r:R ring) s c = ring_sub r (poly_evaluate r p c) (ring_mul r (p t) (ring_product r (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d))))` THENL [ rw[know `s = poly_sub(r:R ring) p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))`] THEN have `ring_polynomial(r:R ring) (poly_const r (p(t:num->num)):(num->num)->R)` [RING_POLYNOMIAL_CONST;ring_polynomial;ring_powerseries] THEN specialize_assuming[ `r:R ring`; `poly_const r (p(t:num->num)):(num->num)->R`; `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))` ]RING_POLYNOMIAL_MUL THEN have `ring_polynomial(r:R ring) (poly_mul r (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))` [] THEN specialize[ `r:R ring`; `p:(num->num)->R`; `poly_mul(r:R ring) (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))`; `c:num->R` ]POLY_EVALUATE_SUB THEN rw[know `poly_evaluate(r:R ring) (poly_sub r p (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))))) c = ring_sub r (poly_evaluate r p c) (poly_evaluate r (poly_mul r (poly_const r (p t)) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))) c)`] THEN subgoal `poly_evaluate(r:R ring) (poly_mul r (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))) c = ring_mul r (p t) (ring_product r (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d)))` THENL [ subgoal `!i. i IN poly_vars r (poly_const r (p(t:num->num))) UNION poly_vars r (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) ==> c i IN ring_carrier (r:R ring)` THENL [ rw[IN_UNION;POLY_VARS_CONST;EMPTY;IN_ELIM_THM] THEN have `!d. elementary_sympoly_range(r:R ring) n d IN ring_carrier(poly_ring r (range n))` [elementary_sympoly_range_in_poly_ring] THEN have `!d. d IN 1..n ==> poly_pow r (elementary_sympoly_range(r:R ring) n d) (e d) IN ring_carrier(poly_ring r (range n))` [poly_pow_in_poly_ring] THEN specialize_assuming[ `r:R ring`; `\d. poly_pow(r:R ring) (elementary_sympoly_range r n d) (e d)`; `range n`; `1..n` ]poly_product_in_poly_ring THEN have `ring_product(poly_ring (r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN ring_carrier(poly_ring r (range n))` [FINITE_NUMSEG] THEN specialize[`r:R ring`;`range n`](CONJUNCT1 POLY_RING) THEN have `ring_product(poly_ring (r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET range n}` [] THEN intro THEN set_fact `ring_product(poly_ring (r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)) IN {q | ring_polynomial r q /\ poly_vars r q SUBSET range n} ==> i IN poly_vars r (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) ==> i IN range n` THEN qed[range_lt] ; pass ] THEN specialize[ `r:R ring`; `poly_const r (p(t:num->num)):(num->num)->R`; `ring_product (poly_ring(r:R ring) (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))`; `c:num->R` ]POLY_EVALUATE_MUL THEN rw[know `poly_evaluate(r:R ring) (poly_mul r (poly_const r (p(t:num->num))) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d)))) c = ring_mul r (poly_evaluate r (poly_const r (p t)) c) (poly_evaluate r (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) c)`] THEN have_rw `poly_evaluate r (poly_const r (p(t:num->num)):(num->num)->R) c = p t` [POLY_EVALUATE_CONST;ring_polynomial;ring_powerseries] THEN subgoal `poly_evaluate(r:R ring) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) c = ring_product r (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d))` THENL [ have `FINITE(1..n)` [FINITE_NUMSEG] THEN have `!v. v IN range n ==> c v IN ring_carrier(r:R ring)` [range_lt] THEN have `!s. s IN 1..n ==> poly_pow(r:R ring) (elementary_sympoly_range r n s) (e s) IN ring_carrier (poly_ring r (range n))` [poly_pow_in_poly_ring;elementary_sympoly_range_in_poly_ring] THEN specialize[ `r:R ring`; `\d. poly_pow(r:R ring) (elementary_sympoly_range r n d) (e d)`; `c:num->R`; `range n`; `1..n` ]eval_poly_product_multi THEN rw[know `poly_evaluate(r:R ring) (ring_product (poly_ring r (:num)) (1..n) (\d. poly_pow r (elementary_sympoly_range r n d) (e d))) c = ring_product r (1..n) (\s. poly_evaluate r (poly_pow r (elementary_sympoly_range r n s) (e s)) c)`] THEN sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN have `elementary_sympoly_range(r:R ring) n a IN ring_carrier (poly_ring r (range n))` [elementary_sympoly_range_in_poly_ring] THEN have `!v. v IN range n ==> c v IN ring_carrier(r:R ring)` [range_lt] THEN specialize[ `r:R ring`; `elementary_sympoly_range (r:R ring) n a`; `c:num->R`; `range n`; `e(a:num):num` ]eval_poly_pow_multi THEN qed[] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_mul(r:R ring) (p(t:num->num)) (ring_product r (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d))) IN ring_carrier (subring_generated r G)` THENL [ have `(p(t:num->num)) IN ring_carrier (subring_generated(r:R ring) G)` [ring_polynomial;ring_powerseries] THEN subgoal `ring_product(r:R ring) (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d)) IN ring_carrier (subring_generated r G)` THENL [ sufficesby ring_product_in_subring THEN rw[BETA_THM] THEN intro THEN subgoal `poly_evaluate r (elementary_sympoly_range(r:R ring) n s') c IN ring_carrier(subring_generated r G)` THENL [ have `s' <= n:num` [IN_NUMSEG] THEN simp[eval_elementary_sympoly_range_coeff] THEN have `ring_1(r:R ring) IN ring_carrier(subring_generated r G)` [SUBRING_GENERATED;RING_1] THEN have `ring_neg(subring_generated(r:R ring) G) (ring_1 r) IN ring_carrier(subring_generated r G)` [RING_NEG] THEN have `ring_neg(r:R ring) (ring_1 r) IN ring_carrier(subring_generated r G)` [SUBRING_GENERATED] THEN have `ring_pow(subring_generated(r:R ring) G) (ring_neg r (ring_1 r)) s':R IN ring_carrier(subring_generated r G)` [RING_POW] THEN have `ring_pow r (ring_neg r (ring_1 r)) s':R IN ring_carrier(subring_generated r G)` [RING_POW_SUBRING_GENERATED] THEN have `(coeff (n - s') (monic_vanishing_at r (range n) c)) IN ring_carrier (subring_generated(r:R ring) G)` [] THEN have `ring_mul(subring_generated r G) (ring_pow r (ring_neg r (ring_1 r)) s':R) (coeff (n - s') (monic_vanishing_at r (range n) c)) IN ring_carrier (subring_generated(r:R ring) G)` [RING_MUL] THEN qed[SUBRING_GENERATED] ; pass ] THEN have `ring_pow(subring_generated(r:R ring) G) (poly_evaluate r (elementary_sympoly_range r n s') c) (e s') IN ring_carrier (subring_generated r G)` [RING_POW] THEN qed[RING_POW_SUBRING_GENERATED] ; pass ] THEN have `ring_mul(subring_generated(r:R ring) G) (p(t:num->num)) (ring_product r (1..n) (\d. ring_pow r (poly_evaluate r (elementary_sympoly_range r n d) c) (e d))) IN ring_carrier (subring_generated r G)` [RING_MUL] THEN qed[SUBRING_GENERATED] ; pass ] THEN qed[sub_in_subring;POLY_EVALUATE] );; let symmetric_subring_if_poly_subring_range = prove(` !(r:R ring) G n c:num->R p. (!i. i < n ==> c i IN ring_carrier r) ==> (!d. coeff d (monic_vanishing_at r (range n) c) IN ring_carrier(subring_generated r G)) ==> ring_polynomial(subring_generated r G) p ==> poly_vars r p SUBSET range n ==> (!f m. f permutes range n ==> p(m o f) = p(m)) ==> poly_evaluate r p c IN ring_carrier(subring_generated r G) `, intro THEN have `ring_polynomial r (p:(num->num)->R)` [ring_polynomial_if_subring] THEN choose `M:num->num` `support_le r (p:(num->num)->R) M` [support_le_exists] THEN specialize[ `r:R ring`; `G:R->bool`; `n:num`; `c:num->R`; `M:num->num`; `p:(num->num)->R` ]symmetric_subring_if_poly_subring_lemma THEN qed[] );; (* ===== poly_reindex: reindexing polynomial variables *) let poly_reindex = new_definition ` poly_reindex (r:R ring) (q:(Y->num)->R) (f:X->Y) (A:X->bool) (B:Y->bool) = (\m. if monomial_vars m SUBSET A then q(\b. if b IN B then m(@a:X. a IN A /\ f a = b) else 0) else ring_0 r) `;; let poly_vars_poly_reindex = prove(` !(r:R ring) A B q f:X->Y. BIJ f A B ==> poly_vars r q SUBSET B ==> poly_vars r (poly_reindex r q f A B) SUBSET A `, rw[poly_reindex] THEN rw[poly_vars;UNIONS_SUBSET;IN_ELIM_THM] THEN rw[monomial_vars;SUBSET;IN_ELIM_THM;EXTENSION] THEN qed[] );; let powerseries_poly_reindex = prove(` !(r:R ring) A B q f:X->Y. BIJ f A B ==> ring_powerseries r q ==> ring_powerseries r (poly_reindex r q f A B) `, rw[poly_reindex] THEN rw[ring_powerseries;INFINITE;monomial_vars;SUBSET;IN_ELIM_THM] THEN intro THENL [ qed[RING_0] ; case `!x:X. ~(m x = 0) ==> x IN A` THENL [ subgoal `~FINITE {b:Y | ~((\b. if b IN B then m (@a:X. a IN A /\ f a = b) else 0) b = 0)}` THENL [ rw[BETA_THM] THEN intro THEN subgoal `{b:Y | ~((if b IN B then m (@a. a IN A /\ f a = b) else 0) = 0)} = IMAGE f {a:X | ~(m a = 0)}` THENL [ rw[EXTENSION;IN_IMAGE;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN have `x:Y IN B` [] THEN have `SURJ (f:X->Y) A B` [BIJ] THEN choose `a:X` `a:X IN A /\ f a = x:Y` [SURJ] THEN witness `a:X` THEN subgoal `(@a:X. a IN A /\ f a = x:Y) = a` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN qed[] ; intro THEN have `x':X IN A` [] THEN have `x:Y IN B` [BIJ;INJ] THEN subgoal `(@a:X. a IN A /\ f a = x:Y) = x'` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN qed[] ] ; pass ] THEN specialize_assuming[ `f:X->Y`; `{a:X | ~(m a = 0)}` ]FINITE_IMAGE_INJ_EQ THEN subgoal `!i j:X. i IN {a | ~(m a = 0)} ==> j IN {a | ~(m a = 0)} ==> f i = f j:Y ==> i = j` THENL [ rw[IN_ELIM_THM] THEN qed[BIJ;INJ] ; pass ] THEN qed[] ; pass ] THEN qed[] ; qed[] ] ] );; let monomials_poly_reindex = prove(` !(r:R ring) A B q f:X->Y. BIJ f A B ==> ring_powerseries r q ==> poly_vars r q SUBSET B ==> {m | ~((poly_reindex r q f A B) m = ring_0 r)} = IMAGE (\m:Y->num. \a:X. if a IN A then m (f a) else 0) {m | ~(q m = ring_0 r)} `, rw[poly_reindex] THEN rw[EXTENSION;in_image_vw;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN witness `(\b:Y. if b IN B then x (@a:X. a IN A /\ f a = b) else 0)` THEN intro THENL [ rw[fun_eq_thm_e] THEN intro THEN case `e:X IN A` THENL [ subgoal `(@a:X. a IN A /\ f a = f e:Y) = e` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN qed[BIJ;INJ] ; pass ] THEN have `monomial_vars(x:X->num) SUBSET A` [] THEN have `~(e:X IN monomial_vars x)` [SUBSET] THEN have `~(e:X IN {i | ~(x i = 0)})` [monomial_vars] THEN set_fact `~(e:X IN {i | ~(x i = 0)}) ==> x e = 0` THEN qed[] ; qed[] ] ; intro THEN subgoal `monomial_vars(x:X->num) SUBSET A` THENL [ rw[monomial_vars;SUBSET;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `(\b:Y. if b IN B then x (@a:X. a IN A /\ f a = b) else 0) = v` THENL [ rw[fun_eq_thm_e] THEN intro THEN case `e:Y IN B` THENL [ have `SURJ (f:X->Y) A B` [BIJ] THEN choose `a:X` `a:X IN A /\ f a = e:Y` [SURJ] THEN subgoal `(@a:X. a IN A /\ f a = e:Y) = a` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN qed[] ; pass ] THEN subgoal `v(e:Y) = 0` THENL [ proven_if `v(e:Y) = 0` [] THEN subgoal `e IN monomial_vars(v:Y->num)` THENL [ rw[monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `e IN poly_vars r (q:(Y->num)->R)` THENL [ rw[poly_vars;IN_UNIONS;IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[SUBSET] ; pass ] THEN qed[] ; pass ] THEN qed[] ] );; let polynomial_poly_reindex = prove(` !(r:R ring) A B q f:X->Y. BIJ f A B ==> ring_polynomial r q ==> poly_vars r q SUBSET B ==> ring_polynomial r (poly_reindex r q f A B) `, rw[ring_polynomial] THEN intro THENL [ qed[powerseries_poly_reindex] ; pass ] THEN specialize[ `r:R ring`; `A:X->bool`; `B:Y->bool`; ]monomials_poly_reindex THEN qed[FINITE_IMAGE] );; let poly_reindex_permutes = prove(` !(r:R ring) A B q f:X->Y. BIJ f A B ==> ring_polynomial r q ==> (!h m. h permutes B ==> q(m o h) = q(m)) ==> (!g m. g permutes A ==> (poly_reindex r q f A B)(m o g) = (poly_reindex r q f A B)(m)) `, intro THEN rw[poly_reindex] THEN case `monomial_vars(m:X->num) SUBSET A` THENL [ subgoal `monomial_vars(m:X->num o g:X->X) SUBSET A` THENL [ rw[SUBSET;monomial_vars;IN_ELIM_THM;o_THM] THEN intro THEN set_fact `~(m(g(x:X):X) = 0) ==> g(x:X):X IN {i | ~(m i = 0)}` THEN have `g(x:X):X IN monomial_vars m` [monomial_vars] THEN have `g(x:X):X IN A` [SUBSET] THEN qed[PERMUTES_IN_IMAGE] ; pass ] THEN simp[] THEN def `h:Y->Y` `\y:Y. if y IN B then f(g(@x:X. x IN A /\ f x = y)) else y` THEN subgoal `h:Y->Y permutes B` THENL [ have `!x y:X. x IN A ==> y IN A ==> f x = f y:Y ==> x = y` [BIJ;INJ] THEN subgoal `!a:X. a IN A ==> h(f a:Y) = f(g a)` THENL [ intro THEN simp[o_THM] THEN subgoal `(@x:X. x IN A /\ f x = f a:Y) = a` THENL [ sufficesby SELECT_UNIQUE THEN qed[] ; pass ] THEN qed[BIJ;INJ] ; pass ] THEN subgoal `!b:Y. ~(b IN IMAGE (f:X->Y) A) ==> h b = b` THENL [ have `SURJ (f:X->Y) A B` [BIJ] THEN simp[image_surj] ; pass ] THEN specialize[ `f:X->Y`; `g:X->X`; `h:Y->Y`; `A:X->bool` ]PERMUTES_TRANSFER THEN qed[BIJ;image_surj] ; pass ] THEN subgoal `(\b:Y. if b IN B then (m o g) (@a:X. a IN A /\ f a = b) else 0) = (\b. if b IN B then m (@a. a IN A /\ f a = b) else 0) o h` THENL [ rw[fun_eq_thm_e;o_THM] THEN intro THEN case `e:Y IN B` THENL [ have `h(e:Y):Y IN B` [PERMUTES_IN_IMAGE] THEN simp[] THEN have `SURJ (f:X->Y) A B` [BIJ] THEN choose `a:X` `a:X IN A /\ f a = e:Y` [SURJ] THEN subgoal `(@x:X. x IN A /\ f x = e:Y) = a` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN simp[] THEN subgoal `(@a':X. a' IN A /\ f a' = f (g a):Y) = g(a:X):X` THENL [ sufficesby SELECT_UNIQUE THEN have `g(a:X):X IN A` [PERMUTES_IN_IMAGE] THEN qed[BIJ;INJ] ; pass ] THEN simp[] ; qed[] ] ; pass ] THEN qed[] ; subgoal `~(monomial_vars(m:X->num o g:X->X) SUBSET A)` THENL [ rw[SUBSET;monomial_vars;IN_ELIM_THM;o_THM] THEN rw[NOT_FORALL_THM;NOT_IMP] THEN choose `a:X` `a IN monomial_vars m /\ ~(a:X IN A)` [SUBSET] THEN witness `a:X` THEN have `g(a:X) = a` [permutes] THEN have `a:X IN {i | ~(m i = 0)}` [monomial_vars] THEN set_fact `a:X IN {i | ~(m i = 0)} ==> ~(m a = 0)` THEN qed[] ; pass ] THEN qed[] ] );; let poly_evaluate_poly_reindex = prove(` !(r:R ring) A B q f:X->Y c:Y->R. BIJ f A B ==> ring_polynomial r q ==> poly_vars r q SUBSET B ==> poly_evaluate r (poly_reindex r q f A B) (c o f) = poly_evaluate r q c `, rw[poly_evaluate;poly_extend] THEN intro THEN have `ring_powerseries r (q:(Y->num)->R)` [ring_polynomial] THEN specialize[ `r:R ring`; `A:X->bool`; `B:Y->bool`; `q:(Y->num)->R`; `f:X->Y` ]monomials_poly_reindex THEN simp[] THEN subgoal `!x y:Y->num. x IN {m | ~(q m:R = ring_0 r)} ==> y IN {m | ~(q m = ring_0 r)} ==> (\a:X. if a IN A then x (f a) else 0) = (\a. if a IN A then y (f a) else 0) ==> x = y` THENL [ rw[IN_ELIM_THM;fun_eq_thm_e] THEN intro THEN case `e:Y IN B` THENL [ choose `a:X` `a:X IN A /\ f a = e:Y` [BIJ;SURJ] THEN qed[] ; pass ] THEN case `~(x(e:Y) = 0)` THENL [ subgoal `e:Y IN monomial_vars x` THENL [ rw[monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `e IN poly_vars r (q:(Y->num)->R)` THENL [ rw[poly_vars;IN_UNIONS;IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[SUBSET] ; pass ] THEN case `~(y(e:Y) = 0)` THENL [ subgoal `e:Y IN monomial_vars y` THENL [ rw[monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `e IN poly_vars r (q:(Y->num)->R)` THENL [ rw[poly_vars;IN_UNIONS;IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[SUBSET] ; pass ] THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `\m a:X. if a IN A then m (f a:Y) else 0`; `\m:X->num. ring_mul(r:R ring) (I (poly_reindex r q (f:X->Y) A B m)) (ring_product r (monomial_vars m) (\i. ring_pow r ((c o f) i) (m i)))`; `{m:Y->num | ~(q m:R = ring_0 r)}` ]RING_SUM_IMAGE THEN simp[] THEN sufficesby ring_sum_eq_name_d THEN rw[BETA_THM;o_THM;I_THM;IN_ELIM_THM] THEN intro THEN subgoal `!x:Y. ~(d x = 0) ==> x IN B` THENL [ intro THEN subgoal `x:Y IN monomial_vars d` THENL [ rw[monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `x IN poly_vars r (q:(Y->num)->R)` THENL [ rw[poly_vars;IN_UNIONS;IN_ELIM_THM] THEN qed[] ; pass ] THEN qed[SUBSET] ; pass ] THEN subgoal `poly_reindex r q f A B (\a:X. if a IN A then d (f a:Y) else 0) = q d:R` THENL [ rw[poly_reindex;monomial_vars;SUBSET;IN_ELIM_THM] THEN subgoal `(\b:Y. if b IN B then if (@a:X. a IN A /\ f a = b) IN A then d (f (@a. a IN A /\ f a = b)) else 0 else 0) = d` THENL [ rw[FUN_EQ_THM] THEN intro THEN case `x:Y IN B` THENL [ choose `a:X` `a:X IN A /\ f a = x:Y` [BIJ;SURJ] THEN subgoal `(@a:X. a IN A /\ f a = x:Y) = a` THENL [ sufficesby SELECT_UNIQUE THEN qed[BIJ;INJ] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN subgoal `monomial_vars (d:Y->num) = IMAGE f (monomial_vars (\a:X. if a IN A then d(f a) else 0))` THENL [ rw[EXTENSION;in_image_vw;monomial_vars;IN_ELIM_THM] THEN intro THEN splitiff THENL [ intro THEN choose `a:X` `a:X IN A /\ f a = x:Y` [BIJ;SURJ] THEN witness `a:X` THEN qed[] ; qed[] ] ; pass ] THEN simp[] THEN subgoal `!x y:X. x IN monomial_vars (\a. if a IN A then d (f a:Y) else 0) ==> y IN monomial_vars (\a:X. if a IN A then d (f a) else 0) ==> f x = f y ==> x = y` THENL [ rw[IN_ELIM_THM;fun_eq_thm_e;monomial_vars] THEN intro THEN qed[BIJ;INJ] ; pass ] THEN specialize[ `r:R ring`; `f:X->Y`; `\i:Y. ring_pow(r:R ring) (c i) (d i)`; `monomial_vars (\a:X. if a IN A then d (f a:Y) else 0)` ]RING_PRODUCT_IMAGE THEN simp[] THEN subgoal `ring_product r (monomial_vars (\a:X. if a IN A then d (f a:Y) else 0)) (\i. ring_pow r (c (f i):R) (if i IN A then d (f i) else 0)) = ring_product r (monomial_vars (\a. if a IN A then d (f a) else 0)) ((\i. ring_pow r (c i) (d i)) o f)` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM;o_THM;monomial_vars;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[] );; let poly_reindex_subring = prove(` !(r:R ring) G A B q f:X->Y. poly_reindex(subring_generated r G) q f A B = poly_reindex r q f A B `, rw[poly_reindex] THEN rw[SUBRING_GENERATED] );; let symmetric_subring_if_poly_subring = prove(` !(r:R ring) G S c:X->R q. FINITE S ==> (!s. s IN S ==> c s IN ring_carrier r) ==> (!d. coeff d (monic_vanishing_at r S c) IN ring_carrier(subring_generated r G)) ==> ring_polynomial(subring_generated r G) q ==> poly_vars r q SUBSET S ==> (!f m. f permutes S ==> q(m o f) = q(m)) ==> poly_evaluate r q c IN ring_carrier(subring_generated r G) `, intro THEN choose `f:num->X` `IMAGE f (range(CARD S)) = S:X->bool` [finite_ordering] THEN def `p:(num->num)->R` `poly_reindex (r:R ring) q (f:num->X) (range(CARD S)) (S:X->bool)` THEN subgoal `!i. i < CARD(S:X->bool) ==> c(f i:X) IN ring_carrier(r:R ring)` THENL [ intro THEN have `i IN range(CARD(S:X->bool))` [range_lt] THEN have `f(i:num) IN (S:X->bool)` [IN_IMAGE] THEN qed[] ; pass ] THEN subgoal `!d. coeff d (monic_vanishing_at r (range(CARD(S:X->bool))) (\i:num. c(f i:X):R)) IN ring_carrier(subring_generated r G)` THENL [ intro THEN have `(\i:num. c(f i:X):R) = c o f` [FUN_EQ_THM;o_THM] THEN specialize[ `S:X->bool`; `f:num->X` ]injective_finite_ordering THEN have `!x y. x IN range(CARD(S:X->bool)) ==> y IN range(CARD S) ==> f x:X = f y ==> x = y` [range_lt] THEN specialize[ `r:R ring`; `range(CARD(S:X->bool))`; `f:num->X`; `c:X->R` ]monic_vanishing_at_image THEN qed[] ; pass ] THEN have `p:(num->num)->R = poly_reindex(subring_generated r G) q (f:num->X) (range(CARD S)) S` [poly_reindex_subring] THEN have `poly_vars(subring_generated r G) (q:(X->num)->R) SUBSET S` [poly_vars_subring] THEN have `BIJ(f:num->X) (range(CARD S)) (S:X->bool)` [bij_finite_ordering] THEN have `ring_polynomial(subring_generated r G) (p:(num->num)->R)` [polynomial_poly_reindex] THEN specialize[ `subring_generated (r:R ring) G`; `range(CARD(S:X->bool))`; `S:X->bool`; `q:(X->num)->R`; `f:num->X` ]poly_vars_poly_reindex THEN have `poly_vars r (p:(num->num)->R) SUBSET range(CARD(S:X->bool))` [poly_vars_subring] THEN specialize[ `subring_generated (r:R ring) G`; `range(CARD(S:X->bool))`; `S:X->bool`; `q:(X->num)->R`; `f:num->X` ]poly_reindex_permutes THEN have `!f m:num->num. f permutes range (CARD(S:X->bool)) ==> p (m o f) = p m:R` [] THEN specialize[ `r:R ring`; `G:R->bool`; `CARD(S:X->bool)`; `\i:num. c(f i:X):R`; `p:(num->num)->R` ]symmetric_subring_if_poly_subring_range THEN have `ring_polynomial r (q:(X->num)->R)` [ring_polynomial_if_subring] THEN specialize[ `r:R ring`; `range(CARD(S:X->bool))`; `S:X->bool`; `q:(X->num)->R`; `f:num->X`; `c:X->R` ](GSYM poly_evaluate_poly_reindex) THEN simp[o_DEF] THEN qed[] );; (* ===== ring_ord *) (* XXX: should factor poly_ord and some of the UFD theorems above via this *) let ring_ord = new_definition ` ring_ord (r:R ring) (p:R) (a:R) = @e:num. ( ?u:R. ( u IN ring_carrier r /\ ~(ring_divides r p u) /\ a = ring_mul r (ring_pow r p e) u ) ) `;; (* could prove via ring_ord_unit *) let ring_ord_1 = prove(` !(r:R ring) p. p IN ring_carrier r ==> ~(ring_unit r p) ==> ring_ord r p (ring_1 r) = 0 `, intro THEN rw[ring_ord] THEN sufficesby SELECT_UNIQUE THEN rw[BETA_THM] THEN intro THEN splitiff THENL [ intro THEN proven_if `y = 0` [] THEN num_linear_fact `~(y = 0) ==> y = 1+(y-1)` THEN have `ring_pow(r:R ring) p y = ring_mul r (ring_pow r p 1) (ring_pow r p (y-1))` [RING_POW_ADD] THEN have `ring_1(r:R ring) = ring_mul r (ring_mul r p (ring_pow r p (y-1))) u` [RING_POW_1] THEN have `ring_1(r:R ring) = ring_mul r p (ring_mul r (ring_pow r p (y-1)) u)` [RING_MUL_ASSOC;RING_POW] THEN subgoal `ring_divides(r:R ring) p (ring_1 r)` THENL [ rw[ring_divides] THEN simp[RING_MUL;RING_POW] THEN witness `ring_mul(r:R ring) (ring_pow r p (y-1)) u` THEN simp[RING_MUL;RING_POW] THEN RING_TAC ; pass ] THEN qed[ring_unit;ring_divides] ; intro THEN witness `ring_1(r:R ring)` THEN qed[RING_1;ring_unit;ring_divides;RING_POW_0;RING_MUL_LID] ] );; let ring_ord_notdivides = prove(` !(r:R ring) p v. p IN ring_carrier r ==> v IN ring_carrier r ==> ~(ring_divides r p v) ==> ring_ord r p v = 0 `, intro THEN rw[ring_ord] THEN sufficesby SELECT_UNIQUE THEN rw[BETA_THM] THEN intro THEN have `~(ring_unit r (p:R))` [RING_UNIT_DIVIDES_ANY] THEN splitiff THENL [ intro THEN proven_if `y = 0` [] THEN num_linear_fact `~(y = 0) ==> y = 1+(y-1)` THEN have `ring_pow(r:R ring) p y = ring_mul r (ring_pow r p 1) (ring_pow r p (y-1))` [RING_POW_ADD] THEN have `v:R = ring_mul r (ring_mul r p (ring_pow r p (y-1))) u` [RING_POW_1] THEN have `v:R = ring_mul r p (ring_mul r (ring_pow r p (y-1)) u)` [RING_MUL_ASSOC;RING_POW] THEN subgoal `ring_divides(r:R ring) p v` THENL [ rw[ring_divides] THEN simp[RING_MUL;RING_POW] THEN witness `ring_mul(r:R ring) (ring_pow r p (y-1)) u` THEN simp[RING_MUL;RING_POW] THEN have `v IN ring_carrier(r:R ring)` [ring_unit] THEN RING_TAC ; pass ] THEN subgoal `ring_divides(r:R ring) v (ring_1 r)` THENL [ rw[ring_divides] THEN qed[ring_unit;RING_1] ; pass ] THEN qed[] ; intro THEN witness `v:R` THEN intro THENL [ qed[] ; qed[] ; qed[RING_POW_0;RING_MUL_LID] ] ] );; let ring_ord_unit = prove(` !(r:R ring) p v. p IN ring_carrier r ==> ~(ring_unit r p) ==> ring_unit r v ==> ring_ord r p v = 0 `, intro THEN have `v:R IN ring_carrier r` [ring_unit] THEN have `~ring_divides r (p:R) v` [RING_UNIT_DIVISOR] THEN qed[ring_ord_notdivides] );; let ring_ord_pow_refl = prove(` !(r:R ring) p e. integral_domain r ==> p IN ring_carrier r ==> ~(ring_unit r p) ==> ~(p = ring_0 r) ==> ring_ord r p (ring_pow r p e) = e `, intro THEN rw[ring_ord] THEN sufficesby SELECT_UNIQUE THEN rw[BETA_THM] THEN intro THEN splitiff THENL [ intro THEN case `y < e:num` THENL [ num_linear_fact `y < e:num ==> e = y+((e-1-y)+1)` THEN have `ring_pow r (p:R) e = ring_mul r (ring_pow r p y) (ring_pow r p ((e-1-y)+1))` [RING_POW_ADD] THEN have `ring_pow r (p:R) ((e-1-y)+1) = ring_mul r (ring_pow r p (e-1-y)) (ring_pow r p 1)` [RING_POW_ADD] THEN have `ring_pow r (p:R) 1 = p` [RING_POW_1] THEN havetac `ring_mul r (ring_pow r p y) u:R = ring_mul r (ring_pow r p y) (ring_mul r p (ring_pow r p (e-1-y)))` RING_TAC THEN subgoal `u:R = ring_mul r p (ring_pow r p (e - 1 - y))` THENL [ specialize_assuming[ `r:R ring`; `ring_pow r (p:R) y`; `u:R`; `ring_mul r (p:R) (ring_pow r p (e - 1 - y))` ]INTEGRAL_DOMAIN_MUL_LCANCEL THEN have `~(ring_pow r (p:R) y = ring_0 r)` [INTEGRAL_DOMAIN_POW_EQ_0] THEN qed[RING_MUL;RING_POW] ; pass ] THEN qed[ring_divides;RING_POW] ; pass ] THEN case `e < y:num` THENL [ num_linear_fact `e < y:num ==> y = e+((y-1-e)+1)` THEN have `ring_pow r (p:R) y = ring_mul r (ring_pow r p e) (ring_pow r p ((y-1-e)+1))` [RING_POW_ADD] THEN have `ring_pow r (p:R) ((y-1-e)+1) = ring_mul r (ring_pow r p (y-1-e)) (ring_pow r p 1)` [RING_POW_ADD] THEN have `ring_pow r (p:R) 1 = p` [RING_POW_1] THEN havetac `ring_mul r (ring_pow r p e) (ring_1 r):R = ring_mul r (ring_pow r p e) (ring_mul r p (ring_mul r (ring_pow r p (y-1-e)) u))` RING_TAC THEN subgoal `(ring_1 r):R = ring_mul r p (ring_mul r (ring_pow r p (y-1-e)) u)` THENL [ specialize_assuming[ `r:R ring`; `ring_pow r (p:R) e`; `ring_1 r:R`; `ring_mul r (p:R) (ring_mul r (ring_pow r p (y - 1 - e)) u)` ]INTEGRAL_DOMAIN_MUL_LCANCEL THEN have `~(ring_pow r (p:R) e = ring_0 r)` [INTEGRAL_DOMAIN_POW_EQ_0] THEN qed[RING_MUL;RING_POW;RING_1] ; pass ] THEN qed[ring_unit;RING_MUL;RING_POW] ; pass ] THEN ASM_ARITH_TAC ; intro THEN witness `ring_1 r:R` THEN qed[RING_1;ring_unit;ring_divides;RING_MUL_RID;RING_POW] ] );; let ring_ord_refl = prove(` !(r:R ring) p. integral_domain r ==> p IN ring_carrier r ==> ~(ring_unit r p) ==> ~(p = ring_0 r) ==> ring_ord r p p = 1 `, qed[ring_ord_pow_refl;RING_POW_1] );; let ring_ord_exists = prove(` !(r:R ring) p. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> p IN ring_carrier r ==> ~(ring_unit r p) ==> !a. a IN ring_carrier r ==> ~(a = ring_0 r) ==> ?e u:R. ( u IN ring_carrier r /\ ~(ring_divides r p u) /\ a = ring_mul r (ring_pow r p e) u ) `, intro_genonly THEN REPEAT DISCH_TAC THEN specialize[`r:R ring`]RING_DIVIDES_WF THEN specialize_forward[`\x y:R. ring_divides r x y /\ ~ring_divides r y x`](GEN `(<<):A->A->bool` WF_IND) THEN sufficesby (know `!P. (!x:R. (!y. ring_divides r y x ==> ~ring_divides r x y ==> P y) ==> P x) ==> (!x. P x)`) THEN intro THEN case `ring_divides r (p:R) a` THENL [ choose `b:R` `b:R IN ring_carrier r /\ a = ring_mul r p b` [ring_divides] THEN have `ring_divides r (b:R) a` [ring_divides;RING_MUL_SYM] THEN have `~(b:R = ring_0 r)` [RING_MUL_RZERO] THEN subgoal `~ring_divides r (a:R) b` THENL [ intro THEN have `ring_associates r (a:R) b` [ring_associates] THEN qed[INTEGRAL_DOMAIN_DIVIDES_ASSOCIATES_MUL_SELF] ; pass ] THEN have `b:R IN ring_carrier r` [] THEN specialize[`b:R`](know `!y:R. ring_divides r y a /\ ~ring_divides r a y ==> y IN ring_carrier r ==> ~(y = ring_0 r) ==> (?e u. u IN ring_carrier r /\ ~ring_divides r p u /\ y = ring_mul r (ring_pow r p e) u)`) THEN choose2 `e:num` `u:R` `u:R IN ring_carrier r /\ ~ring_divides r p u /\ b = ring_mul r (ring_pow r p e) u` [] THEN witness `e+1` THEN witness `u:R` THEN intro THENL [ qed[] ; qed[] ; have `b:R = ring_mul r (ring_pow r p e) u` [] THEN have `ring_pow r (p:R) (e+1) = ring_mul r (ring_pow r p e) (ring_pow r p 1)` [RING_POW_ADD] THEN have `ring_pow r (p:R) (e+1) = ring_mul r (ring_pow r p e) p` [RING_POW_1] THEN specialize_assuming[ `ring_pow r (p:R) e`; `ring_pow r (p:R) (e+1)`; ]( GENL[`E:R`;`E1:R`]( RING_RULE `E1:R = ring_mul r E p ==> a = ring_mul r p b ==> b = ring_mul r E u ==> a = ring_mul r E1 u` )) THEN qed[RING_POW] ] ; witness `0` THEN witness `a:R` THEN qed[RING_POW_0;RING_MUL_LID] ] );; let ring_ord_unique_lemma = prove(` !(r:R ring) p e f u v. integral_domain r ==> p IN ring_carrier r ==> u IN ring_carrier r ==> v IN ring_carrier r ==> ring_prime r p ==> ~(ring_divides r p v) ==> ring_mul r (ring_pow r p e) u = ring_mul r (ring_pow r p f) v ==> e <= f `, intro THEN proven_if `e <= f:num` [] THEN num_linear_fact `~(e <= f) ==> e = f+(1+(e-1-f))` THEN subgoal `ring_pow r (p:R) e = ring_mul r (ring_pow r p f) (ring_mul r (ring_pow r p 1) (ring_pow r p (e-1-f)))` THENL [ simp[GSYM RING_POW_ADD] THEN qed[] ; pass ] THEN have `ring_pow r (p:R) e = ring_mul r (ring_pow r p f) (ring_mul r p (ring_pow r p (e-1-f)))` [RING_POW_1] THEN subgoal `ring_mul r (ring_pow r p f:R) (ring_mul r p (ring_mul r (ring_pow r p (e - 1 - f)) u)) = ring_mul r (ring_pow r p f) v` THENL [ specialize_assuming[ `ring_pow r (p:R) e`; `ring_pow r (p:R) (e-1-f)`; `ring_pow r (p:R) f`; ](GENL[`E:R`;`E1F:R`;`X:R`]( RING_RULE `ring_mul r (E:R) u = ring_mul r X v ==> E = ring_mul r X (ring_mul r p E1F) ==> ring_mul r X (ring_mul r p (ring_mul r E1F u)) = ring_mul r X v` )) THEN qed[RING_POW;RING_MUL] ; pass ] THEN subgoal `ring_mul r (p:R) (ring_mul r (ring_pow r p (e - 1 - f)) u) = v` THENL [ specialize_assuming[ `r:R ring`; `ring_pow r (p:R) f`; `ring_mul r (p:R) (ring_mul r (ring_pow r p (e - 1 - f)) u)`; `v:R` ]INTEGRAL_DOMAIN_MUL_LCANCEL THEN have `~(ring_pow r (p:R) f = ring_0 r)` [INTEGRAL_DOMAIN_POW_EQ_0;ring_prime] THEN qed[RING_MUL;RING_POW] ; pass ] THEN qed[ring_divides;RING_POW;RING_MUL] );; let ring_ord_unique = prove(` !(r:R ring) p e f u v. integral_domain r ==> p IN ring_carrier r ==> u IN ring_carrier r ==> v IN ring_carrier r ==> ring_prime r p ==> ~(ring_divides r p u) ==> ~(ring_divides r p v) ==> ring_mul r (ring_pow r p e) u = ring_mul r (ring_pow r p f) v ==> e = f `, intro THEN have `e <= f:num` [ring_ord_unique_lemma] THEN have `f <= e:num` [ring_ord_unique_lemma] THEN ASM_ARITH_TAC );; let ring_ord_exists_unique = prove(` !(r:R ring) p a. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> ?u:R. ( u IN ring_carrier r /\ ~(ring_divides r p u) /\ a = ring_mul r (ring_pow r p (ring_ord r p a)) u ) `, intro_genonly THEN REPEAT DISCH_TAC THEN have `p:R IN ring_carrier r` [ring_prime] THEN have `~ring_unit r (p:R)` [ring_prime] THEN specialize[ `r:R ring`; `p:R` ]ring_ord_exists THEN choose2 `e:num` `v:R` `v:R IN ring_carrier r /\ ~ring_divides r p v /\ a = ring_mul r (ring_pow r p e) v` [] THEN subgoal `(@e. ?u:R. u IN ring_carrier r /\ ~ring_divides r p u /\ a = ring_mul r (ring_pow r p e) u) = e` THENL [ sufficesby SELECT_UNIQUE THEN qed[ring_ord_unique] ; pass ] THEN witness `v:R` THEN rw[ring_ord] THEN qed[] );; let divides_pow_ring_ord = prove(` !(r:R ring) p a. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> ring_divides r (ring_pow r p (ring_ord r p a)) a `, intro_gendisch THEN have `~ring_unit r (p:R)` [ring_prime] THEN specialize[ `r:R ring`; `p:R`; `a:R` ]ring_ord_exists_unique THEN choose `u:R` `u:R IN ring_carrier r /\ ~ring_divides r p u /\ a = ring_mul r (ring_pow r p (ring_ord r p a)) u` [] THEN rw[ring_divides] THEN qed[RING_POW;ring_prime] );; let divides_le_pow_ring_ord = prove(` !(r:R ring) p a e. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> e <= ring_ord r p a ==> ring_divides r (ring_pow r p e) a `, intro_gendisch THEN have `ring_divides r (ring_pow r p (ring_ord r p a)) (a:R)` [divides_pow_ring_ord] THEN qed[ring_divides_pow_pow;ring_prime;RING_DIVIDES_TRANS] );; let nonzero_ring_ord_if_divides = prove(` !(r:R ring) p a. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> ~(a = ring_0 r) ==> ring_divides r p a ==> ~(ring_ord r p a = 0) `, intro_gendisch THEN have `~ring_unit r (p:R)` [ring_prime] THEN have `p:R IN ring_carrier r` [ring_prime] THEN have `a:R IN ring_carrier r` [ring_divides] THEN specialize[ `r:R ring`; `p:R`; `a:R` ]ring_ord_exists_unique THEN choose `u:R` `u:R IN ring_carrier r /\ ~ring_divides r p u /\ a = ring_mul r (ring_pow r p (ring_ord r p a)) u` [] THEN have `a = ring_mul r (ring_pow r p 0) u:R` [] THEN have `a = ring_mul r (ring_1 r) u:R` [RING_POW_0] THEN have `a = u:R` [RING_MUL_LID] THEN qed[] );; let ring_ord_mul = prove(` !(r:R ring) p a b. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> b IN ring_carrier r ==> ~(b = ring_0 r) ==> ring_ord r p (ring_mul r a b) = (ring_ord r p a) + (ring_ord r p b) `, intro_genonly THEN REPEAT DISCH_TAC THEN specialize[ `r:R ring`; `p:R`; `a:R` ]ring_ord_exists_unique THEN choose `u:R` `u:R IN ring_carrier r /\ ~(ring_divides r p u) /\ a = ring_mul r (ring_pow r p (ring_ord r p a)) u` [] THEN specialize[ `r:R ring`; `p:R`; `b:R` ]ring_ord_exists_unique THEN choose `v:R` `v:R IN ring_carrier r /\ ~(ring_divides r p v) /\ b = ring_mul r (ring_pow r p (ring_ord r p b)) v` [] THEN have `ring_mul r a b:R IN ring_carrier r` [RING_MUL] THEN have `~(ring_mul r a b:R = ring_0 r)` [integral_domain] THEN specialize[ `r:R ring`; `p:R`; `ring_mul r (a:R) b` ]ring_ord_exists_unique THEN choose `w:R` `w:R IN ring_carrier r /\ ~(ring_divides r p w) /\ ring_mul r a b = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r a b))) w` [] THEN subgoal `ring_mul r (ring_pow r p (ring_ord r p a + ring_ord r p b)) (ring_mul r u v) = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r a b))) w:R` THENL [ have `a:R = ring_mul r (ring_pow r p (ring_ord r p a)) u` [] THEN have `b:R = ring_mul r (ring_pow r p (ring_ord r p b)) v` [] THEN have `ring_mul r a b:R = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r a b))) w` [] THEN have `p:R IN ring_carrier r` [ring_prime] THEN simp[RING_POW_ADD] THEN RING_TAC ; pass ] THEN specialize_assuming[ `r:R ring`; `p:R`; `ring_ord r (p:R) a + ring_ord r (p:R) b`; `ring_ord r (p:R) (ring_mul r a b)`; `ring_mul r u v:R`; `w:R` ]ring_ord_unique THEN qed[ring_prime;RING_MUL] );; let ring_ord_pow = prove(` !(r:R ring) p a e. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> ring_ord r p (ring_pow r a e) = e * ring_ord r p a `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[RING_POW_0] THEN intro_gendisch THEN have `ring_ord r (p:R) (ring_1 r) = 0` [ring_ord_1;ring_prime] THEN ASM_ARITH_TAC ; rw[ARITH_RULE `SUC e = 1+e`] THEN intro_gendisch THEN simp[RING_POW_ADD] THEN subgoal `ring_ord r p (ring_mul r (ring_pow r a 1) (ring_pow r a e)) = ring_ord r p (ring_pow r a 1) + ring_ord r p (ring_pow r a e:R)` THENL [ specialize_assuming[ `r:R ring`; `p:R`; `ring_pow r (a:R) 1`; `ring_pow r (a:R) e` ]ring_ord_mul THEN qed[RING_POW;INTEGRAL_DOMAIN_POW_EQ_0] ; pass ] THEN simp[RING_POW_1] THEN ARITH_TAC ] );; let ring_ord_product_waterfall = prove(` !(r:R ring) p a. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> !S. FINITE S ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> (!s:X. s IN S ==> ~(a s = ring_0 r)) ==> ring_ord r p (ring_product r S a) = nsum S (\s. ring_ord r p (a s)) `, intro_gendisch THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;NSUM_CLAUSES] THEN qed[ring_ord_1;ring_prime] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN simp[RING_PRODUCT_CLAUSES;NSUM_CLAUSES] THEN specialize_assuming[ `r:R ring`; `p:R`; `a(x:X):R`; `ring_product r S (a:X->R)` ]ring_ord_mul THEN qed[RING_PRODUCT;INTEGRAL_DOMAIN_PRODUCT_EQ_0] ] );; let ring_ord_product = prove(` !(r:R ring) p S a. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> FINITE S ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> (!s:X. s IN S ==> ~(a s = ring_0 r)) ==> ring_ord r p (ring_product r S a) = nsum S (\s. ring_ord r p (a s)) `, intro THEN specialize_assuming[ `r:R ring`; `p:R`; `a:X->R` ]ring_ord_product_waterfall THEN qed[] );; let ring_ord_divides = prove(` !(r:R ring) p a b. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> ~(b = ring_0 r) ==> ring_divides r a b ==> ring_ord r p a <= ring_ord r p b `, rw[ring_divides] THEN intro_genonly THEN DISCH_TAC THEN DISCH_TAC THEN intro THEN have `~(a:R = ring_0 r)` [RING_MUL_LZERO] THEN have `~(x:R = ring_0 r)` [RING_MUL_RZERO] THEN have `ring_ord r (p:R) b = ring_ord r p a + ring_ord r p x` [ring_ord_mul] THEN qed[ARITH_RULE `e <= e+f:num`] );; let ring_ord_prime_divides = prove(` !(r:R ring) p b. integral_domain r ==> (noetherian_ring r \/ UFD r) ==> ring_prime r p ==> ~(b = ring_0 r) ==> ring_divides r p b ==> 1 <= ring_ord r p b `, intro_gendisch THEN have `ring_ord r p (p:R) = 1` [ring_ord_refl;ring_prime] THEN qed[ring_ord_divides] );; let ring_ord_gcd = prove(` !(r:R ring) p a b. UFD r ==> ring_prime r p ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> b IN ring_carrier r ==> ~(b = ring_0 r) ==> ring_ord r p (ring_gcd r (a,b)) = MIN (ring_ord r p a) (ring_ord r p b) `, intro_gendisch THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN have `ring_divides r (ring_gcd r (a,b)) (a:R)` [RING_GCD_DIVIDES] THEN have `ring_divides r (ring_gcd r (a,b)) (b:R)` [RING_GCD_DIVIDES] THEN have `ring_ord r (p:R) (ring_gcd r (a,b)) <= ring_ord r p a` [ring_ord_divides;RING_GCD] THEN have `ring_ord r (p:R) (ring_gcd r (a,b)) <= ring_ord r p b` [ring_ord_divides;RING_GCD] THEN have `ring_ord r (p:R) (ring_gcd r (a,b)) <= MIN (ring_ord r p a) (ring_ord r p b)` [le_min] THEN have `MIN (ring_ord r p a) (ring_ord r p b) <= ring_ord r p (a:R)` [min_le] THEN have `MIN (ring_ord r p a) (ring_ord r p b) <= ring_ord r p (b:R)` [min_le] THEN have `ring_divides r (ring_pow r (p:R) (MIN (ring_ord r p a) (ring_ord r p b))) a` [divides_le_pow_ring_ord] THEN have `ring_divides r (ring_pow r (p:R) (MIN (ring_ord r p a) (ring_ord r p b))) b` [divides_le_pow_ring_ord] THEN have `ring_divides r (ring_pow r (p:R) (MIN (ring_ord r p a) (ring_ord r p b))) (ring_gcd r (a,b))` [UFD_DIVIDES_GCD] THEN have `~(ring_gcd r (a,b) = ring_0 r:R)` [ring_divides;RING_MUL_LZERO] THEN have `ring_ord r p (ring_pow r (p:R) (MIN (ring_ord r p a) (ring_ord r p b))) <= ring_ord r p (ring_gcd r (a,b))` [ring_ord_divides] THEN have `MIN (ring_ord r p a) (ring_ord r p b) <= ring_ord r (p:R) (ring_gcd r (a,b))` [ring_ord_pow_refl;ring_prime] THEN ASM_ARITH_TAC );; let ring_ord_associates = prove(` !(r:R ring) p a b. UFD r ==> ring_prime r p ==> ring_associates r a b ==> ~(a = ring_0 r) ==> ring_ord r p a = ring_ord r p b `, intro THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN choose `u:R` `ring_unit r (u:R) /\ ring_mul r a u = b` [INTEGRAL_DOMAIN_ASSOCIATES] THEN have `ring_ord r p (u:R) = 0` [ring_ord_unit;ring_prime] THEN specialize_assuming[ `r:R ring`; `p:R`; `a:R`; `u:R` ]ring_ord_mul THEN have `~(b = ring_0 r:R)` [RING_ASSOCIATES_0] THEN have `a:R IN ring_carrier r` [ring_associates;ring_divides] THEN have `~(u = ring_0 r:R)` [RING_MUL_RZERO] THEN have `ring_ord r (p:R) (ring_mul r a u) = ring_ord r p a + ring_ord r p u` [ring_unit] THEN qed[ARITH_RULE `e+0 = e`] );; let ring_associates_ord = prove(` !(r:R ring) p q a. UFD r ==> ring_prime r p ==> ring_associates r p q ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> ring_ord r p a = ring_ord r q a `, intro THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN choose `u:R` `ring_unit r (u:R) /\ ring_mul r q u = p` [RING_ASSOCIATES_SYM;INTEGRAL_DOMAIN_ASSOCIATES] THEN have `noetherian_ring r \/ UFD(r:R ring)` [] THEN specialize[ `r:R ring`; `p:R`; `a:R` ]ring_ord_exists_unique THEN choose `v:R` `v:R IN ring_carrier r /\ ~ring_divides r p v /\ a = ring_mul r (ring_pow r p (ring_ord r p a)) v` [] THEN subgoal `a:R = ring_mul r (ring_pow r q (ring_ord r p a)) (ring_mul r (ring_pow r u (ring_ord r p a)) v)` THENL [ have `p:R IN ring_carrier r` [ring_prime] THEN have `q:R IN ring_carrier r` [ring_associates;ring_divides] THEN have `u:R IN ring_carrier r` [ring_unit] THEN have `ring_pow r (p:R) (ring_ord r p a) = ring_mul r (ring_pow r q (ring_ord r p a)) (ring_pow r u (ring_ord r p a))` [RING_MUL_POW] THEN have `a = ring_mul r (ring_pow r p (ring_ord r p a)) v:R` [] THEN RING_TAC ; pass ] THEN have `ring_unit r (ring_pow r u (ring_ord r p a):R)` [RING_UNIT_POW] THEN have `ring_associates r (v:R) (ring_mul r (ring_pow r u (ring_ord r p a)) v)` [RING_ASSOCIATES_LMUL] THEN have `~ring_divides r (q:R) (ring_mul r (ring_pow r u (ring_ord r p a)) v)` [RING_ASSOCIATES_DIVIDES] THEN have `ring_prime r (q:R)` [RING_ASSOCIATES_PRIME] THEN specialize[ `r:R ring`; `q:R`; `a:R` ]ring_ord_exists_unique THEN choose `w:R` `w:R IN ring_carrier r /\ ~ring_divides r q w /\ a = ring_mul r (ring_pow r q (ring_ord r q a)) w` [] THEN specialize_assuming[ `r:R ring`; `q:R`; `ring_ord r p (a:R)`; `ring_ord r q (a:R)`; `ring_mul r (ring_pow r u (ring_ord r p a)) v:R`; `w:R` ]ring_ord_unique THEN qed[ring_prime;RING_MUL;RING_POW;ring_unit] );; let ring_ord_prime = prove(` !(r:R ring) p q. UFD r ==> ring_prime r p ==> ring_prime r q ==> ring_ord r p q = if ring_associates r p q then 1 else 0 `, intro THEN case `ring_divides r (p:R) q` THENL [ specialize_assuming[ `r:R ring`; `p:R`; `q:R`; `p:R` ]prime_divides_prime_and THEN have `ring_divides r q (p:R)` [RING_DIVIDES_REFL;ring_prime] THEN have `ring_associates r p (q:R)` [ring_associates] THEN have `ring_ord r (p:R) p = 1` [ring_ord_refl;UFD_IMP_INTEGRAL_DOMAIN;ring_prime] THEN qed[ring_ord_associates;ring_prime] ; pass ] THEN qed[ring_ord_notdivides;ring_prime;ring_associates] );; let primefact_ord_waterfall = prove(` !(r:R ring). UFD r ==> !a. a IN ring_carrier r ==> ~(a = ring_0 r) ==> (?n q. (!i. i IN (1..n) ==> ring_prime r (q i)) /\ (!i. i IN (1..n) ==> ring_divides r (q i) a) /\ (!i j. i IN (1..n) ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))) ) `, intro_gendisch THEN sufficesby UFD_PRIME_FACTOR_INDUCT THEN intro THENL [ qed[] ; qed[] ; witness `0` THEN witness `\i:num. ring_0 r:R` THEN have `!i. ~(i IN 1..0)` [IN_NUMSEG;ARITH_RULE `~(1 <= i /\ i <= 0)`] THEN set_fact `(!i. ~(i IN 1..0)) ==> 1..0 = {}` THEN rw[know `1..0 = {}`] THEN rw[RING_PRODUCT_CLAUSES;EMPTY;IN_ELIM_THM] THEN qed[RING_ASSOCIATES_1] ; have `~(a = ring_0 r:R)` [RING_MUL_RZERO;ring_prime] THEN recall (know `~(a = ring_0 r:R) ==> (?n q. (!i. i IN 1..n ==> ring_prime r (q i)) /\ (!i. i IN 1..n ==> ring_divides r (q i) a) /\ (!i j. i IN 1..n ==> j IN 1..n ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))))`) THEN choose2 `n:num` `q:num->R` `(!i. i IN (1..n) ==> ring_prime r (q i:R)) /\ (!i. i IN (1..n) ==> ring_divides r (q i) a) /\ (!i j. i IN (1..n) ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)))` [] THEN case `!i. i IN 1..n ==> ~ring_associates r p (q i:R)` THENL [ witness `n+1` THEN witness `\i:num. if i = n+1 then p else q i:R` THEN subgoal `!i. i IN (1..n+1) ==> ~(i = n+1) ==> i IN 1..n` THENL [ intro THEN have `i <= n+1` [IN_NUMSEG] THEN have `1 <= i` [IN_NUMSEG] THEN num_linear_fact `i <= n+1 ==> ~(i = n+1) ==> i <= n` THEN qed[IN_NUMSEG] ; pass ] THEN rw[BETA_THM] THEN intro THENL [ proven_if `i = n+1` [] THEN qed[] ; proven_if `i = n+1` [ring_divides;ring_prime;RING_MUL] THEN have `ring_divides r (q(i:num):R) a` [] THEN qed[ring_prime;RING_MUL;RING_DIVIDES_LMUL] ; case `i = n+1` THENL [ proven_if `j = n+1` [] THEN have `j IN 1..n` [] THEN qed[] ; pass ] THEN case `j = n+1` THENL [ proven_if `i = n+1` [] THEN have `i IN 1..n` [] THEN qed[RING_ASSOCIATES_SYM] ; pass ] THEN qed[] ; subgoal `ring_product r (1..n + 1) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a))) = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r p a))) (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))):R` THENL [ have `FINITE(1..n+1)` [FINITE_NUMSEG] THEN have `n+1 IN 1..n+1` [IN_NUMSEG;ARITH_RULE `1 <= n+1 /\ n+1 <= n+1`] THEN have `ring_pow r (p:R) (ring_ord r p (ring_mul r p a)) IN ring_carrier r` [RING_POW;ring_prime] THEN specialize[ `r:R ring`; `1..n+1`; `n+1`; `\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i:R) (ring_mul r p a))` ]ring_product_delete THEN rw[know `ring_product r (1..n + 1) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a))) = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r p a))) (ring_product r ((1..n + 1) DELETE (n + 1)) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a)))):R`] THEN subgoal `(1..n+1) DELETE (n+1) = (1..n)` THENL [ rw[EXTENSION;IN_DELETE;IN_NUMSEG] THEN ARITH_TAC ; pass ] THEN rw[know `(1..n+1) DELETE (n+1) = (1..n)`] THEN subgoal `ring_product r (1..n) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a))) = ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)):R` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN have `a' <= n:num` [IN_NUMSEG] THEN num_linear_fact `a' <= n:num ==> ~(a' = n+1)` THEN have `ring_ord r (q(a':num):R) p = 0` [ring_ord_prime;RING_ASSOCIATES_SYM] THEN specialize_assuming[ `r:R ring`; `q(a':num):R`; `p:R`; `a:R` ]ring_ord_mul THEN have `ring_ord r (q(a':num):R) (ring_mul r p a) = ring_ord r (q a') p + ring_ord r (q a') a` [UFD_IMP_INTEGRAL_DOMAIN;ring_prime] THEN have `ring_ord r (q(a':num):R) (ring_mul r p a) = 0 + ring_ord r (q a') a` [] THEN have `ring_ord r (q(a':num):R) (ring_mul r p a) = ring_ord r (q a') a` [ARITH_RULE `0+e = e`] THEN qed[] ; pass ] THEN rw[know `ring_product r (1..n) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a))) = ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)):R`] ; pass ] THEN rw[know `ring_product r (1..n + 1) (\i. ring_pow r (if i = n + 1 then p else q i) (ring_ord r (if i = n + 1 then p else q i) (ring_mul r p a))) = ring_mul r (ring_pow r p (ring_ord r p (ring_mul r p a))) (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))):R`] THEN subgoal `ring_ord r p (ring_mul r p a:R) = 1` THENL [ have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN have `noetherian_ring r \/ UFD (r:R ring)` [] THEN subgoal `ring_ord r p (a:R) = 0` THENL [ subgoal `ring_ord r p (a:R) = nsum (1..n) (\s. ring_ord r p (ring_pow r (q s) (ring_ord r (q s) a)))` THENL [ have `FINITE(1..n)` [FINITE_NUMSEG] THEN have `!s. s IN 1..n ==> ring_pow r (q s) (ring_ord r (q s) a):R IN ring_carrier r` [RING_POW;ring_prime] THEN have `!s. s IN 1..n ==> ~(q s:R = ring_0 r)` [ring_prime] THEN subgoal `!s. s IN 1..n ==> ~(ring_pow r (q s) (ring_ord r (q s) a):R = ring_0 r)` THENL [ intro THEN have `~(q(s:num):R = ring_0 r)` [ring_prime] THEN have `q(s:num):R IN ring_carrier r` [ring_prime] THEN qed[INTEGRAL_DOMAIN_POW_EQ_0] ; pass ] THEN specialize[ `r:R ring`; `p:R`; `1..n`; `\i:num. ring_pow r (q i) (ring_ord r (q i) a):R` ]ring_ord_product THEN qed[ring_ord_associates] ; pass ] THEN subgoal `nsum (1..n) (\s. ring_ord r p (ring_pow r (q s) (ring_ord r (q s:R) a))) = nsum (1..n) (\s. 0)` THENL [ sufficesby NSUM_EQ THEN rw[BETA_THM] THEN intro THEN have `ring_ord r (p:R) (q(x:num)) = 0` [ring_ord_prime] THEN have `q(x:num):R IN ring_carrier r` [ring_prime] THEN have `~(q(x:num):R = ring_0 r)` [ring_prime] THEN specialize[ `r:R ring`; `p:R`; `q(x:num):R`; `ring_ord r (q(x:num):R) a` ]ring_ord_pow THEN qed[ARITH_RULE `e*0 = 0`] ; pass ] THEN qed[NSUM_0] ; pass ] THEN specialize_assuming[ `r:R ring`; `p:R`; `p:R`; `a:R` ]ring_ord_mul THEN have `ring_ord r (p:R) p = 1` [ring_ord_refl;ring_prime] THEN qed[ARITH_RULE `1+0 = 1`;ring_prime] ; pass ] THEN rw[know `ring_ord r (p:R) (ring_mul r p a) = 1`] THEN qed[RING_POW_1;RING_ASSOCIATES_MUL;RING_ASSOCIATES_REFL;RING_PRODUCT;RING_POW;ring_prime] ] ; pass ] THEN choose `b:num` `b IN 1..n /\ ring_associates r p (q b:R)` [] THEN witness `n:num` THEN witness `q:num->R` THEN intro THENL [ qed[] ; have `ring_divides r (q(i:num):R) a` [] THEN qed[RING_DIVIDES_LMUL;ring_prime] ; qed[] ; subgoal `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) (ring_mul r p a))) = ring_product r (1..n) (\i. ring_mul r (ring_pow r (q i) (ring_ord r (q i) p)) (ring_pow r (q i) (ring_ord r (q i) a)):R)` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN have `p:R IN ring_carrier r` [ring_prime] THEN have `~(p:R = ring_0 r)` [ring_prime] THEN have_rw `ring_ord r (q(a':num):R) (ring_mul r p a) = ring_ord r (q a') p + ring_ord r (q a') a` [ring_ord_mul;UFD_IMP_INTEGRAL_DOMAIN] THEN have `q(a':num):R IN ring_carrier r` [ring_prime] THEN qed[RING_POW_ADD] ; pass ] THEN rw[know `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) (ring_mul r p a))) = ring_product r (1..n) (\i. ring_mul r (ring_pow r (q i) (ring_ord r (q i) p)) (ring_pow r (q i) (ring_ord r (q i) a))):R`] THEN subgoal `ring_product r (1..n) (\i. ring_mul r (ring_pow r (q i) (ring_ord r (q i) p)) (ring_pow r (q i) (ring_ord r (q i) a))) = ring_mul r (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) p))) (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))):R` THENL [ have `FINITE(1..n)` [FINITE_NUMSEG] THEN sufficesby RING_PRODUCT_MUL THEN qed[RING_POW;ring_prime] ; pass ] THEN rw[know `ring_product r (1..n) (\i. ring_mul r (ring_pow r (q i) (ring_ord r (q i) p)) (ring_pow r (q i) (ring_ord r (q i) a))) = ring_mul r (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) p))) (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))):R`] THEN subgoal `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) p)) = ring_product r (1..n) (\i. if i = b then q b else ring_1 r):R` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN case `a' = b:num` THENL [ have `ring_associates r (q(a':num):R) p` [RING_ASSOCIATES_SYM] THEN have_rw `ring_ord r (q(a':num):R) p = 1` [ring_ord_prime] THEN qed[RING_POW_1;ring_prime] ; have `~(ring_associates r (q(a':num):R) p)` [RING_ASSOCIATES_SYM;RING_ASSOCIATES_TRANS] THEN have_rw `ring_ord r (q(a':num):R) p = 0` [ring_ord_prime] THEN qed[RING_POW_0] ] ; pass ] THEN rw[know `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) p)) = ring_product r (1..n) (\i. if i = b then q b else ring_1 r):R`] THEN subgoal `ring_product r (1..n) (\i. if i = b then q b else ring_1 r) = q b:R` THENL [ rw[RING_PRODUCT_DELTA] THEN qed[ring_prime] ; pass ] THEN rw[know `ring_product r (1..n) (\i. if i = b then q b else ring_1 r) = q b:R`] THEN qed[RING_ASSOCIATES_MUL] ] ] );; let primefact_ord = prove(` !(r:R ring) a. UFD r ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> (?n q. (!i. i IN (1..n) ==> ring_prime r (q i)) /\ (!i. i IN (1..n) ==> ring_divides r (q i) a) /\ (!i j. i IN (1..n) ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))) ) `, simp[primefact_ord_waterfall] );; let primefact_divides = prove(` !(r:R ring) n q e b. UFD r ==> (!i. i IN 1..n ==> ring_prime r (q i)) ==> (!i j. i IN 1..n ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) ==> b IN ring_carrier r ==> ~(b = ring_0 r) ==> (!i. i IN 1..n ==> e i <= ring_ord r (q i) b) ==> ring_divides r ( ring_product r (1..n) (\i. ring_pow r (q i) (e i)) ) b `, intro THEN have `FINITE(1..n)` [FINITE_NUMSEG] THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN have `!i. i IN 1..n ==> ring_divides r (ring_pow r (q i) (e i)) (b:R)` [divides_le_pow_ring_ord] THEN subgoal `!i j. i IN 1..n ==> j IN 1..n ==> ~(i = j) ==> ring_coprime r (ring_pow r (q i:R) (e i),ring_pow r (q j) (e j))` THENL [ intro THEN proven_if `ring_coprime r (q(i:num):R,q j)` [ring_coprime_lrpow] THEN qed[ring_coprime_associates_prime] ; pass ] THEN specialize_assuming[ `r:R ring`; `\i:num. ring_pow r (q i) (e i):R`; `b:R`; `1..n` ]ring_product_divides_if_coprime THEN qed[] );; let ring_ord_divides_eq = prove(` !(r:R ring) a b. UFD r ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> b IN ring_carrier r ==> ~(b = ring_0 r) ==> ( ring_divides r a b <=> (!p. ring_prime r p ==> ring_ord r p a <= ring_ord r p b ) ) `, intro_gendisch THEN splitiff THENL [ qed[ring_ord_divides;UFD_IMP_INTEGRAL_DOMAIN] ; specialize[`r:R ring`;`a:R`]primefact_ord THEN intro THEN choose2 `n:num` `q:num->R` `(!i. i IN (1..n) ==> ring_prime r (q i:R)) /\ (!i j. i IN (1..n) ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)))` [primefact_ord] THEN specialize_assuming[ `r:R ring`; `n:num`; `q:num->R`; `\i:num. ring_ord r (q i) (a:R)`; `b:R` ]primefact_divides THEN subgoal `!i. i IN 1..n ==> ring_ord r (q i:R) a <= ring_ord r (q i) b` THENL [ intro THEN have `ring_prime r (q(i:num):R)` [] THEN qed[] ; pass ] THEN have `ring_divides r (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a))) (b:R)` [] THEN qed[RING_ASSOCIATES_DIVIDES;RING_ASSOCIATES_REFL] ] );; let ring_ord_associates_eq = prove(` !(r:R ring) a b. UFD r ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> b IN ring_carrier r ==> ~(b = ring_0 r) ==> ( ring_associates r a b <=> (!p. ring_prime r p ==> ring_ord r p a = ring_ord r p b ) ) `, rw[ring_associates] THEN simp[ring_ord_divides_eq] THEN intro THEN splitiff THENL [ intro THEN have `ring_ord r (p:R) a <= ring_ord r p b` [] THEN have `ring_ord r (p:R) b <= ring_ord r p a` [] THEN ASM_ARITH_TAC ; intro THENL [ have `ring_ord r (p:R) a = ring_ord r p b` [] THEN ASM_ARITH_TAC ; have `ring_ord r (p:R) a = ring_ord r p b` [] THEN ASM_ARITH_TAC ] ] );; (* ===== lcm *) let lcm_exists = prove(` !(r:R ring) a b. UFD r ==> a IN ring_carrier r ==> b IN ring_carrier r ==> ?c. ( c IN ring_carrier r /\ ring_divides r a c /\ ring_divides r b c /\ (!d. d IN ring_carrier r ==> ring_divides r a d ==> ring_divides r b d ==> ring_divides r c d ) ) `, intro THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN case `a = ring_0(r:R ring)` THENL [ witness `ring_0(r:R ring)` THEN qed[RING_DIVIDES_0;RING_0;RING_DIVIDES_ZERO] ; pass ] THEN case `b = ring_0(r:R ring)` THENL [ witness `ring_0(r:R ring)` THEN qed[RING_DIVIDES_0;RING_0;RING_DIVIDES_ZERO] ; pass ] THEN have `ring_divides r (ring_gcd r (a,b)) (a:R)` [RING_GCD_DIVIDES] THEN have `~(ring_gcd r (a,b):R = ring_0 r)` [RING_DIVIDES_ZERO] THEN have `ring_gcd r (a,b):R IN ring_carrier r` [RING_GCD] THEN choose `u:R` `u:R IN ring_carrier r /\ a = ring_mul r (ring_gcd r (a,b)) u` [ring_divides;RING_GCD] THEN have `~(u:R = ring_0 r)` [RING_MUL_RZERO] THEN have `ring_divides r (ring_gcd r (a,b)) (b:R)` [RING_GCD_DIVIDES] THEN choose `v:R` `v:R IN ring_carrier r /\ b = ring_mul r (ring_gcd r (a,b)) v` [ring_divides;RING_GCD] THEN have `~(v:R = ring_0 r)` [RING_MUL_RZERO] THEN subgoal `ring_mul r u (b:R) = ring_mul r a v` THENL [ have `ring_mul r (ring_gcd r (a,b)) v = b:R` [] THEN have `ring_mul r (ring_gcd r (a,b)) u = a:R` [] THEN RING_TAC ; pass ] THEN witness `ring_mul r u b:R` THEN intro THENL [ qed[RING_MUL] ; rw[ring_divides] THEN intro THENL [ qed[] ; qed[RING_MUL] ; witness `v:R` THEN qed[] ] ; rw[ring_divides] THEN intro THENL [ qed[] ; qed[RING_MUL] ; witness `u:R` THEN qed[RING_MUL_SYM] ] ; have `ring_mul r u b:R IN ring_carrier r` [RING_MUL] THEN proven_if `d = ring_0 r:R` [RING_DIVIDES_0] THEN have `~(ring_mul r u b:R = ring_0 r)` [integral_domain] THEN specialize[ `r:R ring`; `ring_mul r u b:R`; `d:R` ]ring_ord_divides_eq THEN rw[know `ring_divides r (ring_mul r u b:R) d <=> (!p. ring_prime r p ==> ring_ord r p (ring_mul r u b) <= ring_ord r p d)`] THEN intro THEN have `ring_ord r (p:R) (ring_gcd r (a,b)) = MIN (ring_ord r p a) (ring_ord r p b)` [ring_ord_gcd] THEN case `ring_ord r p (a:R) <= ring_ord r p b` THENL [ have_rw `ring_ord r p (ring_mul r u b:R) = ring_ord r p u + ring_ord r p b` [ring_ord_mul] THEN have `ring_ord r p (ring_gcd r (a,b):R) = ring_ord r p a` [MIN] THEN subgoal `ring_ord r p a = ring_ord r p (ring_gcd r (a,b)) + ring_ord r p (u:R)` THENL [ specialize_assuming[ `r:R ring`; `p:R`; `ring_gcd r (a,b):R`; `u:R` ]ring_ord_mul THEN qed[] ; pass ] THEN num_linear_fact `ring_ord r (p:R) (ring_gcd r (a,b)) = ring_ord r p a ==> ring_ord r p a = ring_ord r p (ring_gcd r (a,b)) + ring_ord r p u ==> ring_ord r p u = 0` THEN rw[know `ring_ord r p (u:R) = 0`] THEN rw[ARITH_RULE `0+e = e`] THEN have `ring_ord r p (b:R) <= ring_ord r p d` [ring_ord_divides] THEN qed[] ; rw[know `ring_mul r u b = ring_mul r a v:R`] THEN have_rw `ring_ord r p (ring_mul r a v:R) = ring_ord r p a + ring_ord r p v` [ring_ord_mul] THEN have `ring_ord r p (ring_gcd r (a,b):R) = ring_ord r p b` [MIN] THEN subgoal `ring_ord r p b = ring_ord r p (ring_gcd r (a,b)) + ring_ord r p (v:R)` THENL [ specialize_assuming[ `r:R ring`; `p:R`; `ring_gcd r (a,b):R`; `v:R` ]ring_ord_mul THEN qed[] ; pass ] THEN num_linear_fact `ring_ord r (p:R) (ring_gcd r (a,b)) = ring_ord r p b ==> ring_ord r p b = ring_ord r p (ring_gcd r (a,b)) + ring_ord r p v ==> ring_ord r p v = 0` THEN rw[know `ring_ord r p (v:R) = 0`] THEN rw[ARITH_RULE `e+0 = e`] THEN have `ring_ord r p (a:R) <= ring_ord r p d` [ring_ord_divides] THEN qed[] ] ] );; let lcm_set_exists = prove(` !(r:R ring) a S. FINITE S ==> UFD r ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> ?c. ( c IN ring_carrier r /\ (!s. s IN S ==> ring_divides r (a s) c) /\ (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ) `, GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ witness `ring_1 r:R` THEN rw[EMPTY;IN_ELIM_THM;RING_1;RING_DIVIDES_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN choose `b:R` `b:R IN ring_carrier r /\ (!s:X. s IN S ==> ring_divides r (a s) b) /\ (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r b d)` [] THEN specialize_assuming[ `r:R ring`; `a(x:X):R`; `b:R` ]lcm_exists THEN choose `c:R` `c:R IN ring_carrier r /\ ring_divides r (a(x:X)) c /\ ring_divides r b c /\ (!d. d IN ring_carrier r ==> ring_divides r (a x) d ==> ring_divides r b d ==> ring_divides r c d)` [] THEN witness `c:R` THEN intro THENL [ qed[] ; proven_if `s = x:X` [] THEN set_fact `~(s = x:X) ==> s IN x INSERT S ==> s IN S` THEN qed[RING_DIVIDES_TRANS] ; qed[] ] ] );; let lcm_set_units = prove(` !(r:R ring) a S c. (!s:X. s IN S ==> a s IN ring_carrier r) ==> (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ==> (!s. s IN S ==> ring_unit r (a s)) ==> ring_unit r c `, intro THEN have `!s:X. s IN S ==> ring_divides r (a s) (ring_1 r:R)` [RING_UNIT_DIVIDES] THEN have `ring_divides r (c:R) (ring_1 r)` [RING_1] THEN qed[RING_UNIT_DIVIDES] );; let zero_lcm_set = prove(` !(r:R ring) a S c. integral_domain r ==> FINITE S ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> c IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) c) ==> (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ==> ( c = ring_0 r <=> ( ?s. s IN S /\ a s = ring_0 r ) ) `, intro THEN splitiff THENL [ intro THEN def `d:R` `ring_product r S (a:X->R)` THEN have `d:R IN ring_carrier r` [RING_PRODUCT] THEN have `!s:X. s IN S ==> ring_divides r (a s:R) d` [ring_divides_product] THEN have `ring_divides r c (d:R)` [] THEN have `d = ring_0 r:R` [RING_DIVIDES_ZERO] THEN qed[INTEGRAL_DOMAIN_PRODUCT_EQ_0] ; qed[RING_DIVIDES_ZERO] ] );; (* could extract from zero_lcm_set *) let zero_lcm = prove(` !(r:R ring) a b. integral_domain r ==> ring_divides r a c ==> ring_divides r b c ==> (!d. d IN ring_carrier r ==> ring_divides r a d ==> ring_divides r b d ==> ring_divides r c d ) ==> ( c = ring_0 r <=> ( a = ring_0 r \/ b = ring_0 r ) ) `, intro THEN splitiff THENL [ intro THEN def `d:R` `ring_mul r a b:R` THEN have `d:R IN ring_carrier r` [RING_MUL;ring_divides] THEN have `ring_divides r (a:R) d` [ring_divides] THEN have `ring_divides r (b:R) d` [ring_divides;RING_MUL_SYM] THEN have `ring_divides r c (d:R)` [] THEN have `d = ring_0 r:R` [RING_DIVIDES_ZERO] THEN qed[integral_domain;ring_divides] ; qed[RING_DIVIDES_ZERO] ] );; let ring_ord_lcm_set = prove(` !(r:R ring) a S c p t. UFD r ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> c IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) c) ==> (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ==> ~(c = ring_0 r) ==> ring_prime r p ==> t IN S ==> (!s:X. s IN S ==> ring_ord r p (a s) <= ring_ord r p (a t)) ==> ring_ord r p c = ring_ord r p (a t) `, intro THEN have `!s:X. s IN S ==> ~(a s = ring_0 r:R)` [RING_DIVIDES_ZERO] THEN have `!s:X. s IN S ==> ring_ord r p (a s:R) <= ring_ord r p c` [ring_ord_divides;UFD_IMP_INTEGRAL_DOMAIN] THEN have `ring_ord r p (a(t:X):R) <= ring_ord r p c` [] THEN case `ring_ord r p (a(t:X):R) < ring_ord r p c` THENL [ subgoal `ring_divides r p (c:R)` THENL [ num_linear_fact `ring_ord r p (a(t:X):R) < ring_ord r p c ==> 1 <= ring_ord r p c` THEN have `c:R IN ring_carrier r` [ring_divides] THEN specialize_assuming[ `r:R ring`; `p:R`; `c:R`; `1` ]divides_le_pow_ring_ord THEN qed[RING_POW_1;UFD_IMP_INTEGRAL_DOMAIN;ring_prime] ; pass ] THEN choose `u:R` `u IN ring_carrier r /\ c = ring_mul r p u:R` [ring_divides] THEN subgoal `!s:X. s IN S ==> ring_divides r (a s) (u:R)` THENL [ intro_gendisch THEN have `a(s:X):R IN ring_carrier r` [ring_divides] THEN have `~(a(s:X):R = ring_0 r)` [RING_DIVIDES_ZERO] THEN have `u:R IN ring_carrier r` [] THEN have `~(u:R = ring_0 r)` [RING_MUL_RZERO;ring_prime] THEN specialize[ `r:R ring`; `a(s:X):R`; `u:R` ]ring_ord_divides_eq THEN rw[know `ring_divides r (a(s:X)) u <=> (!q:R. ring_prime r q ==> ring_ord r q (a(s:X)) <= ring_ord r q u)`] THEN intro THEN case `ring_associates r p (q:R)` THENL [ have `~(p:R = ring_0 r)` [ring_prime] THEN have `ring_ord r p (c:R) = ring_ord r p p + ring_ord r p u` [ring_ord_mul;UFD_IMP_INTEGRAL_DOMAIN;ring_prime] THEN have `ring_ord r (p:R) p = 1` [ring_ord_refl;ring_prime;UFD_IMP_INTEGRAL_DOMAIN] THEN have `ring_ord r p (c:R) = 1 + ring_ord r p u` [] THEN have `ring_ord r p ((a(s:X)):R) < ring_ord r p c` [LET_TRANS] THEN simp[GSYM ring_associates_ord] THEN ASM_ARITH_TAC ; have `ring_ord r q (c:R) = ring_ord r q p + ring_ord r q u` [ring_ord_mul;UFD_IMP_INTEGRAL_DOMAIN;ring_prime] THEN specialize[ `r:R ring`; `q:R`; `p:R` ]ring_ord_prime THEN have `ring_ord r q (p:R) = 0` [RING_ASSOCIATES_SYM] THEN have `ring_ord r q (c:R) = 0 + ring_ord r q u` [] THEN have `ring_ord r q ((a(s:X)):R) <= ring_ord r q c` [ring_ord_divides;UFD_IMP_INTEGRAL_DOMAIN] THEN ASM_ARITH_TAC ] ; pass ] THEN have `ring_divides r (c:R) u` [] THEN choose `v:R` `v:R IN ring_carrier r /\ u = ring_mul r c v` [ring_divides] THEN subgoal `ring_mul r c (ring_1 r) = ring_mul r c (ring_mul r p v):R` THENL [ specialize[ ](GENL[ `c:R`;`p:R`;`u:R`;`v:R` ](RING_RULE `c = ring_mul r p u:R ==> u = ring_mul r c v ==> ring_mul r c (ring_1 r) = ring_mul r c (ring_mul r p v)` )) THEN qed[ring_prime;ring_divides] ; pass ] THEN subgoal `ring_1 r:R = ring_mul r p v` THENL [ specialize[ `r:R ring`; `c:R` ]INTEGRAL_DOMAIN_MUL_LCANCEL THEN qed[RING_MUL;RING_1;ring_divides;ring_prime;UFD_IMP_INTEGRAL_DOMAIN] ; pass ] THEN qed[ring_unit;ring_prime] ; pass ] THEN ASM_ARITH_TAC );; let ring_ord_lcm = prove(` !(r:R ring) a b c p. UFD r ==> ring_divides r a c ==> ring_divides r b c ==> (!d. d IN ring_carrier r ==> ring_divides r a d ==> ring_divides r b d ==> ring_divides r c d ) ==> ~(c = ring_0 r) ==> ring_prime r p ==> ring_ord r p c = MAX (ring_ord r p a) (ring_ord r p b) `, intro THEN have `!s. s IN 0..1 ==> (if s = 0 then a else b:R) IN ring_carrier r` [ring_divides] THEN have `c:R IN ring_carrier r` [ring_divides] THEN have `!s. s IN 0..1 ==> ring_divides r (if s = 0 then a else b:R) c` [] THEN have `0 IN 0..1` [IN_NUMSEG_0;ARITH_RULE `0 <= 1`] THEN have `1 IN 0..1` [IN_NUMSEG_0;ARITH_RULE `1 <= 1`] THEN subgoal `!d:R. d IN ring_carrier r ==> (!s. s IN 0..1 ==> ring_divides r (if s = 0 then a else b) d) ==> ring_divides r c d` THENL [ intro THEN have `ring_divides r (a:R) d` [] THEN have `ring_divides r (b:R) d` [ARITH_RULE `~(0 = 1)`] THEN qed[] ; pass ] THEN case `ring_ord r p (a:R) <= ring_ord r p b` THENL [ have `!s. s IN 0..1 ==> ring_ord r p (if s = 0 then a else b:R) <= ring_ord r p (if 1 = 0 then a else b)` [ARITH_RULE `~(1 = 0)`;LE_REFL] THEN specialize[ `r:R ring`; `\i:num. if i = 0 then a else b:R`; `0..1`; `c:R`; `p:R`; `1` ]ring_ord_lcm_set THEN simp[MAX] THEN qed[ARITH_RULE `~(1 = 0)`] ; num_linear_fact `~(ring_ord r p (a:R) <= ring_ord r p b) ==> ring_ord r p b <= ring_ord r p a` THEN have `!s. s IN 0..1 ==> ring_ord r p (if s = 0 then a else b:R) <= ring_ord r p a` [LE_REFL] THEN specialize[ `r:R ring`; `\i:num. if i = 0 then a else b:R`; `0..1`; `c:R`; `p:R`; `0` ]ring_ord_lcm_set THEN simp[MAX] ] );; let ring_ord_squarefree = prove(` !(r:R ring) a. UFD r ==> a IN ring_carrier r ==> ~(a = ring_0 r) ==> ( ring_squarefree r a <=> (!p. ring_prime r p ==> ring_ord r p a <= 1 ) ) `, intro THEN have `integral_domain(r:R ring)` [UFD_IMP_INTEGRAL_DOMAIN] THEN splitiff THENL [ intro THEN case `2 <= ring_ord r (p:R) a` THENL [ have `p:R IN ring_carrier r` [ring_prime] THEN have `~(ring_unit r (p:R))` [ring_prime] THEN have `noetherian_ring r \/ UFD(r:R ring)` [] THEN specialize[ `r:R ring`; `p:R`; `a:R`; `2` ]divides_le_pow_ring_ord THEN have `ring_divides r (ring_mul r p p:R) a` [RING_POW_2] THEN specialize[ `r:R ring`; `a:R`; `p:R` ]not_squarefree_if_divisible_by_square THEN qed[] ; pass ] THEN ASM_ARITH_TAC ; intro THEN specialize[ `r:R ring`; `a:R` ]primefact_ord THEN choose2 `n:num` `q:num->R` `(!i. i IN (1..n) ==> ring_prime r (q i:R)) /\ (!i. i IN 1..n ==> ring_divides r (q i) a) /\ (!i j. i IN (1..n) ==> j IN (1..n) ==> ring_associates r (q i) (q j) ==> i = j) /\ ring_associates r a (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)))` [] THEN subgoal `ring_squarefree r (ring_product r (1..n) q:R)` THENL [ have `FINITE(1..n)` [FINITE_NUMSEG] THEN have `!i. i IN 1..n ==> ring_prime r (q i:R)` [] THEN have `!i. i IN 1..n ==> q i:R IN ring_carrier r` [ring_prime] THEN have `!i j. i IN 1..n ==> j IN 1..n ==> ring_divides r (q i) (q j:R) ==> i = j` [ring_divides_associates_prime] THEN specialize[ `r:R ring`; `1..n`; `q:num->R` ]ring_squarefree_if_product_coprime_primes_indexed THEN qed[] ; pass ] THEN subgoal `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a):R) = ring_product r (1..n) q` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `ring_ord r (q(a':num)) (a:R) = 1` THENL [ have `1 <= ring_ord r (q(a':num)) (a:R)` [ring_ord_prime_divides] THEN have `ring_ord r (q(a':num)) (a:R) <= 1` [] THEN num_linear_fact `1 <= ring_ord r (q(a':num)) (a:R) ==> ring_ord r (q a') a <= 1 ==> ring_ord r (q a') a = 1` THEN qed[] ; pass ] THEN qed[RING_POW_1;ring_prime] ; pass ] THEN have `ring_squarefree r (ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)):R)` [] THEN specialize_assuming[ `r:R ring`; `ring_product r (1..n) (\i. ring_pow r (q i) (ring_ord r (q i) a)):R`; `a:R` ]ring_squarefree_associates THEN qed[RING_ASSOCIATES_SYM] ] );; let ring_squarefree_lcm = prove(` !(r:R ring) a b c. UFD r ==> ring_divides r a c ==> ring_divides r b c ==> (!d. d IN ring_carrier r ==> ring_divides r a d ==> ring_divides r b d ==> ring_divides r c d ) ==> ~(c = ring_0 r) ==> ring_squarefree r a ==> ring_squarefree r b ==> ring_squarefree r c `, intro THEN have `~(a = ring_0 r:R)` [RING_DIVIDES_ZERO] THEN have `~(b = ring_0 r:R)` [RING_DIVIDES_ZERO] THEN have `c:R IN ring_carrier r` [ring_divides] THEN simp[ring_ord_squarefree] THEN intro THEN have `a:R IN ring_carrier r` [ring_divides] THEN have `ring_ord r p (a:R) <= 1` [ring_ord_squarefree] THEN have `b:R IN ring_carrier r` [ring_divides] THEN have `ring_ord r p (b:R) <= 1` [ring_ord_squarefree] THEN simp[ring_ord_lcm] THEN simp[max_le] );; let ring_squarefree_lcm_set = prove(` !(r:R ring) a S c. UFD r ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> c IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) c) ==> (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ==> ~(c = ring_0 r) ==> (!s. s IN S ==> ring_squarefree r (a s)) ==> ring_squarefree r c `, intro THEN have `!s:X. s IN S ==> ~(a s = ring_0 r:R)` [RING_DIVIDES_ZERO] THEN have `c:R IN ring_carrier r` [ring_divides] THEN simp[ring_ord_squarefree] THEN intro THEN have `!s:X. s IN S ==> (a s):R IN ring_carrier r` [ring_divides] THEN have `!s:X. s IN S ==> ring_ord r p ((a s):R) <= 1` [ring_ord_squarefree] THEN case `?t:X. t IN S /\ ring_ord r p (a t:R) = 1` THENL [ choose `t:X` `t:X IN S /\ ring_ord r p (a t:R) = 1` [] THEN have `!s:X. s IN S ==> ring_ord r p (a s:R) <= ring_ord r p (a t)` [] THEN specialize_assuming[ `r:R ring`; `a:X->R`; `S:X->bool`; `c:R`; `p:R`; `t:X` ]ring_ord_lcm_set THEN have `ring_ord r p c = ring_ord r p (a(t:X):R)` [] THEN qed[] ; pass ] THEN subgoal `!t:X. t IN S ==> ring_ord r p (a t:R) = 0` THENL [ intro THEN have `ring_ord r p (a(t:X):R) <= 1` [] THEN have `~(ring_ord r p (a(t:X):R) = 1)` [] THEN ASM_ARITH_TAC ; pass ] THEN case `S = {}:X->bool` THENL [ subgoal `ring_divides r c (ring_1 r:R)` THENL [ subgoal `!s:X. s IN {} ==> ring_divides r (a s:R) (ring_1 r)` THENL [ rw[EMPTY;IN_ELIM_THM] ; pass ] THEN qed[RING_1] ; pass ] THEN have `ring_unit r (c:R)` [ring_unit;ring_divides] THEN have `ring_ord r p (c:R) = 0` [ring_ord_unit;ring_prime] THEN qed[ARITH_RULE `0 <= 1`] ; pass ] THEN set_fact `~(S = {}) ==> ?t:X. t IN S` THEN choose `t:X` `t:X IN S` [] THEN have `!s:X. s IN S ==> ring_ord r p (a s:R) <= ring_ord r p (a t)` [ARITH_RULE `0 <= 0`] THEN specialize_assuming[ `r:R ring`; `a:X->R`; `S:X->bool`; `c:R`; `p:R`; `t:X` ]ring_ord_lcm_set THEN have `ring_ord r p c = ring_ord r p (a(t:X):R)` [] THEN qed[ARITH_RULE `0 <= 1`] );; let ring_prime_divides_lcm_set = prove(` !(r:R ring) a S c p. UFD r ==> (!s:X. s IN S ==> a s IN ring_carrier r) ==> c IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) c) ==> (!d. d IN ring_carrier r ==> (!s. s IN S ==> ring_divides r (a s) d) ==> ring_divides r c d ) ==> ~(c = ring_0 r) ==> ring_prime r p ==> ( ring_divides r p c <=> (?s. s IN S /\ ring_divides r p (a s)) ) `, intro THEN splitiff THENL [ intro THEN case `S = {}:X->bool` THENL [ subgoal `!s:X. s IN S ==> ring_unit r (a s:R)` THENL [ simp[IN_ELIM_THM;EMPTY] ; pass ] THEN specialize[ `r:R ring`; `a:X->R`; `S:X->bool`; `c:R` ]lcm_set_units THEN qed[RING_DIVIDES_UNIT;ring_prime] ; pass ] THEN case `!s:X. s IN S ==> ~(ring_divides r (p:R) (a s))` THENL [ choose `t:X` `t:X IN S` [MEMBER_NOT_EMPTY] THEN have `!s:X. s IN S ==> ring_ord r (p:R) (a s) = 0` [ring_ord_notdivides;ring_prime] THEN have `!s:X. s IN S ==> ring_ord r (p:R) (a s) <= ring_ord r p (a t)` [ARITH_RULE `0 <= 0`] THEN specialize[ `r:R ring`; `a:X->R`; `S:X->bool`; `c:R`; `p:R`; `t:X` ]ring_ord_lcm_set THEN have `ring_ord r (p:R) c <= 0` [] THEN num_linear_fact `ring_ord r (p:R) c <= 0 ==> ring_ord r p c = 0` THEN qed[nonzero_ring_ord_if_divides;UFD_IMP_INTEGRAL_DOMAIN] ; pass ] THEN qed[] ; qed[RING_DIVIDES_TRANS] ] );; (* ===== polynomial lcm *) let poly_lcm_set_exists = prove(` !(r:R ring) a S. field r ==> FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (a s)) ==> ?c. ( ring_polynomial r c /\ (!s. s IN S ==> ring_divides(x_poly r) (a s) c) /\ (!d. ring_polynomial r d ==> (!s. s IN S ==> ring_divides(x_poly r) (a s) d) ==> ring_divides(x_poly r) c d ) ) `, intro THEN have `UFD(x_poly(r:R ring))` [UFD_x_poly_field] THEN have `!s:X. s IN S ==> a s:(1->num)->R IN ring_carrier(x_poly r)` [x_poly_use] THEN specialize[ `x_poly(r:R ring)`; `a:X->(1->num)->R`; `S:X->bool` ]lcm_set_exists THEN choose `c:(1->num)->R` `c:(1->num)->R IN ring_carrier(x_poly r) /\ (!s:X. s IN S ==> ring_divides(x_poly r) (a s) c) /\ (!d. d IN ring_carrier(x_poly r) ==> (!s. s IN S ==> ring_divides(x_poly r) (a s) d) ==> ring_divides(x_poly r) c d)` [] THEN witness `c:(1->num)->R` THEN rw[x_poly_use] THEN qed[lcm_set_exists] );; let monic_lcm_set_exists = prove(` !(r:R ring) a S. field r ==> FINITE S ==> (!s:X. s IN S ==> ring_polynomial r (a s)) ==> (!s:X. s IN S ==> ~(a s = poly_0 r)) ==> ?c. ( ring_polynomial r c /\ monic r c /\ (!s. s IN S ==> ring_divides(x_poly r) (a s) c) /\ (!d. ring_polynomial r d ==> (!s. s IN S ==> ring_divides(x_poly r) (a s) d) ==> ring_divides(x_poly r) c d ) ) `, intro THEN specialize[ `r:R ring`; `a:X->(1->num)->R`; `S:X->bool` ]poly_lcm_set_exists THEN choose `c:(1->num)->R` `ring_polynomial (r:R ring) c /\ (!s:X. s IN S ==> ring_divides(x_poly r) (a s) c) /\ (!d. ring_polynomial r d ==> (!s. s IN S ==> ring_divides(x_poly r) (a s) d) ==> ring_divides(x_poly r) c d)` [] THEN subgoal `~(c = poly_0 r:(1->num)->R)` THENL [ specialize_assuming[ `x_poly(r:R ring)`; `a:X->(1->num)->R`; `S:X->bool`; `c:(1->num)->R` ]zero_lcm_set THEN qed [x_poly_use;integral_domain_x_poly_field] ; pass ] THEN specialize_assuming[ `r:R ring`; `c:(1->num)->R` ]x_poly_field_monic_associate THEN choose `q:(1->num)->R` `ring_polynomial r (q:(1->num)->R) /\ monic r q /\ ring_associates(x_poly r) c q` [] THEN witness `q:(1->num)->R` THEN qed[RING_ASSOCIATES_DIVIDES;RING_ASSOCIATES_REFL;x_poly_use] );; (* ===== more on algebraic numbers *) let minimal_squarefree_from_algebraic_set = prove(` !S. FINITE S ==> S SUBSET algebraic_number ==> (?p. ring_polynomial QinC_ring p /\ ring_squarefree(x_poly QinC_ring) p /\ monic QinC_ring p /\ S SUBSET complex_root p /\ (!z. complex_root p z ==> (?s. s IN S /\ minimal_polynomial s = minimal_polynomial z)) ) `, intro THEN recall field_QinC THEN have `!z:complex. z IN S ==> ring_polynomial QinC_ring (minimal_polynomial z)` [SUBSET;IN;algebraic_has_minimal_polynomial] THEN have `!z:complex. z IN S ==> ~(minimal_polynomial z = poly_0 QinC_ring)` [SUBSET;IN;algebraic_has_minimal_polynomial] THEN specialize[ `QinC_ring`; `minimal_polynomial`; `S:complex->bool` ]monic_lcm_set_exists THEN choose `p:(1->num)->complex` `ring_polynomial QinC_ring p /\ monic QinC_ring p /\ (!s:complex. s IN S ==> ring_divides(x_poly QinC_ring) (minimal_polynomial s) p) /\ (!d. ring_polynomial QinC_ring d ==> (!s. s IN S ==> ring_divides(x_poly QinC_ring) (minimal_polynomial s) d) ==> ring_divides(x_poly QinC_ring) p d)` [] THEN witness `p:(1->num)->complex` THEN intro THENL [ qed[] ; recall UFD_x_poly_QinC THEN have `!z:complex. z IN S ==> minimal_polynomial z IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `p IN ring_carrier(x_poly QinC_ring)` [x_poly_use] THEN have `!z:complex. z IN S ==> ring_divides(x_poly QinC_ring) (minimal_polynomial z) p` [] THEN have `~(p = ring_0(x_poly QinC_ring))` [x_poly_use;monic_poly_0;ring_1_0_QinC;in_complex_ring] THEN subgoal `!z:complex. z IN S ==> ring_squarefree(x_poly QinC_ring) (minimal_polynomial z)` THENL [ intro THEN have `algebraic_number z` [SUBSET;IN] THEN have `ring_irreducible(x_poly QinC_ring) (minimal_polynomial z)` [SUBSET;IN;algebraic_has_minimal_polynomial] THEN qed[squarefree_if_irreducible_over_field] ; pass ] THEN have `!z:complex. z IN S ==> ring_divides(x_poly QinC_ring) (minimal_polynomial z) p` [] THEN have `!d. d IN ring_carrier (x_poly QinC_ring) ==> (!s:complex. s IN S ==> ring_divides (x_poly QinC_ring) (minimal_polynomial s) d) ==> ring_divides (x_poly QinC_ring) p d` [x_poly_use] THEN specialize[ `x_poly QinC_ring`; `minimal_polynomial`; `S:complex->bool`; `p:(1->num)->complex` ]ring_squarefree_lcm_set THEN qed[] ; qed[] ; rw[SUBSET] THEN intro THEN rw[IN] THEN have `algebraic_number x` [SUBSET;IN] THEN have `poly_eval complex_ring (minimal_polynomial x) x = Cx(&0)` [algebraic_has_minimal_polynomial] THEN have `ring_divides(x_poly QinC_ring) (minimal_polynomial x) p` [] THEN choose `u:(1->num)->complex` `u IN ring_carrier(x_poly QinC_ring) /\ p = ring_mul(x_poly QinC_ring) (minimal_polynomial x) u` [ring_divides] THEN subgoal `p = poly_mul complex_ring (minimal_polynomial x) u:(1->num)->complex` THENL [ recall subring_complex_QinC THEN have `ring_powerseries QinC_ring (minimal_polynomial x:(1->num)->complex)` [ring_polynomial] THEN have `ring_powerseries QinC_ring (u:(1->num)->complex)` [ring_polynomial;x_poly_use] THEN specialize_assuming[ `complex_ring`; `QinC`; `minimal_polynomial x`; `u:(1->num)->complex` ]poly_mul_subring THEN qed[x_poly_use] ; pass ] THEN subgoal `poly_eval complex_ring (poly_mul complex_ring (minimal_polynomial x) u) x = (poly_eval complex_ring (minimal_polynomial x) x) * (poly_eval complex_ring u x)` THENL [ rw[GSYM complex_ring_clauses] THEN specialize_assuming[ `complex_ring`; `minimal_polynomial x`; `u:(1->num)->complex`; `x:complex` ]POLY_EVAL_MUL THEN qed[poly_complex_if_poly_QinC;x_poly_use;in_complex_ring] ; pass ] THEN have `poly_eval complex_ring (poly_mul complex_ring (minimal_polynomial x) u) x = Cx(&0) * (poly_eval complex_ring u x)` [] THEN complex_field_fact `Cx(&0) * (poly_eval complex_ring u x) = Cx(&0)` THEN have `poly_eval complex_ring (poly_mul complex_ring (minimal_polynomial x) u) x = Cx(&0)` [] THEN have `poly_eval complex_ring p x = Cx(&0)` [] THEN qed[complex_root] ; intro THEN recall UFD_x_poly_QinC THEN have `!s. s IN S ==> minimal_polynomial s IN ring_carrier (x_poly QinC_ring)` [x_poly_use] THEN have `p IN ring_carrier (x_poly QinC_ring)` [x_poly_use] THEN have `!s. s IN S ==> ring_divides (x_poly QinC_ring) (minimal_polynomial s) p` [] THEN subgoal `!d. d IN ring_carrier (x_poly QinC_ring) ==> (!s. s IN S ==> ring_divides (x_poly QinC_ring) (minimal_polynomial s) d) ==> ring_divides (x_poly QinC_ring) p d` THENL [ rw[GSYM x_poly_use] THEN qed[] ; pass ] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN subgoal `~(p = ring_0 (x_poly QinC_ring))` THENL [ rw[GSYM x_poly_use] THEN qed[] ; pass ] THEN have `algebraic_number z` [algebraic_number_root_QinC_poly;complex_root] THEN have `ring_irreducible (x_poly QinC_ring) (minimal_polynomial z)` [algebraic_has_minimal_polynomial] THEN have `ring_prime (x_poly QinC_ring) (minimal_polynomial z)` [prime_iff_irreducible_over_field;field_QinC] THEN specialize[ `x_poly QinC_ring`; `minimal_polynomial`; `S:complex->bool`; `p:(1->num)->complex`; `minimal_polynomial z` ]ring_prime_divides_lcm_set THEN have `ring_divides(x_poly QinC_ring) (minimal_polynomial z) p` [minimal_polynomial_divides] THEN choose `s:complex` `s IN S /\ ring_divides (x_poly QinC_ring) (minimal_polynomial z) (minimal_polynomial s)` [] THEN witness `s:complex` THEN have `algebraic_number s` [SUBSET;IN] THEN have `ring_irreducible (x_poly QinC_ring) (minimal_polynomial s)` [algebraic_has_minimal_polynomial] THEN have `ring_prime (x_poly QinC_ring) (minimal_polynomial s)` [prime_iff_irreducible_over_field;field_QinC] THEN have `ring_associates (x_poly QinC_ring) (minimal_polynomial z) (minimal_polynomial s)` [ring_divides_associates_prime;integral_domain_x_poly_QinC] THEN have `monic QinC_ring (minimal_polynomial s)` [algebraic_has_minimal_polynomial] THEN have `monic QinC_ring (minimal_polynomial z)` [algebraic_has_minimal_polynomial] THEN recall field_QinC THEN specialize[ `QinC_ring`; `minimal_polynomial z`; `minimal_polynomial s` ]monic_associates THEN qed[] ] );; let squarefree_from_algebraic_set = prove(` !S. FINITE S ==> S SUBSET algebraic_number ==> (?p. ring_polynomial QinC_ring p /\ ring_squarefree(x_poly QinC_ring) p /\ monic QinC_ring p /\ S SUBSET complex_root p ) `, qed[minimal_squarefree_from_algebraic_set] );; let squarefree_from_algebraic_set_avoiding_0 = prove(` !S. FINITE S ==> S SUBSET algebraic_number ==> ~(Cx(&0) IN S) ==> (?p. ring_polynomial QinC_ring p /\ ring_squarefree(x_poly QinC_ring) p /\ monic QinC_ring p /\ S SUBSET complex_root p /\ ~(complex_root p (Cx(&0))) ) `, intro THEN specialize[`S:complex->bool`]minimal_squarefree_from_algebraic_set THEN choose `p:(1->num)->complex` `ring_polynomial QinC_ring p /\ ring_squarefree (x_poly QinC_ring) p /\ monic QinC_ring p /\ S SUBSET complex_root p /\ (!z. complex_root p z ==> (?s. s IN S /\ minimal_polynomial s = minimal_polynomial z))` [] THEN witness `p:(1->num)->complex` THEN simp[] THEN case `complex_root p (Cx(&0))` THENL [ choose `s:complex` `s IN S /\ minimal_polynomial s = minimal_polynomial(Cx(&0))` [] THEN have `minimal_polynomial s = x_minus_const QinC_ring (Cx(&0))` [minimal_polynomial_QinC;QinC_0] THEN have `minimal_polynomial s = x_minus_const complex_ring (Cx(&0))` [x_minus_const_QinC_eq_x_minus_const_complex] THEN have `poly_eval complex_ring (x_minus_const complex_ring (Cx(&0))) s = Cx(&0)` [maybe_algebraic_minimal_polynomial] THEN have `poly_eval complex_ring (x_minus_const complex_ring (Cx(&0))) s = ring_sub complex_ring s (Cx(&0))` [eval_x_minus_const;in_complex_ring] THEN have `ring_sub complex_ring s (Cx(&0)) = Cx(&0)` [] THEN have `s - (Cx(&0)) = Cx(&0)` [ring_sub_complex] THEN complex_field_fact `s - (Cx(&0)) = Cx(&0) ==> s = Cx(&0)` THEN qed[] ; pass ] THEN qed[] );; (* ===== more on multivariate polynomials *) let poly_vars_sum_poly_subset = prove(` !(r:R ring) p:X->(V->num)->R U S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> (!s. s IN S ==> poly_vars r (p s) SUBSET U) ==> poly_vars r (ring_sum(poly_ring r (:V)) S p) SUBSET U `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;POLY_RING;POLY_VARS_0;EMPTY_SUBSET] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `p(x:X):(V->num)->R IN ring_carrier(poly_ring r (:V))` [poly_in_full_ring] THEN simp[RING_SUM_CLAUSES] THEN have `ring_powerseries r (p(x:X):(V->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (ring_sum(poly_ring r (:V)) S (p:X->(V->num)->R))` [RING_SUM;ring_polynomial;poly_in_full_ring] THEN rw[POLY_RING] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_sum(poly_ring r (:V)) S (p:X->(V->num)->R)` ]POLY_VARS_ADD THEN ASM SET_TAC[] ] );; let poly_vars_product_poly_subset = prove(` !(r:R ring) p:X->(V->num)->R U S. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> (!s. s IN S ==> poly_vars r (p s) SUBSET U) ==> poly_vars r (ring_product(poly_ring r (:V)) S p) SUBSET U `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POLY_RING;POLY_VARS_1;EMPTY_SUBSET] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN have `p(x:X):(V->num)->R IN ring_carrier(poly_ring r (:V))` [poly_in_full_ring] THEN simp[RING_PRODUCT_CLAUSES] THEN have `ring_powerseries r (p(x:X):(V->num)->R)` [ring_polynomial] THEN have `ring_powerseries r (ring_product(poly_ring r (:V)) S (p:X->(V->num)->R))` [RING_PRODUCT;ring_polynomial;poly_in_full_ring] THEN rw[POLY_RING] THEN specialize[ `r:R ring`; `p(x:X):(V->num)->R`; `ring_product(poly_ring r (:V)) S (p:X->(V->num)->R)` ]POLY_VARS_MUL THEN ASM SET_TAC[] ] );; let poly_vars_pow = prove(` !(r:R ring) p:(V->num)->R n. ring_powerseries r p ==> poly_vars r (poly_pow r p n) SUBSET poly_vars r p `, GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;POLY_VARS_1;EMPTY_SUBSET] ; intro THEN rw[ARITH_RULE `SUC n = 1+n`] THEN simp[poly_pow_add;poly_pow_1] THEN have `ring_powerseries r (poly_pow r p n:(V->num)->R)` [poly_pow_series] THEN specialize[ `r:R ring`; `p:(V->num)->R`; `poly_pow r p n:(V->num)->R` ]POLY_VARS_MUL THEN ASM SET_TAC[] ] );; let poly_vars_pow_subset = prove(` !(r:R ring) p:(V->num)->R n U. ring_powerseries r p ==> poly_vars r p SUBSET U ==> poly_vars r (poly_pow r p n) SUBSET U `, qed[poly_vars_pow;SUBSET_TRANS] );; let poly_evaluate_pow = prove(` !(r:R ring) p:(V->num)->R q:V->R n. ring_polynomial r p ==> (!i. i IN poly_vars r p ==> q i IN ring_carrier r) ==> poly_evaluate r (poly_pow r p n) q = ring_pow r (poly_evaluate r p q) n `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN INDUCT_TAC THENL [ rw[poly_pow_0;ring_pow;POLY_RING_CLAUSES;POLY_EVALUATE_1] ; rw[ring_pow] THEN rw[ARITH_RULE `SUC n = 1+n`] THEN intro THEN have `ring_powerseries r (p:(V->num)->R)` [ring_polynomial] THEN simp[poly_pow_add;poly_pow_1] THEN have `ring_polynomial r (poly_pow r p n:(V->num)->R)` [poly_pow_poly] THEN subgoal `!i. i IN poly_vars r p UNION poly_vars r (poly_pow r p n:(V->num)->R) ==> q i IN ring_carrier r` THENL [ rw[IN_UNION] THEN qed[poly_vars_pow;SUBSET] ; pass ] THEN specialize[ `r:R ring`; `p:(V->num)->R`; `poly_pow r p n:(V->num)->R`; `q:V->R` ]POLY_EVALUATE_MUL THEN qed[] ] );; let poly_evaluate_sum = prove(` !(r:R ring) p:X->(V->num)->R q:V->R S:X->bool. FINITE S ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> poly_evaluate r (ring_sum(poly_ring r (:V)) S p) q = ring_sum r S (\s. poly_evaluate r (p s) q) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_SUM_CLAUSES;POLY_RING_CLAUSES;POLY_EVALUATE_0] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN subgoal `p(x:X):(V->num)->R IN ring_carrier(poly_ring r (:V))` THENL [ rw[POLY_RING_CLAUSES;IN_ELIM_THM;SUBSET_UNIV] THEN qed[] ; pass ] THEN simp[RING_SUM_CLAUSES;POLY_EVALUATE] THEN rw[POLY_RING_CLAUSES] THEN qed[poly_in_full_ring;RING_SUM;POLY_EVALUATE_ADD] ] );; (* XXX: limit q v condition to variables appearing *) let poly_evaluate_product = prove(` !(r:R ring) p:X->(V->num)->R q:V->R S:X->bool. FINITE S ==> (!v. q v IN ring_carrier r) ==> (!s. s IN S ==> ring_polynomial r (p s)) ==> poly_evaluate r (ring_product(poly_ring r (:V)) S p) q = ring_product r S (\s. poly_evaluate r (p s) q) `, GEN_TAC THEN GEN_TAC THEN GEN_TAC THEN sufficesby FINITE_INDUCT_STRONG THEN intro THENL [ rw[RING_PRODUCT_CLAUSES;POLY_RING_CLAUSES;POLY_EVALUATE_1] ; set_fact `(x:X) IN x INSERT S` THEN set_fact `!s:X. s IN S ==> s IN x INSERT S` THEN subgoal `p(x:X):(V->num)->R IN ring_carrier(poly_ring r (:V))` THENL [ rw[POLY_RING_CLAUSES;IN_ELIM_THM;SUBSET_UNIV] THEN qed[] ; pass ] THEN simp[RING_PRODUCT_CLAUSES;POLY_EVALUATE] THEN rw[POLY_RING_CLAUSES] THEN qed[poly_in_full_ring;RING_PRODUCT;POLY_EVALUATE_MUL] ] );; (* ===== vandermonde *) (* can prove this via det *) (* or via transpose and monic_vanishing_at *) let vandermonde_indep_range = prove(` !(r:R ring) n c v. integral_domain r ==> (!i. i < n ==> c i IN ring_carrier r) ==> (!i. i < n ==> v i IN ring_carrier r) ==> (!i j. i < n ==> j < n ==> v i = v j ==> i = j) ==> (!e. e < n ==> ring_sum r (range n) (\i. ring_mul r (c i) (ring_pow r (v i) e)) = ring_0 r) ==> (!i. i < n ==> c i = ring_0 r) `, GEN_TAC THEN INDUCT_TAC THENL [ qed[ARITH_RULE `~(i < 0)`] ; rw[ARITH_RULE `SUC n = n+1`] THEN intro_gendisch THEN def `d:num->R` `\i:num. ring_mul r (c i) (ring_sub r (v i) (v n)):R` THEN num_linear_fact `!i:num. i < n ==> i < n+1` THEN num_linear_fact `n < n+1` THEN subgoal `!i:num. i < n ==> d i:R IN ring_carrier r` THENL [ intro THEN qed[RING_MUL;RING_SUB] ; pass ] THEN have `!i:num. i < n ==> v i:R IN ring_carrier r` [] THEN have `!i j:num. i < n ==> j < n ==> v i = v j:R ==> i = j` [] THEN subgoal `!e:num. e < n ==> ring_sum r (range(n+1)) (\i. ring_mul r (d i) (ring_pow r (v i) e)) = ring_0 r:R` THENL [ intro THEN rw[know `d = (\i:num. ring_mul r (c i) (ring_sub r (v i) (v n)):R)`] THEN num_linear_fact `e < n ==> e < n+1` THEN num_linear_fact `e < n ==> e+1 < n+1` THEN subgoal `ring_sum r (range(n+1)) (\i. ring_mul r (ring_mul r (c i) (ring_sub r (v i) (v n))) (ring_pow r (v i) e)) = ring_sum r (range(n+1)) (\i. ring_sub r (ring_mul r (c i) (ring_pow r (v i) (e+1))) (ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e))):R)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `a < n+1` [range_lt] THEN have `v(a:num):R IN ring_carrier r` [] THEN specialize[ `r:R ring`; `v(a:num):R`; `e:num`; `1` ]RING_POW_ADD THEN rw[know `ring_pow r (v(a:num)) (e + 1) = ring_mul r (ring_pow r (v a) e) (ring_pow r (v a) 1):R`] THEN have `c(a:num):R IN ring_carrier r` [] THEN have `v(n:num):R IN ring_carrier r` [] THEN RING_TAC ; pass ] THEN rw[know `ring_sum r (range(n+1)) (\i. ring_mul r (ring_mul r (c i) (ring_sub r (v i) (v n))) (ring_pow r (v i) e)) = ring_sum r (range(n+1)) (\i. ring_sub r (ring_mul r (c i) (ring_pow r (v i) (e + 1))) (ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e)))):R`] THEN subgoal `ring_sum r (range(n+1)) (\s. ring_sub r (ring_mul r (c s) (ring_pow r (v s) (e + 1))) (ring_mul r (v n) (ring_mul r (c s) (ring_pow r (v s) e)))) = ring_sub r (ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) (e + 1)))) (ring_sum r (range(n+1)) (\i. ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e)))):R` THENL [ specialize_assuming[ `r:R ring`; `\i:num. ring_mul r (c i) (ring_pow r (v i) (e + 1)):R`; `\i:num. ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e)):R`; `range(n+1)` ]ring_sum_sub THEN qed[RING_MUL;RING_POW;range_lt;finite_range] ; pass ] THEN rw[know `ring_sum r (range(n+1)) (\s. ring_sub r (ring_mul r (c s) (ring_pow r (v s) (e + 1))) (ring_mul r (v n) (ring_mul r (c s) (ring_pow r (v s) e)))) = ring_sub r (ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) (e + 1)))) (ring_sum r (range(n+1)) (\i. ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e)))):R`] THEN subgoal `ring_sum r (range(n+1)) (\i. ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e))) = ring_mul r (v n) (ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) e))):R` THENL [ specialize_assuming[ `r:R ring`; `\i:num. ring_mul r (c i) (ring_pow r (v i) e):R`; `v(n:num):R`; `range(n+1)` ]RING_SUM_LMUL THEN qed[finite_range;range_lt;RING_MUL;RING_POW] ; pass ] THEN rw[know `ring_sum r (range(n+1)) (\i. ring_mul r (v n) (ring_mul r (c i) (ring_pow r (v i) e))) = ring_mul r (v n) (ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) e))):R`] THEN have `!i j:num. i < n ==> j < n ==> v i = v j:R ==> i = j` [ARITH_RULE `i < n ==> i < n+1`] THEN have_rw `ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) (e + 1))) = ring_0 r:R` [] THEN have_rw `ring_sum r (range(n+1)) (\i. ring_mul r (c i) (ring_pow r (v i) e)) = ring_0 r:R` [] THEN have `v(n:num):R IN ring_carrier r` [] THEN qed[RING_MUL_RZERO;RING_SUB_RZERO;RING_0] ; pass ] THEN subgoal `!e:num. e < n ==> ring_sum r (range n) (\i. ring_mul r (d i) (ring_pow r (v i) e)) = ring_0 r:R` THENL [ intro THEN num_linear_fact `n < n+1` THEN subgoal `d(n:num):R IN ring_carrier r` THENL [ rw[know `d = (\i:num. ring_mul r (c i) (ring_sub r (v i) (v n)):R)`] THEN qed[RING_MUL;RING_SUB] ; pass ] THEN have `ring_mul r (d(n:num)) (ring_pow r (v n) e):R IN ring_carrier r` [RING_MUL;RING_POW] THEN specialize[ `r:R ring`; `n:num`; `\i:num. ring_mul r (d i) (ring_pow r (v i) e):R` ]ring_sum_range_add_1_sub THEN rw[know `ring_sum r (range n) (\i. ring_mul r (d i) (ring_pow r (v i) e)) = ring_sub r (ring_sum r (range (n + 1)) (\i. ring_mul r (d i) (ring_pow r (v i) e))) (ring_mul r (d n) (ring_pow r (v n) e)):R`] THEN have_rw `ring_sum r (range (n + 1)) (\i. ring_mul r (d i) (ring_pow r (v i) e)) = ring_0 r:R` [] THEN subgoal `d(n:num):R = ring_0 r` THENL [ rw[know `d = (\i:num. ring_mul r (c i) (ring_sub r (v i) (v n)):R)`] THEN have `ring_sub r (v(n:num)) (v n):R = ring_0 r` [RING_SUB_REFL] THEN qed[RING_MUL_RZERO] ; pass ] THEN rw[know `d(n:num):R = ring_0 r`] THEN qed[RING_POW;RING_MUL_LZERO;RING_SUB_RZERO] ; pass ] THEN specialize[ `d:num->R`; `v:num->R` ](know `!c v. integral_domain r ==> (!i. i < n ==> c i IN ring_carrier r) ==> (!i. i < n ==> v i IN ring_carrier r) ==> (!i j. i < n ==> j < n ==> v i = v j ==> i = j) ==> (!e. e < n ==> ring_sum r (range n) (\i. ring_mul r (c i) (ring_pow r (v i) e)) = ring_0 r) ==> (!i. i < n ==> c i = ring_0 r:R)`) THEN subgoal `!i:num. i < n ==> c i = ring_0 r:R` THENL [ intro THEN num_linear_fact `i < n:num ==> ~(i = n)` THEN num_linear_fact `i < n:num ==> i < n+1` THEN have `v(i:num):R IN ring_carrier r` [] THEN have `v(n:num):R IN ring_carrier r` [] THEN have `ring_mul r (c(i:num)) (ring_sub r (v i) (v n)) = ring_0 r:R` [] THEN subgoal `~(ring_sub r (v(i:num)) (v n) = ring_0 r:R)` THENL [ have `~(v(i:num) = v(n:num):R)` [] THEN qed[RING_SUB_EQ_0] ; pass ] THEN qed[integral_domain;RING_SUB] ; pass ] THEN subgoal `ring_sum r (range(n+1)) c = ring_0 r:R` THENL [ num_linear_fact `0 < n+1` THEN have `ring_sum r (range (n + 1)) (\i. ring_mul r (c i) (ring_pow r (v i) 0)) = ring_0 r:R` [] THEN subgoal `ring_sum r (range (n + 1)) (\i. ring_mul r (c i) (ring_pow r (v i) 0)) = ring_sum r (range(n+1)) c:R` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN qed[range_lt;RING_POW_0;RING_MUL_RID] ; pass ] THEN qed[] ; pass ] THEN subgoal `ring_sum r (range(n+1)) c = c n:R` THENL [ have `c(n:num):R IN ring_carrier r` [] THEN specialize[ `r:R ring`; `n:num`; `c:num->R` ]ring_sum_range_add_1 THEN subgoal `ring_sum r (range n) c = ring_0 r:R` THENL [ sufficesby RING_SUM_EQ_0 THEN qed[range_lt] ; pass ] THEN qed[RING_ADD_RZERO] ; pass ] THEN intro THEN proven_if `i < n:num` [] THEN num_linear_fact `i < n+1 ==> ~(i < n) ==> i = n` THEN qed[] ] );; let vandermonde_indep = prove(` !(r:R ring) S c v. integral_domain r ==> FINITE S ==> (!s:X. s IN S ==> c s IN ring_carrier r) ==> (!s. s IN S ==> v s IN ring_carrier r) ==> (!s t. s IN S ==> t IN S ==> v s = v t ==> s = t) ==> (!e. e < CARD S ==> ring_sum r S (\s. ring_mul r (c s) (ring_pow r (v s) e)) = ring_0 r) ==> (!s. s IN S ==> c s = ring_0 r) `, intro THEN choose `f:num->X` `IMAGE f (range(CARD S)) = S:X->bool` [finite_ordering] THEN subgoal `!i. i < CARD(S:X->bool) ==> (c:X->R o f) i = ring_0 r` THENL [ subgoal `!i. i < CARD(S:X->bool) ==> f i:X IN S` THENL [ intro THEN have `i IN range(CARD(S:X->bool))` [range_lt] THEN once_rw[GSYM(know `IMAGE f (range(CARD S)) = S:X->bool`)] THEN rw[IN_IMAGE] THEN qed[] ; pass ] THEN subgoal `!i. i < CARD(S:X->bool) ==> (c:X->R o f) i:R IN ring_carrier r` THENL [ rw[o_THM] THEN qed[] ; pass ] THEN subgoal `!i. i < CARD(S:X->bool) ==> (v:X->R o f) i:R IN ring_carrier r` THENL [ rw[o_THM] THEN qed[] ; pass ] THEN subgoal `!i j. i < CARD(S:X->bool) ==> j < CARD S ==> (v:X->R o f) i = (v o f) j ==> i = j` THENL [ rw[o_THM] THEN intro THEN have `f(i:num):X = f j` [] THEN qed[injective_finite_ordering] ; pass ] THEN subgoal `!e. e < CARD(S:X->bool) ==> ring_sum r (range (CARD S)) (\i. ring_mul r ((c:X->R o f) i) (ring_pow r ((v o f) i) e)) = ring_0 r:R` THENL [ rw[o_THM] THEN intro THEN have `!x y. x IN range (CARD(S:X->bool)) ==> y IN range (CARD S) ==> f x = f y:X ==> x = y` [injective_finite_ordering;range_lt] THEN specialize[ `r:R ring`; `f:num->X`; `\s:X. ring_mul r (c s) (ring_pow r (v s) e):R`; `range(CARD(S:X->bool))` ]RING_SUM_IMAGE THEN have `ring_sum r S (\s:X. ring_mul r (c s) (ring_pow r (v s) e)) = ring_0 r:R` [] THEN have `ring_sum r (IMAGE f (range (CARD(S:X->bool)))) (\s:X. ring_mul r (c s) (ring_pow r (v s) e)) = ring_0 r:R` [] THEN have `ring_sum r (range (CARD(S:X->bool))) ((\s:X. ring_mul r (c s) (ring_pow r (v s) e)) o f) = ring_0 r:R` [] THEN subgoal `ring_sum r (range (CARD S)) (\i. ring_mul r (c (f i)) (ring_pow r (v (f i)) e)) = ring_sum r (range (CARD(S:X->bool))) ((\s:X. ring_mul r (c s) (ring_pow r (v s) e)) o f):R` THENL [ rw[o_DEF] ; pass ] THEN qed[] ; pass ] THEN specialize[ `r:R ring`; `CARD(S:X->bool)`; `c:X->R o f:num->X`; `v:X->R o f:num->X` ]vandermonde_indep_range THEN qed[] ; pass ] THEN choose `i:num` `i IN range(CARD(S:X->bool)) /\ s:X = f i` [IN_IMAGE] THEN have `i < CARD(S:X->bool)` [range_lt] THEN have `c(f(i:num):X):R = ring_0 r` [o_THM] THEN qed[injective_finite_ordering] );; let expformal_linearly_independent = prove(` !S:complex->bool B:complex->complex. FINITE S ==> poly_sum complex_ring S ( \z. poly_mul complex_ring ( poly_const complex_ring (B z) ) ( expformal z ) ) = poly_0 complex_ring ==> (!z. z IN S ==> B z = Cx(&0)) `, intro THEN subgoal `!d. ring_sum complex_ring S (\z:complex. B z * z pow d) = Cx(&0)` THENL [ intro THEN have `coeff d (poly_0 complex_ring) = Cx(&0)` [coeff_poly_0;complex_ring_clauses] THEN subgoal `coeff d (poly_sum complex_ring S (\z. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal z))) = ring_sum complex_ring S (\z. coeff d (poly_mul complex_ring (poly_const complex_ring (B z)) (expformal z)))` THENL [ specialize[ `complex_ring`; `\z. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal z)` ]coeff_poly_sum THEN qed[series_complex] ; pass ] THEN subgoal `ring_sum complex_ring S (\z. coeff d (poly_mul complex_ring (poly_const complex_ring (B z)) (expformal z))) = ring_sum complex_ring S (\z. (B z * z pow d) / Cx(&(FACT d)))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN simp[coeff_poly_const_times;series_complex;in_complex_ring] THEN rw[expformal;coeff_series_from_coeffs;complex_ring_clauses] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN subgoal `ring_sum complex_ring S (\z. (B z * z pow d) / Cx(&(FACT d))) = ring_sum complex_ring S (\z. B z * z pow d) / Cx(&(FACT d))` THENL [ simp[GSYM vsum_ring_sum_complex] THEN rw[complex_div] THEN specialize[ `inv(Cx(&(FACT d)))`; `\z:complex. B z * z pow d`; `S:complex->bool` ]VSUM_COMPLEX_RMUL THEN qed[] ; pass ] THEN have `ring_sum complex_ring S (\z. B z * z pow d) / Cx(&(FACT d)) = Cx (&0)` [] THEN have `~(FACT d = 0)` [FACT_NZ] THEN have `~(&(FACT d) = &0:real)` [REAL_OF_NUM_EQ] THEN have `~(Cx(&(FACT d)) = Cx(&0))` [CX_INJ] THEN complex_field_fact `~(Cx (&(FACT d)) = Cx (&0)) ==> ring_sum complex_ring S (\z. B z * z pow d) / Cx (&(FACT d)) = Cx (&0) ==> ring_sum complex_ring S (\z. B z * z pow d) = Cx (&0)` THEN qed[] ; pass ] THEN recall integral_domain_complex THEN have `!s:complex. s IN S ==> B s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!s:complex. s IN S ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!s t:complex. s IN S ==> t IN S ==> I s = I t ==> s = t` [I_THM] THEN subgoal `!e. e < CARD S ==> ring_sum complex_ring S (\s. ring_mul complex_ring (B s) (ring_pow complex_ring (I s) e)) = ring_0 complex_ring` THENL [ rw[I_THM;ring_pow_complex;complex_ring_clauses] THEN qed[] ; pass ] THEN specialize[ `complex_ring`; `S:complex->bool`; `B:complex->complex`; `I:complex->complex` ]vandermonde_indep THEN qed[complex_ring_clauses] );; let expformal_perm_linearly_independent = prove(` !S:complex->bool B:complex->complex i:complex->complex. FINITE S ==> i permutes S ==> poly_sum complex_ring S ( \z. poly_mul complex_ring ( poly_const complex_ring (B z) ) ( expformal (i z) ) ) = poly_0 complex_ring ==> (!z. z IN S ==> B z = Cx(&0)) `, intro THEN subgoal `poly_sum complex_ring S (\z. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal (i z))) = poly_sum complex_ring S (\z. poly_mul complex_ring ( poly_const complex_ring (B(inverse i z))) (expformal z))` THENL [ rw[poly_sum] THEN have `!x y:complex. x IN S ==> y IN S ==> i x = i y:complex ==> x = y` [PERMUTES_INJECTIVE] THEN specialize[ `x_series complex_ring`; `i:complex->complex`; `\z. poly_mul complex_ring (poly_const complex_ring (B (inverse i z:complex))) (expformal z)`; `S:complex->bool` ]RING_SUM_IMAGE THEN have `IMAGE (i:complex->complex) S = S` [PERMUTES_IMAGE] THEN have_rw` ring_sum (x_series complex_ring) S (\z. poly_mul complex_ring (poly_const complex_ring (B(inverse i z))) (expformal z)) = ring_sum (x_series complex_ring) S ((\z. poly_mul complex_ring (poly_const complex_ring (B (inverse i z))) (expformal z)) o i)` [] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;o_THM] THEN qed[PERMUTES_INVERSES] ; pass ] THEN have `!z:complex. z IN S ==> B(inverse i z:complex) = Cx(&0)` [expformal_linearly_independent] THEN intro THEN have `z = (inverse i) (i z:complex):complex` [PERMUTES_INVERSES] THEN have `i(z:complex):complex IN S` [PERMUTES_IN_IMAGE] THEN qed[] );; (* ===== expanding prod_pi sum c_z exp(pi(z)) *) let ring_product_perm_sum_mul_exp_expand = prove(` !(a:A ring) (b:B ring) E:A->B c:A->B S:A->bool. (!x:A. x IN ring_carrier a ==> E x IN ring_carrier b) ==> E(ring_0 a) = ring_1 b ==> (!x y:A. E(ring_add a x y) = ring_mul b (E x) (E y)) ==> FINITE S ==> S SUBSET ring_carrier a ==> (!s:A. s IN S ==> c s IN ring_carrier b) ==> ring_product b (perm S) ( \i. ring_sum b S ( \z. ring_mul b (c z) (E(i z)) ) ) = ring_sum b ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) ( \e. ring_mul b ( ring_product b {y | ~(e y = 0)} ( \y. ring_pow b (c y) (e y) ) ) ( ring_sum b {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} ( E o (\f. ring_sum a (perm S) (\i. i (f i))) ) ) ) `, intro THEN rw[o_DEF] THEN have `FINITE(perm(S:A->bool))` [finite_perm] THEN have `FINITE(functions (perm S) (S:A->bool))` [finite_functions] THEN subgoal `!p q. p IN perm S ==> q:A IN S ==> p q IN S` THENL [ intro THEN have `p:A->A permutes S` [perm;IN] THEN qed[PERMUTES_IN_IMAGE] ; pass ] THEN have `!f p. f IN functions (perm S) S ==> p IN perm S ==> f p:A IN S` [image_functions_subset;SUBSET;IN_IMAGE] THEN have `!q:A. q IN S ==> q IN ring_carrier a` [SUBSET] THEN have `!p q. p IN perm S ==> q:A IN S ==> ring_mul b (c q) (E(p q)):B IN ring_carrier b` [RING_MUL] THEN specialize[ `b:B ring`; `\i:A->A z:A. ring_mul b (c z) (E(i z)):B`; `S:A->bool`; `perm(S:A->bool)` ]ring_product_sum_expand THEN simp[] THEN subgoal `ring_sum b (functions (perm S) S) (\f. ring_product b (perm S) (\i:A->A. ring_mul b (c (f i)) (E (i (f i))))) = ring_sum b (functions (perm S) S) (\f. ring_mul b (ring_product b (perm S) (c o f)) (E(ring_sum a (perm S) (\i. i(f i))))):B` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;o_DEF] THEN intro THEN have `!i:A->A. i IN perm S ==> c(a' i:A):B IN ring_carrier b /\ E(i(a' i)) IN ring_carrier b` [] THEN specialize[ `b:B ring`; `\i. c ((a':(A->A)->A) i):B`; `\i. E(i((a':(A->A)->A) i)):B`; `perm(S:A->bool)` ]RING_PRODUCT_MUL THEN simp[] THEN subgoal `E (ring_sum a (perm S) (\i:A->A. i (a' i))) = ring_product b (perm S) (E o (\i. i (a' i))):B` THENL [ sufficesby ring_exp_sum THEN qed[] ; pass ] THEN simp[o_DEF] ; pass ] THEN simp[] THEN subgoal `ring_sum b (functions (perm S) S) (\f. ring_mul b (ring_product b (perm S) (c o f)) (E(ring_sum a (perm S) (\i. i(f i))))) = ring_sum b (functions (perm S) S) (\f. ring_mul b (ring_product b {y | ~(numpreimages f (perm S) y = 0)} (\y. ring_pow b (c y) (numpreimages f (perm S) y))) (E(ring_sum a (perm S) (\i:A->A. i(f i)))):B)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `!y:A. y IN IMAGE a' (perm(S:A->bool)) ==> c y:B IN ring_carrier b` THENL [ rw[IN_IMAGE] THEN qed[] ; pass ] THEN specialize[ `b:B ring`; `perm(S:A->bool)`; `a':(A->A)->A`; `c:A->B` ]ring_product_o_v2 THEN qed[] ; pass ] THEN simp[] THEN specialize[ `b:B ring`; `\f:(A->A)->A. numpreimages f (perm S)`; `\f:(A->A)->A. ring_mul b (ring_product b {y | ~(numpreimages f (perm S) y = 0)} (\y. ring_pow b (c y) (numpreimages f (perm S) y))) (E (ring_sum a (perm S) (\i. i (f i)))):B`; `functions (perm S) (S:A->bool)` ]RING_SUM_IMAGE_GEN THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_IMAGE] THEN intro THEN specialize[ `functions (perm S) (S:A->bool)`; `\f:(A->A)->A. numpreimages f (perm S) = a'` ]FINITE_RESTRICT THEN subgoal `ring_sum b {x:(A->A)->A | x IN functions (perm S) S /\ numpreimages x (perm S) = a'} (\g. ring_mul b (ring_product b {y | ~(numpreimages g (perm S) y = 0)} (\y. ring_pow b (c y) (numpreimages g (perm S) y))) (E (ring_sum a (perm S) (\i. i (g i)))):B) = ring_sum b {x | x IN functions (perm S) S /\ numpreimages x (perm S) = a'} (\g. ring_mul b (ring_product b {y | ~(a' y = 0)} (\y. ring_pow b (c y) (a' y))) (E (ring_sum a (perm S) (\i. i (g i)))))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN simp[] ; pass ] THEN simp[] THEN specialize_assuming[ `b:B ring`; `\f:(A->A)->A. E (ring_sum a (perm S) (\i. i (f i))):B`; `ring_product b {y | ~(numpreimages x (perm(S:A->bool)) y = 0)} (\a:A. ring_pow b (c a) (numpreimages x (perm S) a)):B`; `{f:(A->A)->A | f IN functions (perm S) S /\ numpreimages f (perm S) = numpreimages x (perm S)}` ]RING_SUM_LMUL THEN qed[RING_PRODUCT;RING_SUM] );; let product_perm_sum_mul_cexp_expand = prove(` !c:complex->complex S:complex->bool. FINITE S ==> ring_product complex_ring (perm S) ( \i. ring_sum complex_ring S ( \z. ring_mul complex_ring (c z) (cexp(i z)) ) ) = ring_sum complex_ring ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) ( \e. ring_product complex_ring {y | ~(e y = 0)} ( \y. (c y) pow (e y) ) * ring_sum complex_ring {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} ( cexp o (\f. ring_sum complex_ring (perm S) (\i. i (f i))) ) ) `, intro THEN rw[GSYM ring_pow_complex] THEN rw[GSYM complex_ring_clauses] THEN specialize_assuming[ `complex_ring`; `complex_ring`; `cexp`; `c:complex->complex`; `S:complex->bool` ]ring_product_perm_sum_mul_exp_expand THEN qed[in_complex_ring;CEXP_0;CEXP_ADD;SUBSET;complex_ring_clauses] );; let product_perm_sum_mul_expformal_expand = prove(` !c:complex->((1->num)->complex) S:complex->bool. FINITE S ==> ring_product(x_series complex_ring) (perm S) ( \i. ring_sum(x_series complex_ring) S ( \z. poly_mul complex_ring (c z) (expformal(i z)) ) ) = ring_sum(x_series complex_ring) ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) ( \e. ring_mul(x_series complex_ring) ( ring_product(x_series complex_ring) {y | ~(e y = 0)} ( \y. ring_pow(x_series complex_ring) (c y) (e y) ) ) ( ring_sum(x_series complex_ring) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} ( expformal o (\f. ring_sum complex_ring (perm S) (\i. i (f i))) ) ) ) `, intro THEN rw[x_series_use] THEN specialize_assuming[ `complex_ring`; `x_series complex_ring`; `expformal`; `c:complex->(1->num)->complex`; `S:complex->bool` ]ring_product_perm_sum_mul_exp_expand THEN qed[complex_ring_clauses;in_complex_ring;series_complex;x_series_use;expformal_0;mul_expformal;SUBSET] );; let product_perm_sum_mul_expformal_expand_v2 = prove(` !c:complex->((1->num)->complex) S:complex->bool. FINITE S ==> ring_product(x_series complex_ring) (perm S) ( \i. ring_sum(x_series complex_ring) S ( \z. poly_mul complex_ring (c z) (expformal(i z)) ) ) = ring_sum(x_series complex_ring) ( IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) (functions (perm S) S) ) (\y. ring_sum(x_series complex_ring) ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) ( \e. poly_mul complex_ring ( ring_product(x_series complex_ring) {z | ~(e z = 0)} ( \z. ring_pow(x_series complex_ring) (c z) (e z) ) ) ( poly_mul complex_ring ( ring_of_num(x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} y) ) (expformal y) ) ) ) `, simp[product_perm_sum_mul_expformal_expand] THEN rw[x_series_use] THEN intro THEN have `FINITE(perm(S:complex->bool))` [finite_perm] THEN have `FINITE(functions (perm(S:complex->bool)) (S:complex->bool))` [finite_functions] THEN have `FINITE(IMAGE (\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))) (functions (perm S) S))` [FINITE_IMAGE] THEN have `FINITE(IMAGE (\f. numpreimages f (perm S)) (functions (perm S) (S:complex->bool)))` [FINITE_IMAGE] THEN have `!i j. i IN IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) (functions (perm S) S) ==> j IN IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ==> ring_mul (x_series complex_ring) (ring_product (x_series complex_ring) {z | ~(j z = 0)} (\z. ring_pow (x_series complex_ring) (c z) (j z))) (ring_mul (x_series complex_ring) (ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = j} i)) (expformal i)) IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN specialize[ `x_series complex_ring`; `\y e. ring_mul(x_series complex_ring) (ring_product (x_series complex_ring) {z | ~(e z = 0)} (\z. ring_pow (x_series complex_ring) (c z) (e z))) (ring_mul(x_series complex_ring) (ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} y)) (expformal y))`; `IMAGE (\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))) (functions (perm S) S)`; `IMAGE (\f. numpreimages f (perm S)) (functions (perm S) (S:complex->bool))` ]RING_SUM_SWAP THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `ring_product (x_series complex_ring) {z | ~(a z = 0)} (\z:complex. ring_pow (x_series complex_ring) (c z) (a z)) IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN have `!z. z IN IMAGE (\f. ring_sum complex_ring (perm S) (\y. y (f y))) (functions (perm S) S) ==> ring_mul (x_series complex_ring) (ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} z)) (expformal z) IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN specialize[ `x_series complex_ring`; `\y. ring_mul (x_series complex_ring) (ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} y)) (expformal y)`; `ring_product (x_series complex_ring) {z:complex | ~(a z = 0)} (\z. ring_pow (x_series complex_ring) (c z) (a z))`; `IMAGE (\f. ring_sum complex_ring (perm S) (\y:complex->complex. y (f y))) (functions (perm S) S)` ]RING_SUM_LMUL THEN simp[] THEN subgoal `ring_sum (x_series complex_ring) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} (expformal o (\f. ring_sum complex_ring (perm S) (\i. i (f i)))) = ring_sum (x_series complex_ring) (IMAGE (\f. ring_sum complex_ring (perm S) (\y. y (f y))) (functions (perm S) S)) (\y. ring_mul (x_series complex_ring) (ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} y)) (expformal y))` THENL [ specialize[ `functions (perm (S:complex->bool)) (S:complex->bool)`; `\f:(complex->complex)->complex. numpreimages f (perm(S:complex->bool)) = a` ]FINITE_RESTRICT THEN have `!y. y IN IMAGE (\f. ring_sum complex_ring (perm S) (\y:complex->complex. y (f y))) (functions (perm S) S) ==> expformal y IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN have `!y. y IN IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} ==> expformal y IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN subgoal `IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a} SUBSET IMAGE (\f. ring_sum complex_ring (perm S) (\y:complex->complex. y (f y))) (functions (perm S) S)` THENL [ rw[SUBSET;in_image_vw;IN_ELIM_THM] THEN qed[] ; pass ] THEN specialize[ `x_series complex_ring`; `{f | f IN functions (perm S) (S:complex->bool) /\ numpreimages f (perm S) = a}`; `\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))`; `IMAGE (\f. ring_sum complex_ring (perm S) (\y:complex->complex. y (f y))) (functions (perm S) S)`; `expformal` ]ring_sum_o_v3 THEN simp[] ; pass ] THEN simp[] );; let product_perm_sum_mul_expformal_expand_v3 = prove(` !c:complex->complex S:complex->bool. FINITE S ==> ring_product(x_series complex_ring) (perm S) ( \i. ring_sum(x_series complex_ring) S ( \z. poly_mul complex_ring (poly_const complex_ring (c z)) (expformal(i z)) ) ) = ring_sum(x_series complex_ring) ( IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) (functions (perm S) S) ) (\y. poly_mul complex_ring ( poly_const complex_ring ( ring_sum complex_ring ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) (\e. ring_product complex_ring {z | ~(e z = 0)} ( \z. (c z) pow (e z) ) * ( Cx(&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} y)) ) ) ) ) (expformal y) ) `, simp[product_perm_sum_mul_expformal_expand_v2] THEN rw[x_series_use] THEN intro THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `FINITE(perm(S:complex->bool))` [finite_perm] THEN have `FINITE(functions (perm(S:complex->bool)) (S:complex->bool))` [finite_functions] THEN have `FINITE(IMAGE (\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))) (functions (perm S) S))` [FINITE_IMAGE] THEN have `FINITE(IMAGE (\f. numpreimages f (perm S)) (functions (perm S) (S:complex->bool)))` [FINITE_IMAGE] THEN have `!s. s IN IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ==> ring_product complex_ring {z | ~(s z = 0)} (\z. c z pow s z) * Cx (&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `\e. ring_product complex_ring {z | ~(e z = 0)} (\z. c z pow e z) * Cx(&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} a))`; `IMAGE (\f. numpreimages f (perm S)) (functions (perm S) (S:complex->bool))` ]poly_const_sum THEN simp[poly_const_sum] THEN subgoal `poly_sum complex_ring (IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S)) (\s. poly_const complex_ring ((ring_product complex_ring {z | ~(s z = 0)} (\z. c z pow s z)) * (Cx (&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a))))) = poly_sum complex_ring (IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S)) (\s. poly_mul complex_ring (ring_product(x_series complex_ring) {z | ~(s z = 0)} (\z. ring_pow(x_series complex_ring) (poly_const complex_ring (c z)) (s z))) (poly_const complex_ring (Cx (&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a)))))` THENL [ sufficesby poly_sum_eq THEN intro THEN have `ring_product complex_ring {z | ~(s z = 0)} (\z:complex. c z pow s z) IN ring_carrier complex_ring` [in_complex_ring] THEN have `Cx(&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `ring_product complex_ring {z | ~(s z = 0)} (\z:complex. c z pow s z)`; `Cx(&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a))` ]POLY_CONST_MUL THEN rw[GSYM complex_ring_clauses] THEN simp[] THEN subgoal `FINITE {z:complex | ~(s z = 0)}` THENL [ choose `f:(complex->complex)->complex` `s:complex->num = (\f. numpreimages f (perm S)) f /\ f IN functions (perm S) S` [IN_IMAGE] THEN have `s = numpreimages (f:(complex->complex)->complex) (perm S)` [] THEN subgoal `{z:complex | ~(s z = 0)} = {z | ~(numpreimages f (perm(S:complex->bool)) z = 0)}` THENL [ rw[EXTENSION;IN_ELIM_THM] THEN qed[] ; pass ] THEN simp[GSYM image_numpreimages] THEN qed[FINITE_IMAGE] ; pass ] THEN have `!t. t IN {z:complex | ~(s z = 0)} ==> c t pow s t IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `\z:complex. (c z pow s z):complex`; `{z:complex | ~(s z = 0)}` ]poly_const_product THEN simp[poly_product] THEN subgoal `ring_product (x_series complex_ring) {z:complex | ~(s z = 0)} (\t. poly_const complex_ring (c t pow s t)) = ring_product (x_series complex_ring) {z | ~(s z = 0)} (\z. ring_pow (x_series complex_ring) (poly_const complex_ring (c z)) (s z))` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN have `c(a':complex) IN ring_carrier complex_ring` [in_complex_ring] THEN rw[GSYM ring_pow_complex] THEN simp[poly_const_pow;x_series_use_pow] ; pass ] THEN simp[] ; pass ] THEN rw[GSYM poly_sum] THEN simp[] THEN rw[poly_sum;x_series_use] THEN have `expformal a IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN have `!e. e IN IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ==> ring_mul (x_series complex_ring) (ring_product (x_series complex_ring) {z | ~(e z = 0)} (\z. ring_pow (x_series complex_ring) (poly_const complex_ring (c z)) (e z))) (poly_const complex_ring (Cx (&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = e} a)))) IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN specialize[ `x_series complex_ring`; `\s. ring_mul (x_series complex_ring) (ring_product (x_series complex_ring) {z | ~(s z = 0)} (\z. ring_pow (x_series complex_ring) (poly_const complex_ring (c z)) (s z))) (poly_const complex_ring (Cx(&(numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = s} a))))`; `expformal a`; `IMAGE (\f. numpreimages f (perm S)) (functions (perm S) (S:complex->bool))` ](GSYM RING_SUM_RMUL) THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN have `ring_product (x_series complex_ring) {z:complex | ~(a' z = 0)} (\z. ring_pow (x_series complex_ring) (poly_const complex_ring (c z)) (a' z)) IN ring_carrier (x_series complex_ring)` [RING_PRODUCT] THEN have `ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i. i (f i))) {f | f IN functions (perm S) S /\ numpreimages f (perm S) = a'} a) IN ring_carrier (x_series complex_ring)` [RING_OF_NUM] THEN specialize_assuming[ `x_series complex_ring`; `ring_product (x_series complex_ring) {z | ~(a' z = 0)} (\z:complex. ring_pow (x_series complex_ring) (poly_const complex_ring (c z)) (a' z))`; `ring_of_num (x_series complex_ring) (numpreimages (\f. ring_sum complex_ring (perm S) (\i:complex->complex. i (f i))) {f | f IN functions (perm S) (S:complex->bool) /\ numpreimages f (perm S) = a'} a)`; `expformal a` ]RING_MUL_ASSOC THEN simp[] THEN rw[ring_of_num_x_series] THEN rw[ring_of_num_complex;complex_of_num] );; (* ===== more general transcendence results, part 1: alpha symmetrization *) let zero_sum_QinC_exp_squarefree_roots_lemma_sym_powersums = prove(` !p e d. ring_polynomial QinC_ring p ==> ring_squarefree(x_poly QinC_ring) p ==> monic QinC_ring p ==> ring_sum complex_ring {f | f IN functions (perm(complex_root p)) (complex_root p) /\ numpreimages f (perm(complex_root p)) = e} (\f. (ring_sum complex_ring (perm(complex_root p)) (\i. i (f i))) pow d) IN QinC `, intro THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [poly_0_QinC_eq_poly_0_complex] THEN have `FINITE(complex_root p)` [complex_root_le_deg] THEN have `FINITE(perm(complex_root p))` [finite_perm] THEN have `FINITE(functions (perm(complex_root p)) (complex_root p))` [finite_functions] THEN specialize[ `functions (perm (complex_root p)) (complex_root p)`; `\f:(complex->complex)->complex. numpreimages f (perm(complex_root p)) = e` ]FINITE_RESTRICT THEN def `Z:(complex->num)->complex` `ring_sum(poly_ring complex_ring (:complex)) {f | f IN functions (perm(complex_root p)) (complex_root p) /\ numpreimages f (perm(complex_root p)) = e} (\f. poly_pow complex_ring (ring_sum(poly_ring complex_ring (:complex)) (perm(complex_root p)) (\i. poly_var complex_ring (i (f i)))) d)` THEN subgoal `ring_sum complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\f. ring_sum complex_ring (perm (complex_root p)) (\i. i (f i)) pow d) = poly_evaluate complex_ring Z I` THENL [ simp[] THEN subgoal `(!s. s IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> ring_polynomial complex_ring (poly_pow complex_ring (ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d))` THENL [ intro THEN have `ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i))) IN ring_carrier(poly_ring complex_ring (:complex))` [RING_SUM] THEN qed[poly_pow_in_poly_ring;poly_in_full_ring;poly_ring_use_pow] ; pass ] THEN specialize[ `complex_ring`; `\f. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (f i)))) d`; `I:complex->complex`; `{f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` ]poly_evaluate_sum THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN have `ring_polynomial complex_ring (ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i))))` [RING_SUM;poly_in_full_ring] THEN have `(!i. i IN poly_vars complex_ring (ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i)))) ==> I i IN ring_carrier complex_ring)` [in_complex_ring] THEN specialize[ `complex_ring`; `ring_sum (poly_ring complex_ring (:complex)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i)))`; `I:complex->complex`; `d:num` ]poly_evaluate_pow THEN simp[] THEN rw[ring_pow_complex] THEN subgoal `poly_evaluate complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i)))) I = ring_sum complex_ring (perm (complex_root p)) (\i. i (a i))` THENL [ have `!s. s IN perm (complex_root p) ==> ring_polynomial complex_ring (poly_var complex_ring (s (a s)))` [RING_POLYNOMIAL_VAR] THEN specialize[ `complex_ring`; `\i:complex->complex. poly_var complex_ring (i (a i))`; `I:complex->complex`; `perm (complex_root p)` ]poly_evaluate_sum THEN simp[] THEN sufficesby RING_SUM_EQ THEN qed[POLY_EVALUATE_VAR;in_complex_ring;I_THM] ; pass ] THEN qed[] ; pass ] THEN rw[know `ring_sum complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\f. ring_sum complex_ring (perm (complex_root p)) (\i. i (f i)) pow d) = poly_evaluate complex_ring Z I`] THEN have `!s. s IN complex_root p ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN subgoal `!d. coeff d (monic_vanishing_at complex_ring (complex_root p) I) IN ring_carrier (subring_generated complex_ring QinC)` THENL [ have `ring_squarefree(x_poly complex_ring) p` [monic_QinC_squarefree_complex_squarefree] THEN have `monic complex_ring p` [monic_subring;subring_complex_QinC] THEN simp[monic_vanishing_at_complex_root] THEN simp[subring_complex_QinC] THEN qed[coeff_poly_in_ring] ; pass ] THEN subgoal `ring_polynomial (subring_generated complex_ring QinC) (Z:(complex->num)->complex)` THENL [ simp[] THEN subgoal `(!s. s IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> ring_polynomial (subring_generated complex_ring QinC) (poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d))` THENL [ rw[IN_ELIM_THM] THEN intro THEN subgoal `poly_pow (subring_generated complex_ring QinC) (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d = poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d` THENL [ sufficesby poly_pow_subring THEN sufficesby ring_powerseries_if_polynomial THEN sufficesby ring_polynomial_subring_sum THEN qed[ring_polynomial_subring_var] ; pass ] THEN subgoal `ring_polynomial (subring_generated complex_ring QinC) (poly_pow (subring_generated complex_ring QinC) (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d)` THENL [ sufficesby poly_pow_poly THEN sufficesby ring_polynomial_subring_sum THEN qed[ring_polynomial_subring_var] ; pass ] THEN qed[] ; pass ] THEN specialize[ `complex_ring`; `QinC`; `\f. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (f i)))) d`; `{f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` ](GSYM poly_sum_subring_multi) THEN simp[] THEN qed[RING_SUM;poly_in_full_ring] ; pass ] THEN subgoal `poly_vars complex_ring (Z:(complex->num)->complex) SUBSET complex_root p` THENL [ simp[] THEN sufficesby poly_vars_sum_poly_subset THEN simp[IN_ELIM_THM;functions] THEN intro THENL [ sufficesby poly_pow_poly THEN qed[RING_SUM;poly_in_full_ring] ; sufficesby poly_vars_pow_subset THEN intro THENL [ sufficesby ring_powerseries_if_polynomial THEN qed[RING_SUM;poly_in_full_ring] ; sufficesby poly_vars_sum_poly_subset THEN simp[] THEN intro THENL [ qed[RING_POLYNOMIAL_VAR] ; rw[POLY_VARS_VAR;SUBSET] THEN have `s(s':complex->complex) IN complex_root p` [IN_IMAGE;SUBSET] THEN have `s' permutes complex_root p` [IN;perm] THEN have `s'(s(s':complex->complex)) IN complex_root p` [PERMUTES_IN_IMAGE] THEN ASM SET_TAC[] ] ] ] ; pass ] THEN subgoal `!g m. g permutes complex_root p ==> Z (m o g) = (Z:(complex->num)->complex) m` THENL [ intro THEN simp[] THEN subgoal `!s. s IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> ring_polynomial complex_ring (poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d)` THENL [ rw[IN_ELIM_THM] THEN intro THEN sufficesby poly_pow_poly THEN qed[RING_SUM;poly_in_full_ring] ; pass ] THEN specialize[ `complex_ring`; `(\f. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (f i)))) d)`; `m:complex->num`; `g:complex->complex`; `complex_root p`; `{f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` ]ring_sum_poly_o_permutes THEN simp[] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s m. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) d (m o g)) = ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s m. poly_pow complex_ring (\m. (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i)))) (m o g)) d m)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;FUN_EQ_THM;IN_ELIM_THM] THEN intro THEN subgoal `ring_powerseries complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i:complex->complex. poly_var complex_ring (i (a i))))` THENL [ sufficesby ring_powerseries_if_polynomial THEN sufficesby poly_sum_poly_multi THEN qed[RING_POLYNOMIAL_VAR] ; pass ] THEN specialize[ `complex_ring`; `ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i)))`; `complex_root p`; `g:complex->complex`; `d:num`; `x:complex->num` ]poly_pow_o_permutes THEN qed[] ; pass ] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s m. poly_pow complex_ring (\m. ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (s i))) (m o g)) d m) = ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (s i))))) d)` THENL [ sufficesby RING_SUM_EQ THEN once_rw[fun_eq_thm_v] THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN subgoal `(\m. ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i))) (m o g)) = (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (a i)))))` THENL [ rw[FUN_EQ_THM] THEN intro THEN have `!s. s IN perm (complex_root p) ==> ring_polynomial complex_ring (poly_var complex_ring (s (a s)))` [RING_POLYNOMIAL_VAR] THEN specialize[ `complex_ring`; `\i:complex->complex. poly_var complex_ring (i (a i))`; `x:complex->num`; `g:complex->complex`; `complex_root p`; `perm(complex_root p)` ]ring_sum_poly_o_permutes THEN simp[] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\s m. poly_var complex_ring (s (a s)) (m o g)) = ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (a i))))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN rw[FUN_EQ_THM] THEN qed[poly_var_o_permutes] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN simp[] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (s i))))) d) = ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\f. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (f i)))) d)` THENL [ subgoal `!x y. x IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> y IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> (\i. x ((g:complex->complex) o i)) = (\i. y (g o i)) ==> x = y` THENL [ once_rw[fun_eq_thm_v] THEN rw[FUN_EQ_THM] THEN intro THEN have `I:complex->complex = (g:complex->complex) o inverse g` [PERMUTES_INVERSES_o] THEN have `v:complex->complex = ((g:complex->complex) o inverse g) o v` [I_O_ID;I_THM] THEN subgoal `v:complex->complex = (g:complex->complex) o inverse g o v` THENL [ rw[o_ASSOC] THEN qed[] ; pass ] THEN have `x(v:complex->complex):complex = x ((g:complex->complex) o (inverse g o v))` [] THEN have `y(v:complex->complex):complex = x ((g:complex->complex) o (inverse g o v))` [] THEN qed[] ; pass ] THEN specialize[ `poly_ring complex_ring (:complex)`; `\f:(complex->complex)->complex i:complex->complex. f ((g:complex->complex) o i)`; `\s. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (s i))))) d`; `{f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` ]RING_SUM_IMAGE THEN subgoal `IMAGE (\f i. f (g o i)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} = {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` THENL [ rw[EXTENSION;IN_ELIM_THM;IN_IMAGE] THEN intro THEN splitiff THENL [ intro THENL [ specialize[ `complex_root p`; `complex_root p`; `g:complex->complex` ]image_permutes_o_functions_perm THEN ASM SET_TAC[] ; rw[know `x:(complex->complex)->complex = (\i. x' ((g:complex->complex) o i))`] THEN qed[numpreimages_permutes_o_perm] ] ; intro THEN witness `\i:complex->complex. x(inverse g:complex->complex o i):complex` THEN have `inverse g permutes complex_root p` [PERMUTES_INVERSE] THEN intro THENL [ rw[fun_eq_thm_e] THEN qed[inverse_permutes_o_refl_o] ; specialize[ `complex_root p`; `complex_root p`; `inverse g:complex->complex` ]image_permutes_o_functions_perm THEN ASM SET_TAC[] ; qed[numpreimages_permutes_o_perm] ] ] ; pass ] THEN have_rw `ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (s i))))) d) = ring_sum (poly_ring complex_ring (:real^2)) {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ((\s. poly_pow complex_ring (ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (s i))))) d) o (\f i. f (g o i)))` [] THEN sufficesby RING_SUM_EQ THEN rw[IN_ELIM_THM;BETA_THM;o_THM] THEN intro THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (g (i (a (g o i))))) = ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i)))` THENL [ have `!x y. x IN perm (complex_root p) ==> y IN perm (complex_root p) ==> g o x = (g:complex->complex) o y ==> x = y` [inverse_permutes_o_refl_o] THEN specialize[ `poly_ring complex_ring (:complex)`; `\i:complex->complex. (g:complex->complex) o i`; `\i:complex->complex. poly_var complex_ring (i (a i))`; `perm(complex_root p)` ]RING_SUM_IMAGE THEN specialize[ `complex_root p`; `g:complex->complex` ]image_permutes_o_perm THEN have_rw `ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) (\i. poly_var complex_ring (i (a i))) = ring_sum (poly_ring complex_ring (:real^2)) (perm (complex_root p)) ((\i. poly_var complex_ring (i (a i))) o (\i. g o i))` [] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;o_THM] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN specialize[ `complex_ring`; `QinC`; `complex_root p`; `I:complex->complex`; `Z:(complex->num)->complex` ]symmetric_subring_if_poly_subring THEN qed[subring_complex_QinC;QinC_ring_clauses] );; let zero_sum_QinC_exp_squarefree_roots_lemma_sym_coeff_poly = prove(` !p e n. ring_polynomial QinC_ring p ==> ring_squarefree(x_poly QinC_ring) p ==> monic QinC_ring p ==> coeff n ( monic_vanishing_at complex_ring {f | f IN functions (perm(complex_root p)) (complex_root p) /\ numpreimages f (perm(complex_root p)) = e} (\f. (ring_sum complex_ring (perm(complex_root p)) (\i. i (f i)))) ) IN QinC `, intro THEN recall subring_complex_QinC THEN recall QinC_ring_clauses THEN subgoal `!d. ring_sum complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\s. ring_pow complex_ring (ring_sum complex_ring (perm (complex_root p)) (\i. i (s i))) d) IN ring_carrier (subring_generated complex_ring QinC)` THENL [ specialize[ `p:(1->num)->complex`; `e:complex->num` ]zero_sum_QinC_exp_squarefree_roots_lemma_sym_powersums THEN rw[ring_pow_complex] THEN qed[] ; pass ] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [poly_0_QinC_eq_poly_0_complex] THEN have `FINITE(complex_root p)` [complex_root_le_deg] THEN have `FINITE(perm(complex_root p))` [finite_perm] THEN have `FINITE(functions (perm(complex_root p)) (complex_root p))` [finite_functions] THEN specialize[ `functions (perm (complex_root p)) (complex_root p)`; `\f:(complex->complex)->complex. numpreimages f (perm(complex_root p)) = e` ]FINITE_RESTRICT THEN have `ring_hasQ (subring_generated complex_ring QinC)` [subring_complex_QinC;ring_hasQ_QinC] THEN have `!s. s IN {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} ==> ring_sum complex_ring (perm (complex_root p)) (\i. i (s i)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `QinC`; `{f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}`; `\f. ring_sum complex_ring (perm (complex_root p)) (\i:complex->complex. i (f i))` ]coeff_poly_subring_if_powersums_subring THEN qed[] );; let zero_sum_QinC_exp_squarefree_roots_lemma_sym_poly = prove(` !p e. ring_polynomial QinC_ring p ==> ring_squarefree(x_poly QinC_ring) p ==> monic QinC_ring p ==> ring_polynomial QinC_ring ( monic_vanishing_at complex_ring {f | f IN functions (perm(complex_root p)) (complex_root p) /\ numpreimages f (perm(complex_root p)) = e} (\f. (ring_sum complex_ring (perm(complex_root p)) (\i. i (f i)))) ) `, intro THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [poly_0_QinC_eq_poly_0_complex] THEN have `FINITE(complex_root p)` [complex_root_le_deg] THEN have `FINITE(perm(complex_root p))` [finite_perm] THEN have `FINITE(functions (perm(complex_root p)) (complex_root p))` [finite_functions] THEN specialize[ `functions (perm (complex_root p)) (complex_root p)`; `\f:(complex->complex)->complex. numpreimages f (perm(complex_root p)) = e` ]FINITE_RESTRICT THEN rw[GSYM subring_complex_QinC] THEN sufficesby ring_polynomial_subring_if_coeffs THEN specialize[`p:(1->num)->complex`;`e:complex->num`]zero_sum_QinC_exp_squarefree_roots_lemma_sym_coeff_poly THEN specialize_assuming[ `complex_ring`; `{x | x IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages x (perm (complex_root p)) = e}`; `\f. ring_sum complex_ring (perm (complex_root p)) (\i:complex->complex. i (f i))` ]monic_vanishing_at_poly THEN qed[in_complex_ring;subring_complex_QinC;QinC_ring_clauses] );; let zero_sum_QinC_exp_squarefree_roots_lemma_denouement = prove(` !S:complex->bool B:complex->complex. FINITE S ==> ( !a. ring_sum complex_ring ( IMAGE (\f. numpreimages f (perm S)) (functions (perm S) S) ) ( \s. ring_product complex_ring {y | ~(s y = 0)} ( \y. (B y) pow (s y)) * Cx(&( numpreimages ( \f. ring_sum complex_ring (perm S) (\i. i (f i)) ) { f | f IN functions (perm S) (S) /\ numpreimages f (perm S) = s } a )) ) = Cx (&0) ) ==> (!z. z IN S ==> B z = Cx(&0)) `, intro THEN subgoal `ring_product (x_series complex_ring) (perm S) (\i. ring_sum (x_series complex_ring) S (\z. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal (i z)))) = ring_0(x_series complex_ring)` THENL [ specialize[ `B:complex->complex`; `S:complex->bool` ]product_perm_sum_mul_expformal_expand_v3 THEN simp[] THEN rw[GSYM complex_ring_clauses;POLY_CONST_0] THEN subgoal `ring_sum (x_series complex_ring) (IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) (functions (perm S) S)) (\a. poly_mul complex_ring (poly_0 complex_ring) (expformal a)) = ring_sum (x_series complex_ring) (IMAGE (\f. ring_sum complex_ring (perm S) (\i. i (f i))) (functions (perm S) S)) (\a. ring_0(x_series complex_ring))` THENL [ sufficesby RING_SUM_EQ THEN intro THEN have `ring_powerseries complex_ring (expformal a)` [series_complex] THEN simp[POWSER_MUL_0] THEN rw[x_series_use] ; pass ] THEN simp[RING_SUM_0] ; pass ] THEN have `integral_domain complex_ring` [integral_domain_complex] THEN have `integral_domain(x_series complex_ring)` [INTEGRAL_DOMAIN_POWSER_RING;x_series] THEN have `FINITE(perm(S:complex->bool))` [finite_perm] THEN choose `i:complex->complex` `i IN perm(S:complex->bool) /\ ring_sum (x_series complex_ring) S (\z. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal (i z))) = ring_0 (x_series complex_ring)` [INTEGRAL_DOMAIN_PRODUCT_EQ_0] THEN subgoal `poly_sum complex_ring S (\z:complex. poly_mul complex_ring (poly_const complex_ring (B z)) (expformal (i z))) = poly_0 complex_ring` THENL [ rw[poly_sum] THEN qed[x_series_use] ; pass ] THEN qed[expformal_perm_linearly_independent;IN;perm] );; let zero_sum_QinC_exp_squarefree_roots = prove(` !p B. ring_polynomial QinC_ring p ==> ring_squarefree(x_poly QinC_ring) p ==> monic QinC_ring p ==> (!z:complex. complex_root p z ==> B z IN QinC) ==> ring_sum complex_ring (complex_root p) (\z. (B z) * cexp z) = Cx(&0) ==> (!z. complex_root p z ==> B z = Cx(&0)) `, intro_gendisch THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [poly_0_QinC_eq_poly_0_complex] THEN have `FINITE(complex_root p)` [complex_root_le_deg] THEN have `FINITE(perm(complex_root p))` [finite_perm] THEN have `FINITE(functions (perm(complex_root p)) (complex_root p))` [finite_functions] THEN subgoal `ring_sum complex_ring (IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p))) (\e. ring_product complex_ring {y | ~(e y = 0)} (\y. (B y) pow (e y)) * ring_sum complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\f. cexp (ring_sum complex_ring (perm (complex_root p)) (\i. i (f i))))) = Cx (&0)` THENL [ rw[ISPEC `cexp` (GSYM o_DEF)] THEN specialize[ `B:complex->complex`; `complex_root p` ](GSYM product_perm_sum_mul_cexp_expand) THEN simp[] THEN rw[complex_ring_clauses] THEN have `I IN perm(complex_root p)` [IN;perm;PERMUTES_I] THEN subgoal `ring_sum complex_ring (complex_root p) (\z. (B z) * cexp (I z)) = Cx(&0)` THENL [ rw[I_THM] THEN qed[] ; pass ] THEN have `ring_sum complex_ring (complex_root p) (\z. B z * cexp (I z)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `perm(complex_root p)`; `I:complex->complex`; `\i. ring_sum complex_ring (complex_root p) (\z. (B z) * cexp (i z))` ]ring_product_delete THEN simp[] THEN rw[complex_ring_clauses] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN have `!e. ring_polynomial QinC_ring (monic_vanishing_at complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e} (\f. ring_sum complex_ring (perm (complex_root p)) (\i. i (f i))))` [zero_sum_QinC_exp_squarefree_roots_lemma_sym_poly] THEN have `FINITE(IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p)))` [FINITE_IMAGE] THEN subgoal `!s. s IN IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p)) ==> ring_product complex_ring {y | ~(s y = 0)} (\y. (B y) pow (s y)) IN QinC` THENL [ rw[IN_IMAGE] THEN intro THEN rw[GSYM QinC_ring_clauses] THEN rw[GSYM subring_complex_QinC] THEN sufficesby ring_product_in_subring THEN rw[IN_ELIM_THM;BETA_THM] THEN intro THEN rw[GSYM ring_pow_complex] THEN sufficesby ring_pow_in_subring THEN rw[QinC_ring_clauses;subring_complex_QinC] THEN subgoal `s' IN {y:complex | ~(numpreimages x (perm(complex_root p)) y = 0)}` THENL [ rw[IN_ELIM_THM] THEN qed[] ; pass ] THEN subgoal `s':complex IN IMAGE x (perm(complex_root p))` THENL [ simp[image_numpreimages] ; pass ] THEN have `IMAGE x (perm(complex_root p)) SUBSET complex_root p` [image_functions_subset] THEN have `s' IN complex_root p` [SUBSET] THEN qed[IN] ; pass ] THEN subgoal `!s. s IN IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p)) ==> FINITE {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = s}` THENL [ rw[IN_IMAGE] THEN intro THEN specialize[ `functions (perm (complex_root p)) (complex_root p)`; `\f:(complex->complex)->complex. numpreimages f (perm(complex_root p)) = s` ]FINITE_RESTRICT THEN qed[] ; pass ] THEN have `!s. s IN IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p)) ==> ring_polynomial QinC_ring (monic_vanishing_at complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = s} (\f. ring_sum complex_ring (perm (complex_root p)) (\i. i (f i))))` [] THEN subgoal `ring_sum complex_ring (IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p))) (\s. ring_product complex_ring {y | ~(s y = 0)} (\y. (B y) pow (s y)) * ring_sum complex_ring {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = s} (cexp o (\f. ring_sum complex_ring (perm (complex_root p)) (\i. i (f i))))) = Cx (&0)` THENL [ rw[o_DEF] THEN qed[] ; pass ] THEN specialize[ `IMAGE (\f. numpreimages f (perm (complex_root p))) (functions (perm (complex_root p)) (complex_root p))`; `\e. ring_product complex_ring {y:complex | ~(e y = 0)} (\y. (B y) pow (e y))`; `\f. ring_sum complex_ring (perm (complex_root p)) (\i:complex->complex. i (f i))`; `\e. {f | f IN functions (perm (complex_root p)) (complex_root p) /\ numpreimages f (perm (complex_root p)) = e}` ]transcendence_weighted_QinC_monic_vanishing_at THEN specialize_assuming[ `complex_root p`; `B:complex->complex` ]zero_sum_QinC_exp_squarefree_roots_lemma_denouement THEN qed[zero_sum_QinC_exp_squarefree_roots_lemma_denouement;IN] );; let zero_sum_QinC_exp_algebraic = prove(` !S B. FINITE S ==> S SUBSET algebraic_number ==> (!s. s IN S ==> B s IN QinC) ==> ring_sum complex_ring S (\s. (B s) * cexp s) = Cx(&0) ==> (!s. s IN S ==> B s = Cx(&0)) `, intro THEN specialize[`S:complex->bool`]squarefree_from_algebraic_set THEN choose `p:(1->num)->complex` `ring_polynomial QinC_ring p /\ ring_squarefree (x_poly QinC_ring) p /\ monic QinC_ring p /\ S SUBSET complex_root p` [] THEN def `C:complex->complex` `\z:complex. if z IN S then B z else Cx(&0)` THEN subgoal `!z. complex_root p z ==> C z IN QinC` THENL [ intro THEN rw[know `C = (\z:complex. if z IN S then B z else Cx(&0))`] THEN qed[QinC_0] ; pass ] THEN specialize_assuming[ `p:(1->num)->complex`; `C:complex->complex` ]zero_sum_QinC_exp_squarefree_roots THEN subgoal `!z. complex_root p z ==> C z = Cx(&0)` THENL [ subgoal `ring_sum complex_ring (complex_root p) (\z. C z * cexp z) = Cx (&0)` THENL [ subgoal `ring_sum complex_ring (complex_root p) (\z. C z * cexp z) = ring_sum complex_ring (complex_root p) (\s. if s IN S then C s * cexp s else ring_0 complex_ring)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN rw[know `C = (\z:complex. if z IN S then B z else Cx(&0))`] THEN rw[complex_ring_clauses] THEN qed[COMPLEX_MUL_LZERO] ; pass ] THEN rw[know `ring_sum complex_ring (complex_root p) (\z. C z * cexp z) = ring_sum complex_ring (complex_root p) (\s. if s IN S then C s * cexp s else ring_0 complex_ring)`] THEN have `S SUBSET complex_root p` [] THEN specialize[ `complex_ring`; `S:complex->bool`; `complex_root p`; `\z. C z * cexp z` ]ring_sum_restrict_subset THEN rw[know `ring_sum complex_ring (complex_root p) (\s. if s IN S then C s * cexp s else ring_0 complex_ring) = ring_sum complex_ring S (\z. C z * cexp z)`] THEN subgoal `ring_sum complex_ring S (\z. C z * cexp z) = ring_sum complex_ring S (\z. B z * cexp z)` THENL [ sufficesby RING_SUM_EQ THEN rw[know `C = (\z:complex. if z IN S then B z else Cx(&0))`] THEN qed[] ; pass ] THEN simp[] ; pass ] THEN qed[] ; pass ] THEN have `complex_root p s` [IN;SUBSET] THEN have `C(s:complex) = Cx(&0)` [] THEN qed[] );; let transcendental_if_exp_nonzero_algebraic = prove(` !a. algebraic_number a /\ algebraic_number(cexp a) ==> a = Cx(&0) `, intro THEN proven_if `a = Cx(&0)` [] THEN choose `f:(1->num)->complex` `ring_polynomial QinC_ring f /\ ~(f = poly_0 QinC_ring) /\ monic QinC_ring f /\ ring_irreducible(x_poly QinC_ring) f /\ poly_eval complex_ring f (cexp a) = Cx(&0)` [algebraic_number_is_root_monic_irreducible_QinC_poly] THEN have `FINITE(IMAGE (\i. Cx(&i) * a) (0..poly_deg complex_ring(f:(1->num)->complex)))` [FINITE_IMAGE;FINITE_NUMSEG] THEN subgoal `IMAGE (\i. Cx (&i) * a) (0..poly_deg complex_ring(f:(1->num)->complex)) SUBSET algebraic_number` THENL [ rw[SUBSET;IN_IMAGE] THEN qed[algebraic_number_mul;algebraic_number_ZinC;num_in_ZinC;IN] ; pass ] THEN have `!s. s IN IMAGE (\i. Cx (&i) * a) (0..poly_deg complex_ring f) ==> coeff (@i. s = Cx (&i) * a) f IN QinC` [coeff_poly_in_ring;QinC_ring_clauses] THEN subgoal `!i j. Cx(&i)*a = Cx(&j)*a ==> i = j` THENL [ intro THEN complex_field_fact `~(a = Cx(&0)) ==> Cx(&i)*a = Cx(&j)*a ==> Cx(&i) = Cx(&j)` THEN qed[CX_INJ;REAL_OF_NUM_EQ] ; pass ] THEN subgoal `ring_sum complex_ring (IMAGE (\i. Cx (&i) * a) (0..poly_deg complex_ring f)) (\s. coeff (@i. s = Cx (&i) * a) f * cexp s) = Cx (&0)` THENL [ have `cexp a IN ring_carrier complex_ring` [in_complex_ring] THEN subgoal `f IN ring_carrier(poly_ring complex_ring (:1))` THENL [ rw[GSYM x_poly;GSYM x_poly_use] THEN qed[poly_complex_if_poly_QinC] ; pass ] THEN specialize[`complex_ring`;`cexp a`;`f:(1->num)->complex`]POLY_EVAL_EXPAND THEN subgoal `!x y. x IN 0..poly_deg complex_ring(f:(1->num)->complex) ==> y IN 0..poly_deg complex_ring f ==> Cx (&x) * a = Cx (&y) * a ==> x = y` THENL [ intro THEN complex_field_fact `~(a = Cx(&0)) ==> Cx(&x)*a = Cx(&y)*a ==> Cx(&x) = Cx(&y)` THEN qed[CX_INJ;REAL_OF_NUM_EQ] ; pass ] THEN specialize[ `complex_ring`; `\i. Cx(&i) * a`; `\s. coeff (@i. s = Cx (&i) * a) f * cexp s`; `0..poly_deg complex_ring(f:(1->num)->complex)` ]RING_SUM_IMAGE THEN rw[know `ring_sum complex_ring (IMAGE (\i. Cx (&i) * a) (0..poly_deg complex_ring f)) (\s. coeff (@i. s = Cx (&i) * a) f * cexp s) = ring_sum complex_ring (0..poly_deg complex_ring f) ((\s. coeff (@i. s = Cx (&i) * a) f * cexp s) o (\i. Cx (&i) * a))`] THEN rw[o_DEF] THEN subgoal `ring_sum complex_ring (0..poly_deg complex_ring f) (\x. coeff (@i. Cx (&x) * a = Cx (&i) * a) f * cexp (Cx (&x) * a)) = ring_sum complex_ring (0..poly_deg complex_ring f) (\i. ring_mul complex_ring (f (\v. i)) (ring_pow complex_ring (cexp a) i))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `(@i. Cx(&a')*a = Cx(&i)*a) = a'` THENL [ sufficesby SELECT_UNIQUE THEN qed[] ; pass ] THEN rw[know `(@i. Cx(&a')*a = Cx(&i)*a) = a'`] THEN rw[complex_ring_clauses;ring_pow_complex] THEN rw[coeff;x_monomial;CEXP_N] ; pass ] THEN qed[] ; pass ] THEN specialize[ `IMAGE (\i. Cx(&i) * a) (0..poly_deg complex_ring(f:(1->num)->complex))`; `\z. coeff (@i. z = Cx(&i) * a) f:complex` ]zero_sum_QinC_exp_algebraic THEN subgoal `Cx(&(poly_deg complex_ring (f:(1->num)->complex)))*a IN IMAGE (\i.Cx(&i)*a) (0..poly_deg complex_ring f)` THENL [ rw[IN_IMAGE] THEN witness `poly_deg complex_ring (f:(1->num)->complex)` THEN rw[IN_NUMSEG] THEN ARITH_TAC ; pass ] THEN have `coeff (@i. Cx(&(poly_deg complex_ring (f:(1->num)->complex)))*a = Cx (&i)*a) f = Cx (&0)` [] THEN have `coeff (poly_deg complex_ring (f:(1->num)->complex)) f = Cx (&0)` [] THEN have `coeff (poly_deg QinC_ring (f:(1->num)->complex)) f = Cx (&1)` [monic;QinC_ring_clauses] THEN have `coeff (poly_deg complex_ring (f:(1->num)->complex)) f = Cx (&1)` [poly_deg_subring;subring_complex_QinC] THEN qed[CX_INJ;REAL_OF_NUM_EQ;ARITH_RULE `~(1 = 0)`] );; (* ===== expanding prod_f sum_y f(y) exp(y) *) let ring_product_functions_sum_mul_exp_expand = prove(` !(a:A ring) (b:B ring) E:A->B c:C->B Y:A->bool Z:C->bool. (!x:A. x IN ring_carrier a ==> E x IN ring_carrier b) ==> E(ring_0 a) = ring_1 b ==> (!x y:A. E(ring_add a x y) = ring_mul b (E x) (E y)) ==> FINITE Y ==> FINITE Z ==> Y SUBSET ring_carrier a ==> IMAGE c Z SUBSET ring_carrier b ==> ring_product b (functions Y Z) ( \f. ring_sum b Y ( \y. ring_mul b (c (f y)) (E y) ) ) = ring_sum b ( IMAGE (ring_sum a (functions Y Z)) (functions (functions Y Z) Y) ) ( \z. ring_mul b ( ring_sum b {g | g IN functions (functions Y Z) Y /\ ring_sum a (functions Y Z) g = z } (\g. ring_product b (functions Y Z) (\f. c(f(g f))) ) ) (E z) ) `, intro THEN have `FINITE(functions (Y:A->bool) (Z:C->bool))` [finite_functions] THEN subgoal `!p q. p IN functions Y Z ==> q:A IN Y ==> ring_mul b (c (p q:C)) (E q):B IN ring_carrier b` THENL [ intro THEN have `IMAGE (p:A->C) Y SUBSET Z` [image_functions_subset] THEN have `(p:A->C) q IN IMAGE (p:A->C) Y` [IN_IMAGE] THEN have `(p:A->C) q IN Z` [SUBSET] THEN have `c((p:A->C) q):B IN IMAGE c Z` [IN_IMAGE] THEN have `c((p:A->C) q):B IN ring_carrier b` [SUBSET] THEN have `q:A IN ring_carrier a` [SUBSET] THEN have `(E:A->B) q IN ring_carrier b` [] THEN qed[RING_MUL] ; pass ] THEN specialize[ `b:B ring`; `\f:A->C y:A. ring_mul b (c (f y:C)) (E y):B`; `Y:A->bool`; `functions (Y:A->bool) (Z:C->bool)` ]ring_product_sum_expand THEN simp[] THEN have `FINITE(functions (functions (Y:A->bool) (Z:C->bool)) Y)` [finite_functions] THEN specialize[ `b:B ring`; `ring_sum a (functions (Y:A->bool) (Z:C->bool)):((A->C)->A)->A`; `\g. ring_product b (functions Y Z) (\f:A->C. ring_mul b (c (f (g f):C):B) (E (g f)))`; `functions (functions (Y:A->bool) (Z:C->bool)) Y` ]RING_SUM_IMAGE_GEN THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM] THEN intro THEN subgoal `ring_mul b (ring_sum b {g | g IN functions (functions Y Z) Y /\ ring_sum a (functions Y Z) g = a'} (\g. ring_product b (functions Y Z) (\f. c(f (g f):C)))) (E a') = ring_sum b {g | g IN functions (functions Y Z) Y /\ ring_sum a (functions Y Z) g = a'} (\x. ring_mul b (ring_product b (functions Y Z) (\f:A->C. c(f (x f):C))) (E a'):B)` THENL [ have `a':A IN ring_carrier a` [IN_IMAGE;RING_SUM] THEN have `(E:A->B) a' IN ring_carrier b` [] THEN specialize[ `functions (functions (Y:A->bool) (Z:C->bool)) Y`; `\g:(A->C)->A. ring_sum a (functions Y Z) g = a'` ]FINITE_RESTRICT THEN have `!g:(A->C)->A. g IN {g | g IN functions (functions Y Z) Y /\ ring_sum a (functions Y Z) g = a'} ==> ring_product b (functions Y Z) (\f:A->C. c(f (g f):C)):B IN ring_carrier b` [RING_PRODUCT] THEN specialize[ `b:B ring`; `\g. ring_product b (functions Y Z) (\f:A->C. c(f (g f):C)):B`; `(E:A->B) a'`; `{g:(A->C)->A | g IN functions (functions Y Z) Y /\ ring_sum a (functions Y Z) g = a'}` ]RING_SUM_RMUL THEN qed[] ; pass ] THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[IN_ELIM_THM;BETA_THM] THEN intro THEN subgoal `!f:A->C. f IN functions Y Z ==> c(f (a'' f):C):B IN ring_carrier b /\ E (a'' f) IN ring_carrier b` THENL [ intro THEN have `a''(f:A->C):A IN IMAGE a'' (functions Y Z)` [IN_IMAGE] THEN have `IMAGE (a'':(A->C)->A) (functions Y Z) SUBSET Y` [image_functions_subset] THEN have `a''(f:A->C):A IN Y` [SUBSET] THEN have `f(a''(f:A->C)) IN IMAGE f Y` [IN_IMAGE] THEN have `IMAGE (f:A->C) Y SUBSET Z` [image_functions_subset] THEN have `f(a''(f:A->C)) IN Z` [SUBSET] THEN have `c(f(a''(f:A->C))):B IN IMAGE c Z` [IN_IMAGE] THEN qed[SUBSET] ; pass ] THEN specialize[ `b:B ring`; `\f:A->C. c(f(a'' f)):B`; `\f:A->C. E(a'' f:A):B`; `functions (Y:A->bool) (Z:C->bool)` ]RING_PRODUCT_MUL THEN simp[] THEN subgoal `ring_product b (functions Y Z) (\f:A->C. E(a'' f:A):B) = E a'` THENL [ rw[GSYM(know `ring_sum a (functions (Y:A->bool) (Z:C->bool)) a'' = a':A`)] THEN subgoal `!f:A->C. f IN functions Y Z ==> a'' f:A IN ring_carrier a` THENL [ intro THEN have `a''(f:A->C):A IN IMAGE a'' (functions Y Z)` [IN_IMAGE] THEN have `IMAGE (a'':(A->C)->A) (functions Y Z) SUBSET Y` [image_functions_subset] THEN qed[SUBSET] ; pass ] THEN specialize[ `a:A ring`; `b:B ring`; `E:A->B`; `a'':(A->C)->A`; `functions (Y:A->bool) (Z:C->bool)` ]ring_exp_sum THEN simp[] THEN qed[o_DEF] ; pass ] THEN qed[] );; let product_functions_sum_mul_cexp_expand = prove(` !Y:complex->bool Z:complex->bool. FINITE Y ==> FINITE Z ==> ring_product complex_ring (functions Y Z) ( \f. ring_sum complex_ring Y ( \y. ring_mul complex_ring (f y) (cexp y) ) ) = ring_sum complex_ring ( IMAGE (ring_sum complex_ring (functions Y Z)) (functions (functions Y Z) Y) ) ( \z. ring_sum complex_ring {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = z } (\g. ring_product complex_ring (functions Y Z) (\f. f(g f)) ) * cexp z ) `, intro THEN subgoal `ring_product complex_ring (functions Y Z) (\f. ring_sum complex_ring Y (\y. ring_mul complex_ring (f y) (cexp y))) = ring_sum complex_ring (IMAGE (ring_sum complex_ring (functions Y Z)) (functions (functions Y Z) Y)) (\z. ring_mul complex_ring (ring_sum complex_ring {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = z} (\g. ring_product complex_ring (functions Y Z) (\f. f (g f)))) (cexp z))` THENL [ subgoal `ring_product complex_ring (functions Y Z) (\f. ring_sum complex_ring Y (\y. ring_mul complex_ring (f y) (cexp y))) = ring_product complex_ring (functions Y Z) (\f. ring_sum complex_ring Y (\y. ring_mul complex_ring (I(f y)) (cexp y)))` THENL [ rw[I_THM] ; pass ] THEN have `Y SUBSET ring_carrier complex_ring` [SUBSET;in_complex_ring] THEN have `Z SUBSET ring_carrier complex_ring` [SUBSET;in_complex_ring] THEN have `!x. x IN ring_carrier complex_ring ==> cexp x IN ring_carrier complex_ring` [in_complex_ring] THEN have `cexp (ring_0 complex_ring) = ring_1 complex_ring` [CEXP_0;complex_ring_clauses] THEN have `!x y. cexp (ring_add complex_ring x y) = ring_mul complex_ring (cexp x) (cexp y)` [CEXP_ADD;complex_ring_clauses] THEN have `IMAGE I Z SUBSET ring_carrier complex_ring` [IMAGE_I;SUBSET;in_complex_ring] THEN specialize[ `complex_ring`; `complex_ring`; `cexp`; `I:complex->complex` ]ring_product_functions_sum_mul_exp_expand THEN simp[] THEN rw[I_THM] ; pass ] THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[complex_ring_clauses] );; let product_functions_sum_mul_expformal_expand = prove(` !Y:complex->bool Z:complex->bool. FINITE Y ==> FINITE Z ==> ring_product(x_series complex_ring) (functions Y Z) ( \f. ring_sum(x_series complex_ring) Y ( \y. ring_mul(x_series complex_ring) (poly_const complex_ring (f y)) (expformal y) ) ) = ring_sum(x_series complex_ring) ( IMAGE (ring_sum complex_ring (functions Y Z)) (functions (functions Y Z) Y) ) ( \z. poly_mul complex_ring ( poly_const complex_ring ( ring_sum complex_ring {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = z } (\g. ring_product complex_ring (functions Y Z) (\f. f(g f)) ) ) ) (expformal z) ) `, intro THEN have `!x. x IN ring_carrier complex_ring ==> expformal x IN ring_carrier (x_series complex_ring)` [series_complex;x_series_use] THEN have `expformal (ring_0 complex_ring) = ring_1 (x_series complex_ring)` [expformal_0;complex_ring_clauses;x_series_use] THEN have `!x y. expformal (ring_add complex_ring x y) = ring_mul (x_series complex_ring) (expformal x) (expformal y)` [mul_expformal;complex_ring_clauses;x_series_use] THEN have `Y SUBSET ring_carrier complex_ring` [SUBSET;in_complex_ring] THEN have `IMAGE (\c. poly_const complex_ring c) Z SUBSET ring_carrier (x_series complex_ring)` [SUBSET;series_complex;x_series_use] THEN specialize[ `complex_ring`; `x_series complex_ring`; `expformal`; `\c. poly_const complex_ring c:(1->num)->complex`; `Y:complex->bool`; `Z:complex->bool`; ]ring_product_functions_sum_mul_exp_expand THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_IMAGE] THEN intro THEN rw[x_series_use] THEN subgoal `ring_sum (x_series complex_ring) {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = a} (\g. ring_product (x_series complex_ring) (functions Y Z) (\f. poly_const complex_ring (f (g f)))) = poly_const complex_ring (ring_sum complex_ring {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = a} (\g. ring_product complex_ring (functions Y Z) (\f. f (g f))))` THENL [ have `FINITE(functions (Y:complex->bool) (Z:complex->bool))` [finite_functions] THEN have `FINITE(functions (functions (Y:complex->bool) (Z:complex->bool)) Y)` [finite_functions] THEN specialize[ `functions (functions (Y:complex->bool) (Z:complex->bool)) Y`; `\g:(complex->complex)->complex. ring_sum complex_ring (functions Y Z) g = a` ]FINITE_RESTRICT THEN have `!s. s IN {g | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = a} ==> ring_product complex_ring (functions Y Z) (\f. f (s f)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `\g:(complex->complex)->complex. ring_product complex_ring (functions Y Z) (\f. f (g f))`; `{g:(complex->complex)->complex | g IN functions (functions Y Z) Y /\ ring_sum complex_ring (functions Y Z) g = a}` ]poly_const_sum THEN simp[] THEN rw[poly_sum] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN have `!t:complex->complex. t IN functions Y Z ==> t (a' t) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `\f:complex->complex. f(a' f)`; `functions (Y:complex->bool) (Z:complex->bool)` ]poly_const_product THEN simp[] THEN rw[poly_product] ; pass ] THEN simp[] );; (* ===== more general transcendence results, part 2: beta symmetrization *) let zero_sum_nonzero_algebraic_exp_algebraic_lemma_sym = prove(` !s S:complex->bool p. FINITE S ==> ring_polynomial QinC_ring p ==> ring_squarefree(x_poly QinC_ring) p ==> monic QinC_ring p ==> ring_sum complex_ring {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s } (\g. ring_product complex_ring (functions S (complex_root p)) (\f. f (g f)) ) IN QinC `, intro THEN def `Z:(complex->num)->complex` `ring_sum (poly_ring complex_ring (:complex)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\g. ring_product (poly_ring complex_ring (:complex)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f))))` THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `~(p = poly_0 QinC_ring:(1->num)->complex)` [monic_poly_0;ring_1_0_QinC] THEN have `~(p = poly_0 complex_ring:(1->num)->complex)` [poly_0_QinC_eq_poly_0_complex] THEN have `FINITE (complex_root p)` [complex_root_le_deg] THEN have `FINITE (functions (S:complex->bool) (complex_root p))` [finite_functions] THEN have `FINITE (functions (functions (S:complex->bool) (complex_root p)) S)` [finite_functions] THEN specialize[ `functions (functions (S:complex->bool) (complex_root p)) S`; `\g. ring_sum complex_ring (functions (S:complex->bool) (complex_root p)) g = s` ]FINITE_RESTRICT THEN have `!s. s IN complex_root p ==> I s IN ring_carrier complex_ring` [in_complex_ring] THEN have `!d. coeff d (monic_vanishing_at complex_ring (complex_root p) I) IN ring_carrier (subring_generated complex_ring QinC)` [monic_QinC_squarefree_complex_squarefree;monic_subring;subring_complex_QinC;monic_vanishing_at_complex_root;subring_complex_QinC;coeff_poly_in_ring] THEN subgoal `ring_polynomial (subring_generated complex_ring QinC) (Z:(complex->num)->complex)` THENL [ simp[] THEN sufficesby ring_polynomial_subring_sum THEN rw[BETA_THM] THEN intro THENL [ qed[] ; sufficesby ring_polynomial_subring_product THEN qed[ring_polynomial_subring_var] ] ; pass ] THEN subgoal `poly_vars complex_ring (Z:(complex->num)->complex) SUBSET complex_root p` THENL [ simp[] THEN sufficesby poly_vars_sum_poly_subset THEN rw[IN_ELIM_THM] THEN intro THENL [ qed[] ; qed[RING_PRODUCT;poly_in_full_ring] ; sufficesby poly_vars_product_poly_subset THEN intro THENL [ qed[] ; qed[RING_POLYNOMIAL_VAR] ; have `~trivial_ring complex_ring` [TRIVIAL_RING_10;ring_1_0_complex] THEN simp[POLY_VARS_VAR;SUBSET;IN_SING] THEN qed[functions_to] ] ] ; pass ] THEN subgoal `!q m. q permutes complex_root p ==> Z (m o q) = Z(m:complex->num):complex` THENL [ simp[] THEN intro THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\g. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f)))) (m o q) = ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\s. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (q(f (s f))))) m` THENL [ subgoal `!g. g IN {g | g IN functions (functions (S:complex->bool) (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ==> ring_polynomial complex_ring (ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f))))` THENL [ qed[RING_PRODUCT;poly_in_full_ring] ; pass ] THEN specialize[ `complex_ring`; `\g. ring_product (poly_ring complex_ring (:real^2)) (functions (S:complex->bool) (complex_root p)) (\f. poly_var complex_ring (f (g f)))`; `m:complex->num`; `q:complex->complex`; `complex_root p`; `{g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions (S:complex->bool) (complex_root p)) g = s}` ]ring_sum_poly_o_permutes THEN simp[] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\s m. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (s f))) (m o q)) = ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\s. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (q (f (s f)))))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;FUN_EQ_THM;IN_ELIM_THM] THEN intro THEN have `!s:complex->complex. s IN functions S (complex_root p) ==> ring_polynomial complex_ring (poly_var complex_ring (s (a s)))` [RING_POLYNOMIAL_VAR] THEN specialize_assuming[ `complex_ring`; `\f:complex->complex. poly_var complex_ring (f (a f))`; `m:complex->num`; `q:complex->complex`; `complex_root p`; `functions (S:complex->bool) (complex_root p)` ]ring_product_poly_o_permutes THEN simp[] THEN subgoal `ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\s:complex->complex m. poly_var complex_ring (s (a s)) (m o q)) = ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (q (f (a f))))` THENL [ sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM;FUN_EQ_THM] THEN qed[poly_var_o_permutes;o_THM] ; pass ] THEN qed[] ; pass ] THEN qed[] ; pass ] THEN simp[] THEN have `inverse q permutes complex_root p` [PERMUTES_INVERSE] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\g. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f)))) = ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ((\g. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f)))) o ((\g ab. if ab IN functions S (complex_root p) then g (\a. if a IN S then inverse q (ab a) else ARB) else ARB)))` THENL [ subgoal `!x:(complex->complex)->complex y. x IN {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ==> y IN {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ==> (\ab. if ab IN functions S (complex_root p) then x (\a. if a IN S then (inverse q:complex->complex) (ab a) else ARB) else ARB) = (\ab. if ab IN functions S (complex_root p) then y (\a. if a IN S then inverse q (ab a) else ARB) else ARB) ==> x = y` THENL [ rw[IN_ELIM_THM] THEN intro THEN specialize_assuming[ `S:complex->bool`; `complex_root p`; `S:complex->bool`; `inverse q:complex->complex`; `x:(complex->complex)->complex`; `y:(complex->complex)->complex` ]injective_permutes_arbo_functions_functions THEN qed[] ; pass ] THEN specialize[ `poly_ring complex_ring (:complex)`; `\g ab. if ab IN functions (S:complex->bool) (complex_root p) then g (\a. if a IN S then inverse q (ab a) else ARB:complex) else ARB:complex`; `\g:(complex->complex)->complex. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f)))`; `{g:(complex->complex)->complex | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s}` ]RING_SUM_IMAGE THEN specialize[ `S:complex->bool`; `complex_root p`; `S:complex->bool`; `inverse q:complex->complex` ]image_permutes_arbo_functions_functions THEN subgoal `IMAGE (\g ab. if ab IN functions S (complex_root p) then g (\a. if a IN S then inverse q (ab a) else ARB) else ARB) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} = {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s}` THENL [ rw[EXTENSION;in_image_vw;IN_ELIM_THM] THEN intro THEN subgoal `!w:(complex->complex)->complex. w IN functions(functions S (complex_root p)) S ==> ring_sum complex_ring (functions S (complex_root p)) (\ab. if ab IN functions S (complex_root p) then w (\a. if a IN S then inverse q (ab a) else ARB) else ARB) = ring_sum complex_ring (functions S (complex_root p)) w` THENL [ intro THEN subgoal `ring_sum complex_ring (functions S (complex_root p)) (\ab:complex->complex. if ab IN functions S (complex_root p) then w (\a. if a IN S then (inverse q:complex->complex)(ab a) else ARB) else ARB) = ring_sum complex_ring (functions S (complex_root p)) (\ab. w(\a. if a IN S then inverse q(ab a) else ARB))` THENL [ sufficesby RING_SUM_EQ THEN qed[] ; pass ] THEN once_rw[know `ring_sum complex_ring (functions S (complex_root p)) (\ab:complex->complex. if ab IN functions S (complex_root p) then w (\a. if a IN S then (inverse q:complex->complex)(ab a) else ARB) else ARB) = ring_sum complex_ring (functions S (complex_root p)) (\ab. w(\a. if a IN S then inverse q(ab a) else ARB))`] THEN specialize[ `S:complex->bool`; `complex_root p`; `inverse q:complex->complex` ]injective_permutes_arbo_functions THEN have `!x:complex->complex y. x IN functions S (complex_root p) ==> y IN functions S (complex_root p) ==> (\a. if a IN S then (inverse q:complex->complex) (x a) else ARB) = (\a. if a IN S then inverse q (y a) else ARB) ==> x = y` [] THEN specialize[ `complex_ring`; `\ab a:complex. if a IN S then (inverse q:complex->complex)(ab a) else ARB`; `w:(complex->complex)->complex`; `functions (S:complex->bool) (complex_root p)` ]RING_SUM_IMAGE THEN specialize[ `S:complex->bool`; `complex_root p`; `inverse q:complex->complex` ]image_permutes_arbo_functions THEN have_rw `ring_sum complex_ring (functions S (complex_root p)) w = ring_sum complex_ring (functions S (complex_root p)) (w o (\ab a:complex. if a IN S then (inverse q:complex->complex) (ab a) else ARB))` [] THEN sufficesby RING_SUM_EQ THEN rw[BETA_THM;o_THM] ; pass ] THEN splitiff THENL [ qed[IN_IMAGE] ; intro THEN have `x IN IMAGE (\bc ab. if ab IN functions S (complex_root p) then bc (\a. if a IN S then inverse q (ab a) else ARB) else ARB) (functions (functions S (complex_root p)) (S:complex->bool))` [] THEN choose `v:(complex->complex)->complex` `x = (\bc ab. if ab IN functions S (complex_root p) then bc (\a. if a IN S then inverse q (ab a) else ARB) else ARB) v /\ v IN functions (functions S (complex_root p)) (S:complex->bool)` [IN_IMAGE] THEN witness `v:(complex->complex)->complex` THEN qed[] ] ; pass ] THEN qed[] ; pass ] THEN simp[] THEN subgoal `ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\s. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (q (f (s f))))) = ring_sum (poly_ring complex_ring (:real^2)) {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ((\g. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f)))) o (\g ab. if ab IN functions S (complex_root p) then g (\a. if a IN S then inverse q (ab a) else ARB) else ARB))` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM;o_THM] THEN intro THEN subgoal `!x:complex->complex y. x IN functions S (complex_root p) ==> y IN functions S (complex_root p) ==> (\a. if a IN S then (inverse q:complex->complex) (x a) else ARB) = (\a. if a IN S then inverse q (y a) else ARB) ==> x = y` THENL [ intro THEN specialize[ `S:complex->bool`; `complex_root p`; `inverse q:complex->complex`; `x:complex->complex`; `y:complex->complex` ]injective_permutes_arbo_functions THEN qed[] ; pass ] THEN specialize[ `poly_ring complex_ring (:complex)`; `\f:complex->complex a. if a IN S then (inverse q:complex->complex)(f a) else ARB`; `\f:complex->complex. poly_var complex_ring (q(f(a f)):complex)`; `functions (S:complex->bool) (complex_root p)` ]RING_PRODUCT_IMAGE THEN specialize[ `S:complex->bool`; `complex_root p`; `inverse q:complex->complex`; ]image_permutes_arbo_functions THEN have_rw `ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (q (f (a f)))) = ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) ((\f. poly_var complex_ring (q (f (a f)))) o (\f a:complex. if a IN S then inverse q (f a) else ARB))` [] THEN sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM;o_THM] THEN intro THEN have `(\c:complex. if c IN S then inverse (q:complex->complex) (a' c) else ARB) IN functions S (complex_root p)` [IN_IMAGE] THEN have_rw `a(\c:complex. if c IN S then inverse (q:complex->complex) (a' c) else ARB) IN S` [functions_to] THEN rw[know `a':complex->complex IN functions S (complex_root p)`] THEN qed[PERMUTES_INVERSES] ; pass ] THEN simp[] ; pass ] THEN specialize[ `complex_ring`; `QinC`; `complex_root p`; `I:complex->complex`; `Z:(complex->num)->complex` ]symmetric_subring_if_poly_subring THEN subgoal `poly_evaluate complex_ring Z I = ring_sum complex_ring {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\g. ring_product complex_ring (functions S (complex_root p)) (\f. f (g f)))` THENL [ simp[] THEN have `!g. g IN {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} ==> ring_polynomial complex_ring (ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f. poly_var complex_ring (f (g f))))` [RING_PRODUCT;poly_in_full_ring] THEN specialize[ `complex_ring`; `\g. ring_product (poly_ring complex_ring (:real^2)) (functions S (complex_root p)) (\f:complex->complex. poly_var complex_ring (f (g f)))`; `I:complex->complex`; `{g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s}` ]poly_evaluate_sum THEN simp[] THEN sufficesby RING_SUM_EQ THEN rw[IN_ELIM_THM;BETA_THM] THEN intro THEN have `!v:complex. I v IN ring_carrier complex_ring` [in_complex_ring] THEN have `!s:complex->complex. s IN functions S (complex_root p) ==> ring_polynomial complex_ring (poly_var complex_ring (s (a s)))` [RING_POLYNOMIAL_VAR] THEN specialize[ `complex_ring`; `\f:complex->complex. poly_var complex_ring (f (a f))`; `I:complex->complex`; `functions (S:complex->bool) (complex_root p)` ]poly_evaluate_product THEN simp[] THEN sufficesby RING_PRODUCT_EQ THEN rw[BETA_THM] THEN intro THEN simp[POLY_EVALUATE_VAR;in_complex_ring;I_THM] ; pass ] THEN qed[subring_complex_QinC;QinC_ring_clauses] );; let zero_sum_nonzero_algebraic_exp_algebraic = prove(` !S B. FINITE S ==> S SUBSET algebraic_number ==> (!s. s IN S ==> algebraic_number (B s)) ==> (!s. s IN S ==> ~(B s = Cx(&0))) ==> ring_sum complex_ring S (\s. (B s) * cexp s) = Cx(&0) ==> S = {} `, intro THEN have `FINITE(IMAGE (B:complex->complex) S)` [FINITE_IMAGE] THEN have `~(Cx(&0) IN IMAGE (B:complex->complex) S)` [IN_IMAGE] THEN subgoal `IMAGE (B:complex->complex) S SUBSET algebraic_number` THENL [ rw[SUBSET;IN_IMAGE] THEN qed[IN] ; pass ] THEN specialize[ `IMAGE (B:complex->complex) S` ]squarefree_from_algebraic_set_avoiding_0 THEN choose `p:(1->num)->complex` `ring_polynomial QinC_ring p /\ ring_squarefree (x_poly QinC_ring) p /\ monic QinC_ring p /\ IMAGE (B:complex->complex) S SUBSET complex_root p /\ ~complex_root p (Cx (&0))` [] THEN have `ring_polynomial complex_ring (p:(1->num)->complex)` [poly_complex_if_poly_QinC] THEN have `monic complex_ring (p:(1->num)->complex)` [monic_subring;subring_complex_QinC] THEN have `~(p:(1->num)->complex = poly_0 complex_ring)` [monic_poly_0;ring_1_0_complex] THEN have `FINITE (complex_root p)` [complex_root_le_deg] THEN have `FINITE (functions (S:complex->bool) (complex_root p))` [finite_functions] THEN have `!s:complex. s IN S ==> B s IN complex_root p` [IN_IMAGE;SUBSET] THEN subgoal `ring_product complex_ring (functions S (complex_root p)) (\f. ring_sum complex_ring S (\y. ring_mul complex_ring (f y) (cexp y))) = Cx(&0)` THENL [ subgoal `(\y:complex. if y IN S then B y else ARB) IN functions S (complex_root p)` THENL [ rw[in_functions;IN_IMAGE;SUBSET] THEN qed[] ; pass ] THEN have `ring_sum complex_ring S (\y. ring_mul complex_ring (if y IN S then B y else ARB) (cexp y)) IN ring_carrier complex_ring` [in_complex_ring] THEN specialize[ `complex_ring`; `functions (S:complex->bool) (complex_root p)`; `\y:complex. if y IN S then B y else ARB:complex`; `\f. ring_sum complex_ring S (\y. ring_mul complex_ring (f y) (cexp y))` ]ring_product_delete THEN subgoal `ring_sum complex_ring S (\y. ring_mul complex_ring (if y IN S then B y else ARB) (cexp y)) = Cx(&0)` THENL [ subgoal `ring_sum complex_ring S (\y. ring_mul complex_ring (if y IN S then B y else ARB) (cexp y)) = ring_sum complex_ring S (\s. B s * cexp s)` THENL [ sufficesby RING_SUM_EQ THEN qed[complex_ring_clauses] ; pass ] THEN qed[] ; pass ] THEN simp[] THEN rw[complex_ring_clauses] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN specialize[ `S:complex->bool`; `complex_root p` ]product_functions_sum_mul_cexp_expand THEN have `ring_sum complex_ring (IMAGE (ring_sum complex_ring (functions S (complex_root p))) (functions (functions S (complex_root p)) S)) (\z. ring_sum complex_ring {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = z} (\g. ring_product complex_ring (functions S (complex_root p)) (\f. f (g f))) * cexp z) = Cx(&0)` [] THEN have `FINITE (functions (functions (S:complex->bool) (complex_root p)) S)` [finite_functions] THEN have `FINITE (IMAGE (ring_sum complex_ring (functions S (complex_root p))) (functions (functions S (complex_root p)) (S:complex->bool)))` [FINITE_IMAGE] THEN subgoal `IMAGE (ring_sum complex_ring (functions S (complex_root p))) (functions (functions S (complex_root p)) (S:complex->bool)) SUBSET algebraic_number` THENL [ rw[SUBSET;in_image_vw] THEN intro THEN simp[IN] THEN sufficesby algebraic_number_sum THEN simp[] THEN intro THEN have `v s:complex IN IMAGE v (functions (S:complex->bool) (complex_root p))` [IN_IMAGE] THEN have `v(s:complex->complex):complex IN S` [image_functions_subset;SUBSET] THEN qed[SUBSET;IN] ; pass ] THEN have `!s. s IN IMAGE (ring_sum complex_ring (functions S (complex_root p))) (functions (functions S (complex_root p)) S) ==> ring_sum complex_ring {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = s} (\g. ring_product complex_ring (functions S (complex_root p)) (\f. f (g f))) IN QinC` [zero_sum_nonzero_algebraic_exp_algebraic_lemma_sym] THEN specialize[ `IMAGE (ring_sum complex_ring (functions (S:complex->bool) (complex_root p))) (functions (functions S (complex_root p)) S)`; `\z. ring_sum complex_ring {g | g IN functions (functions S (complex_root p)) S /\ ring_sum complex_ring (functions S (complex_root p)) g = z} (\g. ring_product complex_ring (functions S (complex_root p)) (\f. f (g f)))` ]zero_sum_QinC_exp_algebraic THEN subgoal `ring_product (x_series complex_ring) (functions S (complex_root p)) (\f. ring_sum (x_series complex_ring) S (\y. ring_mul (x_series complex_ring) (poly_const complex_ring (f y)) (expformal y))) = ring_0(x_series complex_ring)` THENL [ specialize[ `S:complex->bool`; `complex_root p` ]product_functions_sum_mul_expformal_expand THEN simp[] THEN sufficesby RING_SUM_EQ_0 THEN rw[IN_IMAGE;BETA_THM] THEN rw[GSYM x_series_use;GSYM complex_ring_clauses] THEN rw[GSYM poly_0] THEN qed[POWSER_MUL_0;series_complex] ; pass ] THEN have `integral_domain (x_series complex_ring)` [INTEGRAL_DOMAIN_POWSER_RING;integral_domain_complex;x_series] THEN choose `f:complex->complex` `f:complex->complex IN functions S (complex_root p) /\ ring_sum (x_series complex_ring) S (\y. ring_mul (x_series complex_ring) (poly_const complex_ring (f y)) (expformal y)) = ring_0 (x_series complex_ring)` [INTEGRAL_DOMAIN_PRODUCT_EQ_0] THEN subgoal `poly_sum complex_ring S (\y. poly_mul complex_ring (poly_const complex_ring (f y)) (expformal y)) = poly_0 complex_ring` THENL [ rw[x_series_use;poly_sum] THEN qed[] ; pass ] THEN have `!s:complex. s IN S ==> f s = Cx(&0)` [expformal_linearly_independent] THEN proven_if `S = {}:complex->bool` [] THEN choose `s:complex` `s:complex IN S` [MEMBER_NOT_EMPTY] THEN have `f(s:complex) = Cx(&0)` [] THEN have `f(s:complex) IN complex_root p` [functions_to] THEN qed[IN] );; let zero_sum_algebraic_exp_algebraic = prove(` !S B. FINITE S /\ S SUBSET algebraic_number /\ (!s. s IN S ==> algebraic_number (B s)) /\ ring_sum complex_ring S (\s. (B s) * cexp s) = Cx(&0) ==> (!s. s IN S ==> B s = Cx(&0)) `, intro THEN set_fact `{s:complex | s IN S /\ ~(B s = ring_0 complex_ring)} SUBSET S` THEN subgoal `ring_sum complex_ring {s | s IN S /\ ~(B s = ring_0 complex_ring)} (\s. B s * cexp s) = Cx(&0)` THENL [ rw[GSYM complex_ring_clauses] THEN specialize[ `complex_ring`; `{s:complex | s IN S /\ ~(B s = ring_0 complex_ring)}`; `S:complex->bool`; `\s. ring_mul complex_ring (B s) (cexp s)` ](GSYM ring_sum_restrict_subset) THEN simp[] THEN subgoal `ring_sum complex_ring S (\s. if s IN {s | s IN S /\ ~(B s = ring_0 complex_ring)} then ring_mul complex_ring (B s) (cexp s) else ring_0 complex_ring) = ring_sum complex_ring S (\s. B s * cexp s)` THENL [ sufficesby RING_SUM_EQ THEN rw[BETA_THM;IN_ELIM_THM] THEN intro THEN rw[complex_ring_clauses] THEN simp[] THEN CONV_TAC COMPLEX_FIELD ; pass ] THEN qed[complex_ring_clauses] ; pass ] THEN specialize[ `S:complex->bool`; `\s:complex. ~(B s = ring_0 complex_ring)` ]FINITE_RESTRICT THEN have `{s:complex | s IN S /\ ~(B s = ring_0 complex_ring)} SUBSET algebraic_number` [SUBSET_TRANS] THEN have `!s:complex. s IN {s | s IN S /\ ~(B s = ring_0 complex_ring)} ==> algebraic_number (B s)` [SUBSET] THEN subgoal `!s:complex. s IN {s | s IN S /\ ~(B s = ring_0 complex_ring)} ==> ~(B s = Cx (&0))` THENL [ rw[IN_ELIM_THM] THEN qed[complex_ring_clauses] ; pass ] THEN specialize[ `{s:complex | s IN S /\ ~(B s = ring_0 complex_ring)}`; `B:complex->complex` ]zero_sum_nonzero_algebraic_exp_algebraic THEN ASM SET_TAC[complex_ring_clauses] );; (* ===== re-print highlighted theorems *) let e_is_irrational = e_is_irrational;; let e_is_transcendental = e_is_transcendental;; let pi_is_transcendental = pi_is_transcendental;; let transcendental_if_exp_nonzero_algebraic = transcendental_if_exp_nonzero_algebraic;; let zero_sum_algebraic_exp_algebraic = zero_sum_algebraic_exp_algebraic;;